removed obsolete HOL-Boogie session;
authorboehmes
Tue, 23 Jul 2013 18:36:23 +0200
changeset 52722 2c81f7baf8c4
parent 52721 6bafe21b13b2
child 52724 f547266a9338
removed obsolete HOL-Boogie session; keep examples that also test SMT solvers, using a minimal version of the old Boogie loader
src/HOL/Boogie/Boogie.thy
src/HOL/Boogie/Examples/Boogie_Dijkstra.b2i
src/HOL/Boogie/Examples/Boogie_Dijkstra.certs
src/HOL/Boogie/Examples/Boogie_Dijkstra.thy
src/HOL/Boogie/Examples/Boogie_Max.b2i
src/HOL/Boogie/Examples/Boogie_Max.certs
src/HOL/Boogie/Examples/Boogie_Max.thy
src/HOL/Boogie/Examples/Boogie_Max_Stepwise.thy
src/HOL/Boogie/Examples/VCC_Max.b2i
src/HOL/Boogie/Examples/VCC_Max.certs
src/HOL/Boogie/Examples/VCC_Max.thy
src/HOL/Boogie/Tools/boogie_commands.ML
src/HOL/Boogie/Tools/boogie_loader.ML
src/HOL/Boogie/Tools/boogie_tactics.ML
src/HOL/Boogie/Tools/boogie_vcs.ML
src/HOL/ROOT
src/HOL/SMT_Examples/Boogie.thy
src/HOL/SMT_Examples/Boogie_Dijkstra.b2i
src/HOL/SMT_Examples/Boogie_Dijkstra.certs
src/HOL/SMT_Examples/Boogie_Max.b2i
src/HOL/SMT_Examples/Boogie_Max.certs
src/HOL/SMT_Examples/VCC_Max.b2i
src/HOL/SMT_Examples/VCC_Max.certs
src/HOL/SMT_Examples/boogie.ML
--- a/src/HOL/Boogie/Boogie.thy	Tue Jul 23 13:14:14 2013 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,104 +0,0 @@
-(*  Title:      HOL/Boogie/Boogie.thy
-    Author:     Sascha Boehme, TU Muenchen
-*)
-
-header {* Integration of the Boogie program verifier *}
-
-theory Boogie
-imports Word
-keywords
-  "boogie_open" :: thy_load and
-  "boogie_end" :: thy_decl and
-  "boogie_vc" :: thy_goal and
-  "boogie_status" :: diag
-begin
-
-text {*
-HOL-Boogie and its applications are described in:
-\begin{itemize}
-
-\item Sascha B"ohme, K. Rustan M. Leino, and Burkhart Wolff.
-HOL-Boogie --- An Interactive Prover for the Boogie Program-Verifier.
-Theorem Proving in Higher Order Logics, 2008.
-
-\item Sascha B"ohme, Micha{\l} Moskal, Wolfram Schulte, and Burkhart Wolff.
-HOL-Boogie --- An Interactive Prover-Backend for the Verifying C Compiler.
-Journal of Automated Reasoning, 2009.
-
-\end{itemize}
-*}
-
-
-
-section {* Built-in types and functions of Boogie *}
-
-subsection {* Labels *}
-
-text {*
-See "Generating error traces from verification-condition counterexamples"
-by Leino e.a. (2004) for a description of Boogie's labelling mechanism and
-semantics.
-*}
-
-definition assert_at :: "bool \<Rightarrow> bool \<Rightarrow> bool" where "assert_at l P = P"
-definition block_at :: "bool \<Rightarrow> bool \<Rightarrow> bool" where "block_at l P = P"
-
-lemmas labels = assert_at_def block_at_def
-
-
-subsection {* Integer arithmetics *}
-
-text {*
-The operations @{text div} and @{text mod} are built-in in Boogie, but
-without a particular semantics due to different interpretations in
-programming languages. We assume that each application comes with a
-proper axiomatization.
-*}
-
-axiomatization
-  boogie_div :: "int \<Rightarrow> int \<Rightarrow> int" (infixl "div'_b" 70) and
-  boogie_mod :: "int \<Rightarrow> int \<Rightarrow> int" (infixl "mod'_b" 70)
-
-
-
-section {* Boogie environment *}
-
-text {*
-Proving Boogie-generated verification conditions happens inside
-a Boogie environment:
-
-  boogie_open "b2i file with extension"
-  boogie_vc "verification condition 1" ...
-  boogie_vc "verification condition 2" ...
-  boogie_vc "verification condition 3" ...
-  boogie_end
-
-See the Boogie examples for more details.
- 
-At most one Boogie environment should occur per theory,
-otherwise unexpected effects may arise.
-*}
-
-
-
-section {* Setup *}
-
-ML {*
-structure Boogie_Axioms = Named_Thms
-(
-  val name = @{binding boogie}
-  val description =
-    "Boogie background axioms loaded along with Boogie verification conditions"
-)
-*}
-setup Boogie_Axioms.setup
-
-ML_file "Tools/boogie_vcs.ML"
-ML_file "Tools/boogie_loader.ML"
-ML_file "Tools/boogie_tactics.ML"
-setup Boogie_Tactics.setup
-
-ML_file "Tools/boogie_commands.ML"
-setup Boogie_Commands.setup
-
-end
--- a/src/HOL/Boogie/Examples/Boogie_Dijkstra.b2i	Tue Jul 23 13:14:14 2013 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,1879 +0,0 @@
-type-decl Vertex 0 0
-fun-decl G 1 0
-    array 3
-      type-con Vertex 0
-      type-con Vertex 0
-      int
-fun-decl Infinity 1 0
-    int
-fun-decl Source 1 0
-    type-con Vertex 0
-axiom 0
-    forall 2 0 3
-      var x
-        type-con Vertex 0
-      var y
-        type-con Vertex 0
-      attribute qid 1
-        string-attr BoogieDi.3:15
-      attribute uniqueId 1
-        string-attr 0
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    not
-    =
-    var x
-      type-con Vertex 0
-    var y
-      type-con Vertex 0
-    <
-    int-num 0
-    select 3
-    fun G 0
-    var x
-      type-con Vertex 0
-    var y
-      type-con Vertex 0
-axiom 0
-    forall 2 0 3
-      var x
-        type-con Vertex 0
-      var y
-        type-con Vertex 0
-      attribute qid 1
-        string-attr BoogieDi.4:15
-      attribute uniqueId 1
-        string-attr 1
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    =
-    var x
-      type-con Vertex 0
-    var y
-      type-con Vertex 0
-    =
-    select 3
-    fun G 0
-    var x
-      type-con Vertex 0
-    var y
-      type-con Vertex 0
-    int-num 0
-axiom 0
-    <
-    int-num 0
-    fun Infinity 0
-var-decl SP 0
-    array 2
-      type-con Vertex 0
-      int
-vc Dijkstra 1
-    implies
-    label pos 26 3
-    true
-    implies
-    true
-    implies
-    forall 1 0 3
-      var x
-        type-con Vertex 0
-      attribute qid 1
-        string-attr BoogieDi.27:18
-      attribute uniqueId 1
-        string-attr 5
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    =
-    var x
-      type-con Vertex 0
-    fun Source 0
-    =
-    select 2
-    var SP@0
-      array 2
-        type-con Vertex 0
-        int
-    var x
-      type-con Vertex 0
-    int-num 0
-    implies
-    forall 1 0 3
-      var x
-        type-con Vertex 0
-      attribute qid 1
-        string-attr BoogieDi.28:18
-      attribute uniqueId 1
-        string-attr 6
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    not
-    =
-    var x
-      type-con Vertex 0
-    fun Source 0
-    =
-    select 2
-    var SP@0
-      array 2
-        type-con Vertex 0
-        int
-    var x
-      type-con Vertex 0
-    fun Infinity 0
-    implies
-    forall 1 0 3
-      var x
-        type-con Vertex 0
-      attribute qid 1
-        string-attr BoogieDi.31:18
-      attribute uniqueId 1
-        string-attr 7
-      attribute bvZ3Native 1
-        string-attr False
-    not
-    select 2
-    var Visited@0
-      array 2
-        type-con Vertex 0
-        bool
-    var x
-      type-con Vertex 0
-    implies
-    true
-    and 2
-    label neg 34 5
-    =
-    select 2
-    var SP@0
-      array 2
-        type-con Vertex 0
-        int
-    fun Source 0
-    int-num 0
-    implies
-    =
-    select 2
-    var SP@0
-      array 2
-        type-con Vertex 0
-        int
-    fun Source 0
-    int-num 0
-    and 2
-    label neg 35 5
-    forall 1 0 3
-      var x
-        type-con Vertex 0
-      attribute qid 1
-        string-attr BoogieDi.35:23
-      attribute uniqueId 1
-        string-attr 9
-      attribute bvZ3Native 1
-        string-attr False
-    >=
-    select 2
-    var SP@0
-      array 2
-        type-con Vertex 0
-        int
-    var x
-      type-con Vertex 0
-    int-num 0
-    implies
-    forall 1 0 3
-      var x
-        type-con Vertex 0
-      attribute qid 1
-        string-attr BoogieDi.35:23
-      attribute uniqueId 1
-        string-attr 9
-      attribute bvZ3Native 1
-        string-attr False
-    >=
-    select 2
-    var SP@0
-      array 2
-        type-con Vertex 0
-        int
-    var x
-      type-con Vertex 0
-    int-num 0
-    and 2
-    label neg 36 5
-    forall 2 0 3
-      var y
-        type-con Vertex 0
-      var z
-        type-con Vertex 0
-      attribute qid 1
-        string-attr BoogieDi.36:23
-      attribute uniqueId 1
-        string-attr 10
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 2
-    not
-    select 2
-    var Visited@0
-      array 2
-        type-con Vertex 0
-        bool
-    var z
-      type-con Vertex 0
-    select 2
-    var Visited@0
-      array 2
-        type-con Vertex 0
-        bool
-    var y
-      type-con Vertex 0
-    <=
-    select 2
-    var SP@0
-      array 2
-        type-con Vertex 0
-        int
-    var y
-      type-con Vertex 0
-    select 2
-    var SP@0
-      array 2
-        type-con Vertex 0
-        int
-    var z
-      type-con Vertex 0
-    implies
-    forall 2 0 3
-      var y
-        type-con Vertex 0
-      var z
-        type-con Vertex 0
-      attribute qid 1
-        string-attr BoogieDi.36:23
-      attribute uniqueId 1
-        string-attr 10
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 2
-    not
-    select 2
-    var Visited@0
-      array 2
-        type-con Vertex 0
-        bool
-    var z
-      type-con Vertex 0
-    select 2
-    var Visited@0
-      array 2
-        type-con Vertex 0
-        bool
-    var y
-      type-con Vertex 0
-    <=
-    select 2
-    var SP@0
-      array 2
-        type-con Vertex 0
-        int
-    var y
-      type-con Vertex 0
-    select 2
-    var SP@0
-      array 2
-        type-con Vertex 0
-        int
-    var z
-      type-con Vertex 0
-    and 2
-    label neg 38 5
-    forall 2 0 3
-      var z
-        type-con Vertex 0
-      var y
-        type-con Vertex 0
-      attribute qid 1
-        string-attr BoogieDi.38:23
-      attribute uniqueId 1
-        string-attr 11
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 2
-    select 2
-    var Visited@0
-      array 2
-        type-con Vertex 0
-        bool
-    var y
-      type-con Vertex 0
-    <
-    select 3
-    fun G 0
-    var y
-      type-con Vertex 0
-    var z
-      type-con Vertex 0
-    fun Infinity 0
-    <=
-    select 2
-    var SP@0
-      array 2
-        type-con Vertex 0
-        int
-    var z
-      type-con Vertex 0
-    +
-    select 2
-    var SP@0
-      array 2
-        type-con Vertex 0
-        int
-    var y
-      type-con Vertex 0
-    select 3
-    fun G 0
-    var y
-      type-con Vertex 0
-    var z
-      type-con Vertex 0
-    implies
-    forall 2 0 3
-      var z
-        type-con Vertex 0
-      var y
-        type-con Vertex 0
-      attribute qid 1
-        string-attr BoogieDi.38:23
-      attribute uniqueId 1
-        string-attr 11
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 2
-    select 2
-    var Visited@0
-      array 2
-        type-con Vertex 0
-        bool
-    var y
-      type-con Vertex 0
-    <
-    select 3
-    fun G 0
-    var y
-      type-con Vertex 0
-    var z
-      type-con Vertex 0
-    fun Infinity 0
-    <=
-    select 2
-    var SP@0
-      array 2
-        type-con Vertex 0
-        int
-    var z
-      type-con Vertex 0
-    +
-    select 2
-    var SP@0
-      array 2
-        type-con Vertex 0
-        int
-    var y
-      type-con Vertex 0
-    select 3
-    fun G 0
-    var y
-      type-con Vertex 0
-    var z
-      type-con Vertex 0
-    and 2
-    label neg 40 5
-    forall 1 0 3
-      var z
-        type-con Vertex 0
-      attribute qid 1
-        string-attr BoogieDi.40:23
-      attribute uniqueId 1
-        string-attr 13
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 2
-    not
-    =
-    var z
-      type-con Vertex 0
-    fun Source 0
-    <
-    select 2
-    var SP@0
-      array 2
-        type-con Vertex 0
-        int
-    var z
-      type-con Vertex 0
-    fun Infinity 0
-    exists 1 0 3
-      var y
-        type-con Vertex 0
-      attribute qid 1
-        string-attr BoogieDi.41:15
-      attribute uniqueId 1
-        string-attr 12
-      attribute bvZ3Native 1
-        string-attr False
-    and 3
-    <
-    select 2
-    var SP@0
-      array 2
-        type-con Vertex 0
-        int
-    var y
-      type-con Vertex 0
-    select 2
-    var SP@0
-      array 2
-        type-con Vertex 0
-        int
-    var z
-      type-con Vertex 0
-    select 2
-    var Visited@0
-      array 2
-        type-con Vertex 0
-        bool
-    var y
-      type-con Vertex 0
-    =
-    select 2
-    var SP@0
-      array 2
-        type-con Vertex 0
-        int
-    var z
-      type-con Vertex 0
-    +
-    select 2
-    var SP@0
-      array 2
-        type-con Vertex 0
-        int
-    var y
-      type-con Vertex 0
-    select 3
-    fun G 0
-    var y
-      type-con Vertex 0
-    var z
-      type-con Vertex 0
-    implies
-    forall 1 0 3
-      var z
-        type-con Vertex 0
-      attribute qid 1
-        string-attr BoogieDi.40:23
-      attribute uniqueId 1
-        string-attr 13
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 2
-    not
-    =
-    var z
-      type-con Vertex 0
-    fun Source 0
-    <
-    select 2
-    var SP@0
-      array 2
-        type-con Vertex 0
-        int
-    var z
-      type-con Vertex 0
-    fun Infinity 0
-    exists 1 0 3
-      var y
-        type-con Vertex 0
-      attribute qid 1
-        string-attr BoogieDi.41:15
-      attribute uniqueId 1
-        string-attr 12
-      attribute bvZ3Native 1
-        string-attr False
-    and 3
-    <
-    select 2
-    var SP@0
-      array 2
-        type-con Vertex 0
-        int
-    var y
-      type-con Vertex 0
-    select 2
-    var SP@0
-      array 2
-        type-con Vertex 0
-        int
-    var z
-      type-con Vertex 0
-    select 2
-    var Visited@0
-      array 2
-        type-con Vertex 0
-        bool
-    var y
-      type-con Vertex 0
-    =
-    select 2
-    var SP@0
-      array 2
-        type-con Vertex 0
-        int
-    var z
-      type-con Vertex 0
-    +
-    select 2
-    var SP@0
-      array 2
-        type-con Vertex 0
-        int
-    var y
-      type-con Vertex 0
-    select 3
-    fun G 0
-    var y
-      type-con Vertex 0
-    var z
-      type-con Vertex 0
-    implies
-    label pos 33 3
-    true
-    implies
-    true
-    implies
-    =
-    select 2
-    var SP@1
-      array 2
-        type-con Vertex 0
-        int
-    fun Source 0
-    int-num 0
-    implies
-    forall 1 0 3
-      var x
-        type-con Vertex 0
-      attribute qid 1
-        string-attr BoogieDi.35:23
-      attribute uniqueId 1
-        string-attr 9
-      attribute bvZ3Native 1
-        string-attr False
-    >=
-    select 2
-    var SP@1
-      array 2
-        type-con Vertex 0
-        int
-    var x
-      type-con Vertex 0
-    int-num 0
-    implies
-    forall 2 0 3
-      var y
-        type-con Vertex 0
-      var z
-        type-con Vertex 0
-      attribute qid 1
-        string-attr BoogieDi.36:23
-      attribute uniqueId 1
-        string-attr 10
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 2
-    not
-    select 2
-    var Visited@1
-      array 2
-        type-con Vertex 0
-        bool
-    var z
-      type-con Vertex 0
-    select 2
-    var Visited@1
-      array 2
-        type-con Vertex 0
-        bool
-    var y
-      type-con Vertex 0
-    <=
-    select 2
-    var SP@1
-      array 2
-        type-con Vertex 0
-        int
-    var y
-      type-con Vertex 0
-    select 2
-    var SP@1
-      array 2
-        type-con Vertex 0
-        int
-    var z
-      type-con Vertex 0
-    implies
-    forall 2 0 3
-      var z
-        type-con Vertex 0
-      var y
-        type-con Vertex 0
-      attribute qid 1
-        string-attr BoogieDi.38:23
-      attribute uniqueId 1
-        string-attr 11
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 2
-    select 2
-    var Visited@1
-      array 2
-        type-con Vertex 0
-        bool
-    var y
-      type-con Vertex 0
-    <
-    select 3
-    fun G 0
-    var y
-      type-con Vertex 0
-    var z
-      type-con Vertex 0
-    fun Infinity 0
-    <=
-    select 2
-    var SP@1
-      array 2
-        type-con Vertex 0
-        int
-    var z
-      type-con Vertex 0
-    +
-    select 2
-    var SP@1
-      array 2
-        type-con Vertex 0
-        int
-    var y
-      type-con Vertex 0
-    select 3
-    fun G 0
-    var y
-      type-con Vertex 0
-    var z
-      type-con Vertex 0
-    implies
-    forall 1 0 3
-      var z
-        type-con Vertex 0
-      attribute qid 1
-        string-attr BoogieDi.40:23
-      attribute uniqueId 1
-        string-attr 13
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 2
-    not
-    =
-    var z
-      type-con Vertex 0
-    fun Source 0
-    <
-    select 2
-    var SP@1
-      array 2
-        type-con Vertex 0
-        int
-    var z
-      type-con Vertex 0
-    fun Infinity 0
-    exists 1 0 3
-      var y
-        type-con Vertex 0
-      attribute qid 1
-        string-attr BoogieDi.41:15
-      attribute uniqueId 1
-        string-attr 12
-      attribute bvZ3Native 1
-        string-attr False
-    and 3
-    <
-    select 2
-    var SP@1
-      array 2
-        type-con Vertex 0
-        int
-    var y
-      type-con Vertex 0
-    select 2
-    var SP@1
-      array 2
-        type-con Vertex 0
-        int
-    var z
-      type-con Vertex 0
-    select 2
-    var Visited@1
-      array 2
-        type-con Vertex 0
-        bool
-    var y
-      type-con Vertex 0
-    =
-    select 2
-    var SP@1
-      array 2
-        type-con Vertex 0
-        int
-    var z
-      type-con Vertex 0
-    +
-    select 2
-    var SP@1
-      array 2
-        type-con Vertex 0
-        int
-    var y
-      type-con Vertex 0
-    select 3
-    fun G 0
-    var y
-      type-con Vertex 0
-    var z
-      type-con Vertex 0
-    implies
-    true
-    and 2
-    implies
-    label pos 33 3
-    true
-    implies
-    true
-    implies
-    not
-    exists 1 0 3
-      var x
-        type-con Vertex 0
-      attribute qid 1
-        string-attr BoogieDi.33:18
-      attribute uniqueId 1
-        string-attr 8
-      attribute bvZ3Native 1
-        string-attr False
-    and 2
-    not
-    select 2
-    var Visited@1
-      array 2
-        type-con Vertex 0
-        bool
-    var x
-      type-con Vertex 0
-    <
-    select 2
-    var SP@1
-      array 2
-        type-con Vertex 0
-        int
-    var x
-      type-con Vertex 0
-    fun Infinity 0
-    implies
-    true
-    implies
-    label pos 0 0
-    true
-    implies
-    =
-    var Visited@3
-      array 2
-        type-con Vertex 0
-        bool
-    var Visited@1
-      array 2
-        type-con Vertex 0
-        bool
-    implies
-    =
-    var v@2
-      type-con Vertex 0
-    var v@0
-      type-con Vertex 0
-    implies
-    =
-    var SP@3
-      array 2
-        type-con Vertex 0
-        int
-    var SP@1
-      array 2
-        type-con Vertex 0
-        int
-    implies
-    =
-    var oldSP@1
-      array 2
-        type-con Vertex 0
-        int
-    var oldSP@0
-      array 2
-        type-con Vertex 0
-        int
-    implies
-    label pos 0 0
-    true
-    and 2
-    label neg 17 3
-    forall 1 0 3
-      var z
-        type-con Vertex 0
-      attribute qid 1
-        string-attr BoogieDi.17:19
-      attribute uniqueId 1
-        string-attr 4
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 2
-    not
-    =
-    var z
-      type-con Vertex 0
-    fun Source 0
-    <
-    select 2
-    var SP@3
-      array 2
-        type-con Vertex 0
-        int
-    var z
-      type-con Vertex 0
-    fun Infinity 0
-    exists 1 0 3
-      var y
-        type-con Vertex 0
-      attribute qid 1
-        string-attr BoogieDi.18:13
-      attribute uniqueId 1
-        string-attr 3
-      attribute bvZ3Native 1
-        string-attr False
-    and 2
-    <
-    select 2
-    var SP@3
-      array 2
-        type-con Vertex 0
-        int
-    var y
-      type-con Vertex 0
-    select 2
-    var SP@3
-      array 2
-        type-con Vertex 0
-        int
-    var z
-      type-con Vertex 0
-    =
-    select 2
-    var SP@3
-      array 2
-        type-con Vertex 0
-        int
-    var z
-      type-con Vertex 0
-    +
-    select 2
-    var SP@3
-      array 2
-        type-con Vertex 0
-        int
-    var y
-      type-con Vertex 0
-    select 3
-    fun G 0
-    var y
-      type-con Vertex 0
-    var z
-      type-con Vertex 0
-    implies
-    forall 1 0 3
-      var z
-        type-con Vertex 0
-      attribute qid 1
-        string-attr BoogieDi.17:19
-      attribute uniqueId 1
-        string-attr 4
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 2
-    not
-    =
-    var z
-      type-con Vertex 0
-    fun Source 0
-    <
-    select 2
-    var SP@3
-      array 2
-        type-con Vertex 0
-        int
-    var z
-      type-con Vertex 0
-    fun Infinity 0
-    exists 1 0 3
-      var y
-        type-con Vertex 0
-      attribute qid 1
-        string-attr BoogieDi.18:13
-      attribute uniqueId 1
-        string-attr 3
-      attribute bvZ3Native 1
-        string-attr False
-    and 2
-    <
-    select 2
-    var SP@3
-      array 2
-        type-con Vertex 0
-        int
-    var y
-      type-con Vertex 0
-    select 2
-    var SP@3
-      array 2
-        type-con Vertex 0
-        int
-    var z
-      type-con Vertex 0
-    =
-    select 2
-    var SP@3
-      array 2
-        type-con Vertex 0
-        int
-    var z
-      type-con Vertex 0
-    +
-    select 2
-    var SP@3
-      array 2
-        type-con Vertex 0
-        int
-    var y
-      type-con Vertex 0
-    select 3
-    fun G 0
-    var y
-      type-con Vertex 0
-    var z
-      type-con Vertex 0
-    and 2
-    label neg 15 3
-    forall 2 0 3
-      var z
-        type-con Vertex 0
-      var y
-        type-con Vertex 0
-      attribute qid 1
-        string-attr BoogieDi.15:19
-      attribute uniqueId 1
-        string-attr 2
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 2
-    <
-    select 2
-    var SP@3
-      array 2
-        type-con Vertex 0
-        int
-    var y
-      type-con Vertex 0
-    fun Infinity 0
-    <
-    select 3
-    fun G 0
-    var y
-      type-con Vertex 0
-    var z
-      type-con Vertex 0
-    fun Infinity 0
-    <=
-    select 2
-    var SP@3
-      array 2
-        type-con Vertex 0
-        int
-    var z
-      type-con Vertex 0
-    +
-    select 2
-    var SP@3
-      array 2
-        type-con Vertex 0
-        int
-    var y
-      type-con Vertex 0
-    select 3
-    fun G 0
-    var y
-      type-con Vertex 0
-    var z
-      type-con Vertex 0
-    implies
-    forall 2 0 3
-      var z
-        type-con Vertex 0
-      var y
-        type-con Vertex 0
-      attribute qid 1
-        string-attr BoogieDi.15:19
-      attribute uniqueId 1
-        string-attr 2
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 2
-    <
-    select 2
-    var SP@3
-      array 2
-        type-con Vertex 0
-        int
-    var y
-      type-con Vertex 0
-    fun Infinity 0
-    <
-    select 3
-    fun G 0
-    var y
-      type-con Vertex 0
-    var z
-      type-con Vertex 0
-    fun Infinity 0
-    <=
-    select 2
-    var SP@3
-      array 2
-        type-con Vertex 0
-        int
-    var z
-      type-con Vertex 0
-    +
-    select 2
-    var SP@3
-      array 2
-        type-con Vertex 0
-        int
-    var y
-      type-con Vertex 0
-    select 3
-    fun G 0
-    var y
-      type-con Vertex 0
-    var z
-      type-con Vertex 0
-    and 2
-    label neg 14 3
-    =
-    select 2
-    var SP@3
-      array 2
-        type-con Vertex 0
-        int
-    fun Source 0
-    int-num 0
-    implies
-    =
-    select 2
-    var SP@3
-      array 2
-        type-con Vertex 0
-        int
-    fun Source 0
-    int-num 0
-    true
-    implies
-    label pos 44 5
-    true
-    implies
-    true
-    implies
-    exists 1 0 3
-      var x
-        type-con Vertex 0
-      attribute qid 1
-        string-attr BoogieDi.33:18
-      attribute uniqueId 1
-        string-attr 8
-      attribute bvZ3Native 1
-        string-attr False
-    and 2
-    not
-    select 2
-    var Visited@1
-      array 2
-        type-con Vertex 0
-        bool
-    var x
-      type-con Vertex 0
-    <
-    select 2
-    var SP@1
-      array 2
-        type-con Vertex 0
-        int
-    var x
-      type-con Vertex 0
-    fun Infinity 0
-    implies
-    not
-    select 2
-    var Visited@1
-      array 2
-        type-con Vertex 0
-        bool
-    var v@1
-      type-con Vertex 0
-    implies
-    <
-    select 2
-    var SP@1
-      array 2
-        type-con Vertex 0
-        int
-    var v@1
-      type-con Vertex 0
-    fun Infinity 0
-    implies
-    forall 1 0 3
-      var x
-        type-con Vertex 0
-      attribute qid 1
-        string-attr BoogieDi.47:20
-      attribute uniqueId 1
-        string-attr 14
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    not
-    select 2
-    var Visited@1
-      array 2
-        type-con Vertex 0
-        bool
-    var x
-      type-con Vertex 0
-    <=
-    select 2
-    var SP@1
-      array 2
-        type-con Vertex 0
-        int
-    var v@1
-      type-con Vertex 0
-    select 2
-    var SP@1
-      array 2
-        type-con Vertex 0
-        int
-    var x
-      type-con Vertex 0
-    implies
-    =
-    var Visited@2
-      array 2
-        type-con Vertex 0
-        bool
-    store 3
-    var Visited@1
-      array 2
-        type-con Vertex 0
-        bool
-    var v@1
-      type-con Vertex 0
-    true
-    implies
-    forall 1 0 3
-      var u
-        type-con Vertex 0
-      attribute qid 1
-        string-attr BoogieDi.53:20
-      attribute uniqueId 1
-        string-attr 15
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 2
-    <
-    select 3
-    fun G 0
-    var v@1
-      type-con Vertex 0
-    var u
-      type-con Vertex 0
-    fun Infinity 0
-    <
-    +
-    select 2
-    var SP@1
-      array 2
-        type-con Vertex 0
-        int
-    var v@1
-      type-con Vertex 0
-    select 3
-    fun G 0
-    var v@1
-      type-con Vertex 0
-    var u
-      type-con Vertex 0
-    select 2
-    var SP@1
-      array 2
-        type-con Vertex 0
-        int
-    var u
-      type-con Vertex 0
-    =
-    select 2
-    var SP@2
-      array 2
-        type-con Vertex 0
-        int
-    var u
-      type-con Vertex 0
-    +
-    select 2
-    var SP@1
-      array 2
-        type-con Vertex 0
-        int
-    var v@1
-      type-con Vertex 0
-    select 3
-    fun G 0
-    var v@1
-      type-con Vertex 0
-    var u
-      type-con Vertex 0
-    implies
-    forall 1 0 3
-      var u
-        type-con Vertex 0
-      attribute qid 1
-        string-attr BoogieDi.56:20
-      attribute uniqueId 1
-        string-attr 16
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    not
-    and 2
-    <
-    select 3
-    fun G 0
-    var v@1
-      type-con Vertex 0
-    var u
-      type-con Vertex 0
-    fun Infinity 0
-    <
-    +
-    select 2
-    var SP@1
-      array 2
-        type-con Vertex 0
-        int
-    var v@1
-      type-con Vertex 0
-    select 3
-    fun G 0
-    var v@1
-      type-con Vertex 0
-    var u
-      type-con Vertex 0
-    select 2
-    var SP@1
-      array 2
-        type-con Vertex 0
-        int
-    var u
-      type-con Vertex 0
-    =
-    select 2
-    var SP@2
-      array 2
-        type-con Vertex 0
-        int
-    var u
-      type-con Vertex 0
-    select 2
-    var SP@1
-      array 2
-        type-con Vertex 0
-        int
-    var u
-      type-con Vertex 0
-    and 2
-    label neg 59 5
-    forall 1 0 3
-      var z
-        type-con Vertex 0
-      attribute qid 1
-        string-attr BoogieDi.59:20
-      attribute uniqueId 1
-        string-attr 17
-      attribute bvZ3Native 1
-        string-attr False
-    <=
-    select 2
-    var SP@2
-      array 2
-        type-con Vertex 0
-        int
-    var z
-      type-con Vertex 0
-    select 2
-    var SP@1
-      array 2
-        type-con Vertex 0
-        int
-    var z
-      type-con Vertex 0
-    implies
-    forall 1 0 3
-      var z
-        type-con Vertex 0
-      attribute qid 1
-        string-attr BoogieDi.59:20
-      attribute uniqueId 1
-        string-attr 17
-      attribute bvZ3Native 1
-        string-attr False
-    <=
-    select 2
-    var SP@2
-      array 2
-        type-con Vertex 0
-        int
-    var z
-      type-con Vertex 0
-    select 2
-    var SP@1
-      array 2
-        type-con Vertex 0
-        int
-    var z
-      type-con Vertex 0
-    and 2
-    label neg 60 5
-    forall 1 0 3
-      var y
-        type-con Vertex 0
-      attribute qid 1
-        string-attr BoogieDi.60:20
-      attribute uniqueId 1
-        string-attr 18
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    select 2
-    var Visited@2
-      array 2
-        type-con Vertex 0
-        bool
-    var y
-      type-con Vertex 0
-    =
-    select 2
-    var SP@2
-      array 2
-        type-con Vertex 0
-        int
-    var y
-      type-con Vertex 0
-    select 2
-    var SP@1
-      array 2
-        type-con Vertex 0
-        int
-    var y
-      type-con Vertex 0
-    implies
-    forall 1 0 3
-      var y
-        type-con Vertex 0
-      attribute qid 1
-        string-attr BoogieDi.60:20
-      attribute uniqueId 1
-        string-attr 18
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    select 2
-    var Visited@2
-      array 2
-        type-con Vertex 0
-        bool
-    var y
-      type-con Vertex 0
-    =
-    select 2
-    var SP@2
-      array 2
-        type-con Vertex 0
-        int
-    var y
-      type-con Vertex 0
-    select 2
-    var SP@1
-      array 2
-        type-con Vertex 0
-        int
-    var y
-      type-con Vertex 0
-    implies
-    true
-    implies
-    label pos 0 0
-    true
-    and 2
-    label neg 34 5
-    =
-    select 2
-    var SP@2
-      array 2
-        type-con Vertex 0
-        int
-    fun Source 0
-    int-num 0
-    implies
-    =
-    select 2
-    var SP@2
-      array 2
-        type-con Vertex 0
-        int
-    fun Source 0
-    int-num 0
-    and 2
-    label neg 35 5
-    forall 1 0 3
-      var x
-        type-con Vertex 0
-      attribute qid 1
-        string-attr BoogieDi.35:23
-      attribute uniqueId 1
-        string-attr 9
-      attribute bvZ3Native 1
-        string-attr False
-    >=
-    select 2
-    var SP@2
-      array 2
-        type-con Vertex 0
-        int
-    var x
-      type-con Vertex 0
-    int-num 0
-    implies
-    forall 1 0 3
-      var x
-        type-con Vertex 0
-      attribute qid 1
-        string-attr BoogieDi.35:23
-      attribute uniqueId 1
-        string-attr 9
-      attribute bvZ3Native 1
-        string-attr False
-    >=
-    select 2
-    var SP@2
-      array 2
-        type-con Vertex 0
-        int
-    var x
-      type-con Vertex 0
-    int-num 0
-    and 2
-    label neg 36 5
-    forall 2 0 3
-      var y
-        type-con Vertex 0
-      var z
-        type-con Vertex 0
-      attribute qid 1
-        string-attr BoogieDi.36:23
-      attribute uniqueId 1
-        string-attr 10
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 2
-    not
-    select 2
-    var Visited@2
-      array 2
-        type-con Vertex 0
-        bool
-    var z
-      type-con Vertex 0
-    select 2
-    var Visited@2
-      array 2
-        type-con Vertex 0
-        bool
-    var y
-      type-con Vertex 0
-    <=
-    select 2
-    var SP@2
-      array 2
-        type-con Vertex 0
-        int
-    var y
-      type-con Vertex 0
-    select 2
-    var SP@2
-      array 2
-        type-con Vertex 0
-        int
-    var z
-      type-con Vertex 0
-    implies
-    forall 2 0 3
-      var y
-        type-con Vertex 0
-      var z
-        type-con Vertex 0
-      attribute qid 1
-        string-attr BoogieDi.36:23
-      attribute uniqueId 1
-        string-attr 10
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 2
-    not
-    select 2
-    var Visited@2
-      array 2
-        type-con Vertex 0
-        bool
-    var z
-      type-con Vertex 0
-    select 2
-    var Visited@2
-      array 2
-        type-con Vertex 0
-        bool
-    var y
-      type-con Vertex 0
-    <=
-    select 2
-    var SP@2
-      array 2
-        type-con Vertex 0
-        int
-    var y
-      type-con Vertex 0
-    select 2
-    var SP@2
-      array 2
-        type-con Vertex 0
-        int
-    var z
-      type-con Vertex 0
-    and 2
-    label neg 38 5
-    forall 2 0 3
-      var z
-        type-con Vertex 0
-      var y
-        type-con Vertex 0
-      attribute qid 1
-        string-attr BoogieDi.38:23
-      attribute uniqueId 1
-        string-attr 11
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 2
-    select 2
-    var Visited@2
-      array 2
-        type-con Vertex 0
-        bool
-    var y
-      type-con Vertex 0
-    <
-    select 3
-    fun G 0
-    var y
-      type-con Vertex 0
-    var z
-      type-con Vertex 0
-    fun Infinity 0
-    <=
-    select 2
-    var SP@2
-      array 2
-        type-con Vertex 0
-        int
-    var z
-      type-con Vertex 0
-    +
-    select 2
-    var SP@2
-      array 2
-        type-con Vertex 0
-        int
-    var y
-      type-con Vertex 0
-    select 3
-    fun G 0
-    var y
-      type-con Vertex 0
-    var z
-      type-con Vertex 0
-    implies
-    forall 2 0 3
-      var z
-        type-con Vertex 0
-      var y
-        type-con Vertex 0
-      attribute qid 1
-        string-attr BoogieDi.38:23
-      attribute uniqueId 1
-        string-attr 11
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 2
-    select 2
-    var Visited@2
-      array 2
-        type-con Vertex 0
-        bool
-    var y
-      type-con Vertex 0
-    <
-    select 3
-    fun G 0
-    var y
-      type-con Vertex 0
-    var z
-      type-con Vertex 0
-    fun Infinity 0
-    <=
-    select 2
-    var SP@2
-      array 2
-        type-con Vertex 0
-        int
-    var z
-      type-con Vertex 0
-    +
-    select 2
-    var SP@2
-      array 2
-        type-con Vertex 0
-        int
-    var y
-      type-con Vertex 0
-    select 3
-    fun G 0
-    var y
-      type-con Vertex 0
-    var z
-      type-con Vertex 0
-    and 2
-    label neg 40 5
-    forall 1 0 3
-      var z
-        type-con Vertex 0
-      attribute qid 1
-        string-attr BoogieDi.40:23
-      attribute uniqueId 1
-        string-attr 13
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 2
-    not
-    =
-    var z
-      type-con Vertex 0
-    fun Source 0
-    <
-    select 2
-    var SP@2
-      array 2
-        type-con Vertex 0
-        int
-    var z
-      type-con Vertex 0
-    fun Infinity 0
-    exists 1 0 3
-      var y
-        type-con Vertex 0
-      attribute qid 1
-        string-attr BoogieDi.41:15
-      attribute uniqueId 1
-        string-attr 12
-      attribute bvZ3Native 1
-        string-attr False
-    and 3
-    <
-    select 2
-    var SP@2
-      array 2
-        type-con Vertex 0
-        int
-    var y
-      type-con Vertex 0
-    select 2
-    var SP@2
-      array 2
-        type-con Vertex 0
-        int
-    var z
-      type-con Vertex 0
-    select 2
-    var Visited@2
-      array 2
-        type-con Vertex 0
-        bool
-    var y
-      type-con Vertex 0
-    =
-    select 2
-    var SP@2
-      array 2
-        type-con Vertex 0
-        int
-    var z
-      type-con Vertex 0
-    +
-    select 2
-    var SP@2
-      array 2
-        type-con Vertex 0
-        int
-    var y
-      type-con Vertex 0
-    select 3
-    fun G 0
-    var y
-      type-con Vertex 0
-    var z
-      type-con Vertex 0
-    implies
-    forall 1 0 3
-      var z
-        type-con Vertex 0
-      attribute qid 1
-        string-attr BoogieDi.40:23
-      attribute uniqueId 1
-        string-attr 13
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 2
-    not
-    =
-    var z
-      type-con Vertex 0
-    fun Source 0
-    <
-    select 2
-    var SP@2
-      array 2
-        type-con Vertex 0
-        int
-    var z
-      type-con Vertex 0
-    fun Infinity 0
-    exists 1 0 3
-      var y
-        type-con Vertex 0
-      attribute qid 1
-        string-attr BoogieDi.41:15
-      attribute uniqueId 1
-        string-attr 12
-      attribute bvZ3Native 1
-        string-attr False
-    and 3
-    <
-    select 2
-    var SP@2
-      array 2
-        type-con Vertex 0
-        int
-    var y
-      type-con Vertex 0
-    select 2
-    var SP@2
-      array 2
-        type-con Vertex 0
-        int
-    var z
-      type-con Vertex 0
-    select 2
-    var Visited@2
-      array 2
-        type-con Vertex 0
-        bool
-    var y
-      type-con Vertex 0
-    =
-    select 2
-    var SP@2
-      array 2
-        type-con Vertex 0
-        int
-    var z
-      type-con Vertex 0
-    +
-    select 2
-    var SP@2
-      array 2
-        type-con Vertex 0
-        int
-    var y
-      type-con Vertex 0
-    select 3
-    fun G 0
-    var y
-      type-con Vertex 0
-    var z
-      type-con Vertex 0
-    implies
-    false
-    true
--- a/src/HOL/Boogie/Examples/Boogie_Dijkstra.certs	Tue Jul 23 13:14:14 2013 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,6048 +0,0 @@
-a549accb56a0340bea4e867d66fdfbbac187f40c 6047 0
-#2 := false
-#7 := 0::Int
-decl f30 :: (-> S2 Int)
-decl ?v1!7 :: (-> S2 S2)
-decl ?v0!20 :: S2
-#2034 := ?v0!20
-#5342 := (?v1!7 ?v0!20)
-#13325 := (f30 #5342)
-#923 := -1::Int
-#13327 := (* -1::Int #13325)
-decl f4 :: (-> S3 Int)
-decl f5 :: (-> S4 S2 S3)
-decl f6 :: (-> S5 S2 S4)
-decl f7 :: S5
-#13 := f7
-#5350 := (f6 f7 #5342)
-#5351 := (f5 #5350 ?v0!20)
-#5352 := (f4 #5351)
-#5353 := (* -1::Int #5352)
-#14209 := (+ #5353 #13327)
-#2037 := (f30 ?v0!20)
-#14210 := (+ #2037 #14209)
-#18243 := (>= #14210 0::Int)
-decl f19 :: (-> S11 S2 Int)
-decl f20 :: S11
-#109 := f20
-#5343 := (f19 f20 #5342)
-#5344 := (* -1::Int #5343)
-#5354 := (+ #5344 #5353)
-#5080 := (f19 f20 ?v0!20)
-#5355 := (+ #5080 #5354)
-#15530 := (>= #5355 0::Int)
-#5356 := (= #5355 0::Int)
-#5357 := (not #5356)
-decl f1 :: S1
-#3 := f1
-decl f11 :: (-> S7 S2 S1)
-decl f21 :: S7
-#115 := f21
-#5347 := (f11 f21 #5342)
-#5348 := (= #5347 f1)
-#5349 := (not #5348)
-#5345 := (+ #5080 #5344)
-#5346 := (<= #5345 0::Int)
-#5358 := (or #5346 #5349 #5357)
-#5359 := (not #5358)
-#5105 := (* -1::Int #5080)
-decl f3 :: Int
-#8 := f3
-#5340 := (+ f3 #5105)
-#5341 := (<= #5340 0::Int)
-#15968 := (not #5341)
-#5106 := (+ #2037 #5105)
-#12022 := (>= #5106 0::Int)
-#5087 := (= #2037 #5080)
-decl f28 :: S2
-#181 := f28
-#195 := (f6 f7 f28)
-#5121 := (f5 #195 ?v0!20)
-#5122 := (f4 #5121)
-#5139 := (+ #5105 #5122)
-#185 := (f19 f20 f28)
-#5140 := (+ #185 #5139)
-#5141 := (>= #5140 0::Int)
-#5123 := (* -1::Int #5122)
-#5124 := (+ f3 #5123)
-#5125 := (<= #5124 0::Int)
-#5146 := (or #5125 #5141)
-#2038 := (* -1::Int #2037)
-#5178 := (+ #2038 #5122)
-#5179 := (+ #185 #5178)
-#5182 := (= #5179 0::Int)
-#15374 := (not #5182)
-#12121 := (<= #5179 0::Int)
-#15379 := (not #12121)
-#4078 := (f30 f28)
-#4079 := (* -1::Int #4078)
-#5068 := (+ #2037 #4079)
-#5229 := (<= #5068 0::Int)
-#15402 := (not #5229)
-#15340 := (<= #5122 0::Int)
-#15358 := (not #15340)
-#15327 := (= f28 ?v0!20)
-#15368 := (not #15327)
-#9945 := (= ?v0!20 f28)
-#16393 := [hypothesis]: #15327
-#16394 := [symm #16393]: #9945
-#9687 := (not #9945)
-#5309 := (f11 f21 ?v0!20)
-#5310 := (= #5309 f1)
-#9951 := (or #9945 #5310)
-#9685 := (not #9951)
-decl f12 :: (-> S8 S1 S7)
-decl f13 :: (-> S9 S2 S8)
-decl f14 :: (-> S10 S7 S9)
-decl f15 :: S10
-#39 := f15
-#191 := (f14 f15 f21)
-#192 := (f13 #191 f28)
-#193 := (f12 #192 f1)
-#9943 := (f11 #193 ?v0!20)
-#9944 := (= #9943 f1)
-#9956 := (iff #9944 #9951)
-#11 := (:var 0 S2)
-#54 := (:var 1 S1)
-#52 := (:var 2 S2)
-#50 := (:var 3 S7)
-#51 := (f14 f15 #50)
-#53 := (f13 #51 #52)
-#55 := (f12 #53 #54)
-#56 := (f11 #55 #11)
-#3640 := (pattern #56)
-#60 := (f11 #50 #11)
-#61 := (= #60 f1)
-#59 := (= #54 f1)
-#58 := (= #11 #52)
-#62 := (if #58 #59 #61)
-#57 := (= #56 f1)
-#63 := (iff #57 #62)
-#3641 := (forall (vars (?v0 S7) (?v1 S2) (?v2 S1) (?v3 S2)) (:pat #3640) #63)
-#64 := (forall (vars (?v0 S7) (?v1 S2) (?v2 S1) (?v3 S2)) #63)
-#3644 := (iff #64 #3641)
-#3642 := (iff #63 #63)
-#3643 := [refl]: #3642
-#3645 := [quant-intro #3643]: #3644
-#1496 := (~ #64 #64)
-#1524 := (~ #63 #63)
-#1525 := [refl]: #1524
-#1497 := [nnf-pos #1525]: #1496
-#342 := [asserted]: #64
-#1526 := [mp~ #342 #1497]: #64
-#3646 := [mp #1526 #3645]: #3641
-#7518 := (not #3641)
-#9763 := (or #7518 #9956)
-#4057 := (= f1 f1)
-#9946 := (if #9945 #4057 #5310)
-#9947 := (iff #9944 #9946)
-#9775 := (or #7518 #9947)
-#9715 := (iff #9775 #9763)
-#9722 := (iff #9763 #9763)
-#9724 := [rewrite]: #9722
-#9957 := (iff #9947 #9956)
-#9954 := (iff #9946 #9951)
-#1 := true
-#9948 := (if #9945 true #5310)
-#9952 := (iff #9948 #9951)
-#9953 := [rewrite]: #9952
-#9949 := (iff #9946 #9948)
-#4059 := (iff #4057 true)
-#4060 := [rewrite]: #4059
-#9950 := [monotonicity #4060]: #9949
-#9955 := [trans #9950 #9953]: #9954
-#9958 := [monotonicity #9955]: #9957
-#9706 := [monotonicity #9958]: #9715
-#9734 := [trans #9706 #9724]: #9715
-#9776 := [quant-inst #115 #181 #3 #2034]: #9775
-#9519 := [mp #9776 #9734]: #9763
-#16374 := [unit-resolution #9519 #3646]: #9956
-#9410 := (not #9944)
-decl f29 :: S7
-#190 := f29
-#4533 := (f11 f29 ?v0!20)
-#4534 := (= #4533 f1)
-#4541 := (not #4534)
-#16390 := (iff #4541 #9410)
-#16382 := (iff #4534 #9944)
-#16359 := (iff #9944 #4534)
-#16289 := (= #9943 #4533)
-#8145 := (= #193 f29)
-#194 := (= f29 #193)
-#91 := (f6 f7 #11)
-#3693 := (pattern #91)
-#212 := (f11 f29 #11)
-#3854 := (pattern #212)
-#202 := (f30 #11)
-#3829 := (pattern #202)
-#2047 := (f5 #91 ?v0!20)
-#2048 := (f4 #2047)
-#2383 := (+ #2038 #2048)
-#2384 := (+ #202 #2383)
-#2387 := (= #2384 0::Int)
-#2941 := (not #2387)
-#213 := (= #212 f1)
-#220 := (not #213)
-#2044 := (+ #202 #2038)
-#2045 := (>= #2044 0::Int)
-#2942 := (or #2045 #220 #2941)
-#3888 := (forall (vars (?v1 S2)) (:pat #3829 #3854 #3693) #2942)
-#3893 := (not #3888)
-#2039 := (+ f3 #2038)
-#2040 := (<= #2039 0::Int)
-decl f16 :: S2
-#65 := f16
-#2035 := (= ?v0!20 f16)
-#10 := (:var 1 S2)
-#92 := (f5 #91 #10)
-#3684 := (pattern #92)
-#224 := (f30 #10)
-#1186 := (* -1::Int #224)
-#1187 := (+ #202 #1186)
-#93 := (f4 #92)
-#1207 := (+ #93 #1187)
-#1205 := (>= #1207 0::Int)
-#938 := (* -1::Int #93)
-#939 := (+ f3 #938)
-#940 := (<= #939 0::Int)
-#2933 := (or #220 #940 #1205)
-#3880 := (forall (vars (?v0 S2) (?v1 S2)) (:pat #3684) #2933)
-#3885 := (not #3880)
-#3896 := (or #3885 #2035 #2040 #3893)
-#3899 := (not #3896)
-decl ?v0!19 :: S2
-#2003 := ?v0!19
-#2016 := (f30 ?v0!19)
-#2017 := (* -1::Int #2016)
-decl ?v1!18 :: S2
-#2002 := ?v1!18
-#2015 := (f30 ?v1!18)
-#2018 := (+ #2015 #2017)
-#2006 := (f6 f7 ?v1!18)
-#2007 := (f5 #2006 ?v0!19)
-#2008 := (f4 #2007)
-#2019 := (+ #2008 #2018)
-#2020 := (>= #2019 0::Int)
-#2009 := (* -1::Int #2008)
-#2010 := (+ f3 #2009)
-#2011 := (<= #2010 0::Int)
-#2004 := (f11 f29 ?v1!18)
-#2005 := (= #2004 f1)
-#2896 := (not #2005)
-#2911 := (or #2896 #2011 #2020)
-#2916 := (not #2911)
-#3902 := (or #2916 #3899)
-#3905 := (not #3902)
-#3871 := (pattern #202 #224)
-#1185 := (>= #1187 0::Int)
-#221 := (f11 f29 #10)
-#222 := (= #221 f1)
-#2873 := (not #222)
-#2888 := (or #213 #2873 #1185)
-#3872 := (forall (vars (?v0 S2) (?v1 S2)) (:pat #3871) #2888)
-#3877 := (not #3872)
-#3908 := (or #3877 #3905)
-#3911 := (not #3908)
-decl ?v0!17 :: S2
-#1976 := ?v0!17
-#1985 := (f30 ?v0!17)
-#1986 := (* -1::Int #1985)
-decl ?v1!16 :: S2
-#1975 := ?v1!16
-#1984 := (f30 ?v1!16)
-#1987 := (+ #1984 #1986)
-#1988 := (>= #1987 0::Int)
-#1980 := (f11 f29 ?v0!17)
-#1981 := (= #1980 f1)
-#2850 := (not #1981)
-#1977 := (f11 f29 ?v1!16)
-#1978 := (= #1977 f1)
-#2865 := (or #1978 #2850 #1988)
-#2870 := (not #2865)
-#3914 := (or #2870 #3911)
-#3917 := (not #3914)
-#1176 := (>= #202 0::Int)
-#3863 := (forall (vars (?v0 S2)) (:pat #3829) #1176)
-#3868 := (not #3863)
-#3920 := (or #3868 #3917)
-#3923 := (not #3920)
-decl ?v0!15 :: S2
-#1960 := ?v0!15
-#1961 := (f30 ?v0!15)
-#1962 := (>= #1961 0::Int)
-#1963 := (not #1962)
-#3926 := (or #1963 #3923)
-#3929 := (not #3926)
-#216 := (f30 f16)
-#217 := (= #216 0::Int)
-#661 := (not #217)
-#3932 := (or #661 #3929)
-#3935 := (not #3932)
-#3938 := (or #661 #3935)
-#3941 := (not #3938)
-#112 := (f19 f20 #11)
-#3716 := (pattern #112)
-#207 := (= #202 #112)
-#560 := (or #220 #207)
-#3855 := (forall (vars (?v0 S2)) (:pat #3854 #3829 #3716) #560)
-#3860 := (not #3855)
-#3944 := (or #3860 #3941)
-#3947 := (not #3944)
-decl ?v0!14 :: S2
-#1935 := ?v0!14
-#1940 := (f19 f20 ?v0!14)
-#1939 := (f30 ?v0!14)
-#1941 := (= #1939 #1940)
-#1936 := (f11 f29 ?v0!14)
-#1937 := (= #1936 f1)
-#1938 := (not #1937)
-#1942 := (or #1938 #1941)
-#1943 := (not #1942)
-#3950 := (or #1943 #3947)
-#3953 := (not #3950)
-#1166 := (* -1::Int #202)
-#1167 := (+ #112 #1166)
-#1165 := (>= #1167 0::Int)
-#3846 := (forall (vars (?v0 S2)) (:pat #3716 #3829) #1165)
-#3851 := (not #3846)
-#3956 := (or #3851 #3953)
-#3959 := (not #3956)
-decl ?v0!13 :: S2
-#1917 := ?v0!13
-#1919 := (f30 ?v0!13)
-#1920 := (* -1::Int #1919)
-#1918 := (f19 f20 ?v0!13)
-#1921 := (+ #1918 #1920)
-#1922 := (>= #1921 0::Int)
-#1923 := (not #1922)
-#3962 := (or #1923 #3959)
-#3965 := (not #3962)
-#196 := (f5 #195 #11)
-#3828 := (pattern #196)
-#197 := (f4 #196)
-#1140 := (* -1::Int #197)
-#1146 := (* -1::Int #185)
-#1147 := (+ #1146 #1140)
-#1148 := (+ #112 #1147)
-#1149 := (<= #1148 0::Int)
-#1141 := (+ f3 #1140)
-#1142 := (<= #1141 0::Int)
-#2822 := (or #1142 #1149)
-#2823 := (not #2822)
-#2844 := (or #2823 #207)
-#3838 := (forall (vars (?v0 S2)) (:pat #3828 #3716 #3829) #2844)
-#3843 := (not #3838)
-#1296 := (+ #197 #1166)
-#1297 := (+ #185 #1296)
-#1294 := (= #1297 0::Int)
-#2836 := (or #1142 #1149 #1294)
-#3830 := (forall (vars (?v0 S2)) (:pat #3828 #3716 #3829) #2836)
-#3835 := (not #3830)
-#715 := (not #194)
-#116 := (f11 f21 #11)
-#3750 := (pattern #116)
-#1309 := (+ #112 #1146)
-#1308 := (>= #1309 0::Int)
-#117 := (= #116 f1)
-#1312 := (or #117 #1308)
-#3820 := (forall (vars (?v0 S2)) (:pat #3750 #3716) #1312)
-#3825 := (not #3820)
-#1321 := (+ f3 #1146)
-#1322 := (<= #1321 0::Int)
-#182 := (f11 f21 f28)
-#183 := (= #182 f1)
-decl ?v0!12 :: S2
-#1872 := ?v0!12
-#1876 := (f19 f20 ?v0!12)
-#1877 := (* -1::Int #1876)
-#1878 := (+ f3 #1877)
-#1879 := (<= #1878 0::Int)
-#1873 := (f11 f21 ?v0!12)
-#1874 := (= #1873 f1)
-#3968 := (or #1874 #1879 #183 #1322 #3825 #715 #3835 #3843 #3965)
-#3971 := (not #3968)
-decl f25 :: S11
-#148 := f25
-#168 := (f19 f25 f16)
-#169 := (= #168 0::Int)
-#156 := (f19 f25 #10)
-#1060 := (* -1::Int #156)
-#153 := (f19 f25 #11)
-#1061 := (+ #153 #1060)
-#1067 := (+ #93 #1061)
-#1090 := (>= #1067 0::Int)
-#1047 := (* -1::Int #153)
-#1048 := (+ f3 #1047)
-#1049 := (<= #1048 0::Int)
-#2776 := (or #1049 #940 #1090)
-#3782 := (forall (vars (?v0 S2) (?v1 S2)) (:pat #3684) #2776)
-#3787 := (not #3782)
-#3790 := (or #3787 #169)
-#3793 := (not #3790)
-decl ?v0!11 :: S2
-#1816 := ?v0!11
-#1831 := (f19 f25 ?v0!11)
-#1832 := (* -1::Int #1831)
-decl ?v1!10 :: S2
-#1815 := ?v1!10
-#1822 := (f6 f7 ?v1!10)
-#1823 := (f5 #1822 ?v0!11)
-#1824 := (f4 #1823)
-#2352 := (+ #1824 #1832)
-#1817 := (f19 f25 ?v1!10)
-#2353 := (+ #1817 #2352)
-#2356 := (>= #2353 0::Int)
-#1825 := (* -1::Int #1824)
-#1826 := (+ f3 #1825)
-#1827 := (<= #1826 0::Int)
-#1818 := (* -1::Int #1817)
-#1819 := (+ f3 #1818)
-#1820 := (<= #1819 0::Int)
-#2754 := (or #1820 #1827 #2356)
-#2759 := (not #2754)
-#3796 := (or #2759 #3793)
-#3799 := (not #3796)
-#3759 := (pattern #153)
-decl ?v1!9 :: (-> S2 S2)
-#1791 := (?v1!9 #11)
-#1796 := (f6 f7 #1791)
-#1797 := (f5 #1796 #11)
-#1798 := (f4 #1797)
-#2335 := (* -1::Int #1798)
-#1792 := (f19 f25 #1791)
-#2318 := (* -1::Int #1792)
-#2336 := (+ #2318 #2335)
-#2337 := (+ #153 #2336)
-#2338 := (= #2337 0::Int)
-#2724 := (not #2338)
-#2319 := (+ #153 #2318)
-#2320 := (<= #2319 0::Int)
-#2725 := (or #2320 #2724)
-#2726 := (not #2725)
-#66 := (= #11 f16)
-#2732 := (or #66 #1049 #2726)
-#3774 := (forall (vars (?v0 S2)) (:pat #3759) #2732)
-#3779 := (not #3774)
-#3802 := (or #3779 #3799)
-#3805 := (not #3802)
-decl ?v0!8 :: S2
-#1751 := ?v0!8
-#1764 := (f5 #91 ?v0!8)
-#1765 := (f4 #1764)
-#1754 := (f19 f25 ?v0!8)
-#1755 := (* -1::Int #1754)
-#2288 := (+ #1755 #1765)
-#2289 := (+ #153 #2288)
-#2292 := (= #2289 0::Int)
-#2688 := (not #2292)
-#1761 := (+ #153 #1755)
-#1762 := (>= #1761 0::Int)
-#2689 := (or #1762 #2688)
-#3760 := (forall (vars (?v1 S2)) (:pat #3759 #3693) #2689)
-#3765 := (not #3760)
-#1756 := (+ f3 #1755)
-#1757 := (<= #1756 0::Int)
-#1752 := (= ?v0!8 f16)
-#3768 := (or #1752 #1757 #3765)
-#3771 := (not #3768)
-#3808 := (or #3771 #3805)
-#3811 := (not #3808)
-decl f27 :: S11
-#151 := f27
-decl f26 :: S11
-#150 := f26
-#152 := (= f26 f27)
-#494 := (not #152)
-#149 := (= f25 f20)
-#503 := (not #149)
-decl f24 :: S2
-#146 := f24
-decl f23 :: S2
-#145 := f23
-#147 := (= f23 f24)
-#512 := (not #147)
-decl f22 :: S7
-#143 := f22
-#144 := (= f22 f21)
-#521 := (not #144)
-#1002 := (* -1::Int #112)
-#1003 := (+ f3 #1002)
-#1004 := (<= #1003 0::Int)
-#2674 := (or #117 #1004)
-#3751 := (forall (vars (?v0 S2)) (:pat #3750 #3716) #2674)
-#3756 := (not #3751)
-#3814 := (or #3756 #521 #512 #503 #494 #3811)
-#3817 := (not #3814)
-#3974 := (or #3817 #3971)
-#3977 := (not #3974)
-#1707 := (?v1!7 #11)
-#1714 := (f6 f7 #1707)
-#1715 := (f5 #1714 #11)
-#1716 := (f4 #1715)
-#2261 := (* -1::Int #1716)
-#1708 := (f19 f20 #1707)
-#2244 := (* -1::Int #1708)
-#2262 := (+ #2244 #2261)
-#2263 := (+ #112 #2262)
-#2264 := (= #2263 0::Int)
-#2658 := (not #2264)
-#1712 := (f11 f21 #1707)
-#1713 := (= #1712 f1)
-#2657 := (not #1713)
-#2245 := (+ #112 #2244)
-#2246 := (<= #2245 0::Int)
-#2659 := (or #2246 #2657 #2658)
-#2660 := (not #2659)
-#2666 := (or #66 #1004 #2660)
-#3742 := (forall (vars (?v0 S2)) (:pat #3716) #2666)
-#3747 := (not #3742)
-#122 := (f19 f20 #10)
-#1016 := (* -1::Int #122)
-#1017 := (+ #112 #1016)
-#1018 := (+ #93 #1017)
-#1371 := (>= #1018 0::Int)
-#118 := (not #117)
-#2638 := (or #118 #940 #1371)
-#3734 := (forall (vars (?v0 S2) (?v1 S2)) (:pat #3684) #2638)
-#3739 := (not #3734)
-#119 := (f11 f21 #10)
-#3725 := (pattern #116 #119)
-#1020 := (>= #1017 0::Int)
-#120 := (= #119 f1)
-#2601 := (not #120)
-#2616 := (or #117 #2601 #1020)
-#3726 := (forall (vars (?v0 S2) (?v1 S2)) (:pat #3725) #2616)
-#3731 := (not #3726)
-#1394 := (>= #112 0::Int)
-#3717 := (forall (vars (?v0 S2)) (:pat #3716) #1394)
-#3722 := (not #3717)
-#110 := (f19 f20 f16)
-#111 := (= #110 0::Int)
-#804 := (not #111)
-decl f17 :: (-> S2 Int)
-#67 := (f17 #11)
-#3647 := (pattern #67)
-decl ?v1!6 :: (-> S2 S2)
-#1654 := (?v1!6 #11)
-#1661 := (f6 f7 #1654)
-#1662 := (f5 #1661 #11)
-#1663 := (f4 #1662)
-#2219 := (* -1::Int #1663)
-#1655 := (f17 #1654)
-#2202 := (* -1::Int #1655)
-#2220 := (+ #2202 #2219)
-#2221 := (+ #67 #2220)
-#2222 := (= #2221 0::Int)
-#2585 := (not #2222)
-decl f18 :: S7
-#75 := f18
-#1659 := (f11 f18 #1654)
-#1660 := (= #1659 f1)
-#2584 := (not #1660)
-#2203 := (+ #67 #2202)
-#2204 := (<= #2203 0::Int)
-#2586 := (or #2204 #2584 #2585)
-#2587 := (not #2586)
-#964 := (* -1::Int #67)
-#965 := (+ f3 #964)
-#966 := (<= #965 0::Int)
-#2593 := (or #66 #966 #2587)
-#3708 := (forall (vars (?v0 S2)) (:pat #3647) #2593)
-#3713 := (not #3708)
-#3980 := (or #3713 #804 #3722 #3731 #3739 #3747 #3977)
-#3983 := (not #3980)
-#76 := (f11 f18 #11)
-#3660 := (pattern #76)
-decl ?v0!5 :: S2
-#1613 := ?v0!5
-#1626 := (f5 #91 ?v0!5)
-#1627 := (f4 #1626)
-#1616 := (f17 ?v0!5)
-#1617 := (* -1::Int #1616)
-#1628 := (+ #1617 #1627)
-#1629 := (+ #67 #1628)
-#1630 := (= #1629 0::Int)
-#2548 := (not #1630)
-#77 := (= #76 f1)
-#78 := (not #77)
-#1623 := (+ #67 #1617)
-#1624 := (>= #1623 0::Int)
-#2549 := (or #1624 #78 #2548)
-#3694 := (forall (vars (?v1 S2)) (:pat #3647 #3660 #3693) #2549)
-#3699 := (not #3694)
-#1618 := (+ f3 #1617)
-#1619 := (<= #1618 0::Int)
-#1614 := (= ?v0!5 f16)
-#3702 := (or #1614 #1619 #3699)
-#6176 := (= f3 #1616)
-#6574 := (= #1616 f3)
-#6629 := (iff #6574 #6176)
-#6627 := (iff #6176 #6574)
-#6628 := [commutativity]: #6627
-#6630 := [symm #6628]: #6629
-#1615 := (not #1614)
-#3705 := (not #3702)
-#5790 := [hypothesis]: #3705
-#3262 := (or #3702 #1615)
-#3263 := [def-axiom]: #3262
-#6625 := [unit-resolution #3263 #5790]: #1615
-#72 := (= #67 f3)
-#350 := (or #66 #72)
-#3654 := (forall (vars (?v0 S2)) (:pat #3647) #350)
-#353 := (forall (vars (?v0 S2)) #350)
-#3657 := (iff #353 #3654)
-#3655 := (iff #350 #350)
-#3656 := [refl]: #3655
-#3658 := [quant-intro #3656]: #3657
-#1500 := (~ #353 #353)
-#1530 := (~ #350 #350)
-#1531 := [refl]: #1530
-#1501 := [nnf-pos #1531]: #1500
-#1229 := (= #1207 0::Int)
-#1232 := (not #1185)
-#1241 := (and #1232 #213 #1229)
-#1246 := (exists (vars (?v1 S2)) #1241)
-#1218 := (+ f3 #1166)
-#1219 := (<= #1218 0::Int)
-#1220 := (not #1219)
-#71 := (not #66)
-#1223 := (and #71 #1220)
-#1226 := (not #1223)
-#1249 := (or #1226 #1246)
-#1252 := (forall (vars (?v0 S2)) #1249)
-#941 := (not #940)
-#1199 := (and #213 #941)
-#1202 := (not #1199)
-#1209 := (or #1202 #1205)
-#1212 := (forall (vars (?v0 S2) (?v1 S2)) #1209)
-#1215 := (not #1212)
-#1255 := (or #1215 #1252)
-#1258 := (and #1212 #1255)
-#223 := (and #220 #222)
-#566 := (not #223)
-#1190 := (or #566 #1185)
-#1193 := (forall (vars (?v0 S2) (?v1 S2)) #1190)
-#1196 := (not #1193)
-#1261 := (or #1196 #1258)
-#1264 := (and #1193 #1261)
-#1179 := (forall (vars (?v0 S2)) #1176)
-#1182 := (not #1179)
-#1267 := (or #1182 #1264)
-#1270 := (and #1179 #1267)
-#1273 := (or #661 #1270)
-#1276 := (and #217 #1273)
-#563 := (forall (vars (?v0 S2)) #560)
-#673 := (not #563)
-#1279 := (or #673 #1276)
-#1282 := (and #563 #1279)
-#1170 := (forall (vars (?v0 S2)) #1165)
-#1173 := (not #1170)
-#1285 := (or #1173 #1282)
-#1288 := (and #1170 #1285)
-#1150 := (not #1149)
-#1143 := (not #1142)
-#1153 := (and #1143 #1150)
-#1156 := (or #1153 #207)
-#1159 := (forall (vars (?v0 S2)) #1156)
-#1162 := (not #1159)
-#1291 := (not #1153)
-#1299 := (or #1291 #1294)
-#1302 := (forall (vars (?v0 S2)) #1299)
-#1305 := (not #1302)
-#1315 := (forall (vars (?v0 S2)) #1312)
-#1318 := (not #1315)
-#1005 := (not #1004)
-#1114 := (and #118 #1005)
-#1117 := (exists (vars (?v0 S2)) #1114)
-#1333 := (not #1117)
-#1357 := (or #1333 #183 #1322 #1318 #715 #1305 #1162 #1288)
-#1050 := (not #1049)
-#1084 := (and #1050 #941)
-#1087 := (not #1084)
-#1093 := (or #1087 #1090)
-#1096 := (forall (vars (?v0 S2) (?v1 S2)) #1093)
-#1099 := (not #1096)
-#1102 := (or #1099 #169)
-#1105 := (and #1096 #1102)
-#1065 := (= #1067 0::Int)
-#1059 := (>= #1061 0::Int)
-#1062 := (not #1059)
-#1069 := (and #1062 #1065)
-#1072 := (exists (vars (?v1 S2)) #1069)
-#1053 := (and #71 #1050)
-#1056 := (not #1053)
-#1075 := (or #1056 #1072)
-#1078 := (forall (vars (?v0 S2)) #1075)
-#1081 := (not #1078)
-#1108 := (or #1081 #1105)
-#1111 := (and #1078 #1108)
-#1135 := (or #1117 #521 #512 #503 #494 #1111)
-#1362 := (and #1135 #1357)
-#1014 := (= #1018 0::Int)
-#1021 := (not #1020)
-#1030 := (and #1021 #117 #1014)
-#1035 := (exists (vars (?v1 S2)) #1030)
-#1008 := (and #71 #1005)
-#1011 := (not #1008)
-#1038 := (or #1011 #1035)
-#1041 := (forall (vars (?v0 S2)) #1038)
-#1044 := (not #1041)
-#1365 := (and #117 #941)
-#1368 := (not #1365)
-#1374 := (or #1368 #1371)
-#1377 := (forall (vars (?v0 S2) (?v1 S2)) #1374)
-#1380 := (not #1377)
-#121 := (and #118 #120)
-#377 := (not #121)
-#1385 := (or #377 #1020)
-#1388 := (forall (vars (?v0 S2) (?v1 S2)) #1385)
-#1391 := (not #1388)
-#1397 := (forall (vars (?v0 S2)) #1394)
-#1400 := (not #1397)
-#87 := (f17 #10)
-#926 := (* -1::Int #87)
-#953 := (+ #926 #93)
-#954 := (+ #67 #953)
-#976 := (= #954 0::Int)
-#927 := (+ #67 #926)
-#925 := (>= #927 0::Int)
-#979 := (not #925)
-#988 := (and #979 #77 #976)
-#993 := (exists (vars (?v1 S2)) #988)
-#967 := (not #966)
-#970 := (and #71 #967)
-#973 := (not #970)
-#996 := (or #973 #993)
-#999 := (forall (vars (?v0 S2)) #996)
-#1403 := (not #999)
-#1424 := (or #1403 #804 #1400 #1391 #1380 #1044 #1362)
-#1429 := (and #999 #1424)
-#951 := (>= #954 0::Int)
-#944 := (and #77 #941)
-#947 := (not #944)
-#955 := (or #947 #951)
-#958 := (forall (vars (?v0 S2) (?v1 S2)) #955)
-#961 := (not #958)
-#1432 := (or #961 #1429)
-#1435 := (and #958 #1432)
-#84 := (f11 f18 #10)
-#85 := (= #84 f1)
-#86 := (and #78 #85)
-#356 := (not #86)
-#929 := (or #356 #925)
-#932 := (forall (vars (?v0 S2) (?v1 S2)) #929)
-#935 := (not #932)
-#1438 := (or #935 #1435)
-#1441 := (and #932 #1438)
-#916 := (>= #67 0::Int)
-#917 := (forall (vars (?v0 S2)) #916)
-#920 := (not #917)
-#1444 := (or #920 #1441)
-#1447 := (and #917 #1444)
-#80 := (f17 f16)
-#81 := (= #80 0::Int)
-#868 := (not #81)
-#1450 := (or #868 #1447)
-#1453 := (and #81 #1450)
-#79 := (forall (vars (?v0 S2)) #78)
-#880 := (not #79)
-#889 := (not #353)
-#68 := (= #67 0::Int)
-#344 := (or #71 #68)
-#347 := (forall (vars (?v0 S2)) #344)
-#898 := (not #347)
-#1465 := (or #898 #889 #880 #1453)
-#1470 := (not #1465)
-#229 := (+ #202 #93)
-#236 := (= #224 #229)
-#237 := (and #213 #236)
-#235 := (< #202 #224)
-#238 := (and #235 #237)
-#239 := (exists (vars (?v1 S2)) #238)
-#233 := (< #202 f3)
-#234 := (and #71 #233)
-#240 := (implies #234 #239)
-#241 := (forall (vars (?v0 S2)) #240)
-#242 := (and #241 true)
-#230 := (<= #224 #229)
-#94 := (< #93 f3)
-#228 := (and #213 #94)
-#231 := (implies #228 #230)
-#232 := (forall (vars (?v0 S2) (?v1 S2)) #231)
-#243 := (implies #232 #242)
-#244 := (and #232 #243)
-#225 := (<= #224 #202)
-#226 := (implies #223 #225)
-#227 := (forall (vars (?v0 S2) (?v1 S2)) #226)
-#245 := (implies #227 #244)
-#246 := (and #227 #245)
-#218 := (<= 0::Int #202)
-#219 := (forall (vars (?v0 S2)) #218)
-#247 := (implies #219 #246)
-#248 := (and #219 #247)
-#249 := (implies #217 #248)
-#250 := (and #217 #249)
-#214 := (implies #213 #207)
-#215 := (forall (vars (?v0 S2)) #214)
-#251 := (implies #215 #250)
-#252 := (and #215 #251)
-#210 := (<= #202 #112)
-#211 := (forall (vars (?v0 S2)) #210)
-#253 := (implies #211 #252)
-#254 := (and #211 #253)
-#199 := (+ #185 #197)
-#200 := (< #199 #112)
-#198 := (< #197 f3)
-#201 := (and #198 #200)
-#206 := (not #201)
-#208 := (implies #206 #207)
-#209 := (forall (vars (?v0 S2)) #208)
-#255 := (implies #209 #254)
-#203 := (= #202 #199)
-#204 := (implies #201 #203)
-#205 := (forall (vars (?v0 S2)) #204)
-#256 := (implies #205 #255)
-#257 := (implies #194 #256)
-#187 := (<= #185 #112)
-#188 := (implies #118 #187)
-#189 := (forall (vars (?v0 S2)) #188)
-#258 := (implies #189 #257)
-#186 := (< #185 f3)
-#259 := (implies #186 #258)
-#184 := (not #183)
-#260 := (implies #184 #259)
-#131 := (< #112 f3)
-#140 := (and #118 #131)
-#141 := (exists (vars (?v0 S2)) #140)
-#261 := (implies #141 #260)
-#262 := (implies true #261)
-#170 := (and #169 true)
-#158 := (+ #153 #93)
-#165 := (<= #156 #158)
-#154 := (< #153 f3)
-#164 := (and #154 #94)
-#166 := (implies #164 #165)
-#167 := (forall (vars (?v0 S2) (?v1 S2)) #166)
-#171 := (implies #167 #170)
-#172 := (and #167 #171)
-#159 := (= #156 #158)
-#157 := (< #153 #156)
-#160 := (and #157 #159)
-#161 := (exists (vars (?v1 S2)) #160)
-#155 := (and #71 #154)
-#162 := (implies #155 #161)
-#163 := (forall (vars (?v0 S2)) #162)
-#173 := (implies #163 #172)
-#174 := (and #163 #173)
-#175 := (implies #152 #174)
-#176 := (implies #149 #175)
-#177 := (implies #147 #176)
-#178 := (implies #144 #177)
-#142 := (not #141)
-#179 := (implies #142 #178)
-#180 := (implies true #179)
-#263 := (and #180 #262)
-#127 := (+ #112 #93)
-#134 := (= #122 #127)
-#135 := (and #117 #134)
-#133 := (< #112 #122)
-#136 := (and #133 #135)
-#137 := (exists (vars (?v1 S2)) #136)
-#132 := (and #71 #131)
-#138 := (implies #132 #137)
-#139 := (forall (vars (?v0 S2)) #138)
-#264 := (implies #139 #263)
-#128 := (<= #122 #127)
-#126 := (and #117 #94)
-#129 := (implies #126 #128)
-#130 := (forall (vars (?v0 S2) (?v1 S2)) #129)
-#265 := (implies #130 #264)
-#123 := (<= #122 #112)
-#124 := (implies #121 #123)
-#125 := (forall (vars (?v0 S2) (?v1 S2)) #124)
-#266 := (implies #125 #265)
-#113 := (<= 0::Int #112)
-#114 := (forall (vars (?v0 S2)) #113)
-#267 := (implies #114 #266)
-#268 := (implies #111 #267)
-#269 := (implies true #268)
-#96 := (+ #67 #93)
-#103 := (= #87 #96)
-#104 := (and #77 #103)
-#102 := (< #67 #87)
-#105 := (and #102 #104)
-#106 := (exists (vars (?v1 S2)) #105)
-#100 := (< #67 f3)
-#101 := (and #71 #100)
-#107 := (implies #101 #106)
-#108 := (forall (vars (?v0 S2)) #107)
-#270 := (implies #108 #269)
-#271 := (and #108 #270)
-#97 := (<= #87 #96)
-#95 := (and #77 #94)
-#98 := (implies #95 #97)
-#99 := (forall (vars (?v0 S2) (?v1 S2)) #98)
-#272 := (implies #99 #271)
-#273 := (and #99 #272)
-#88 := (<= #87 #67)
-#89 := (implies #86 #88)
-#90 := (forall (vars (?v0 S2) (?v1 S2)) #89)
-#274 := (implies #90 #273)
-#275 := (and #90 #274)
-#82 := (<= 0::Int #67)
-#83 := (forall (vars (?v0 S2)) #82)
-#276 := (implies #83 #275)
-#277 := (and #83 #276)
-#278 := (implies #81 #277)
-#279 := (and #81 #278)
-#280 := (implies #79 #279)
-#73 := (implies #71 #72)
-#74 := (forall (vars (?v0 S2)) #73)
-#281 := (implies #74 #280)
-#69 := (implies #66 #68)
-#70 := (forall (vars (?v0 S2)) #69)
-#282 := (implies #70 #281)
-#283 := (implies true #282)
-#284 := (not #283)
-#1473 := (iff #284 #1470)
-#573 := (+ #93 #202)
-#591 := (= #224 #573)
-#594 := (and #213 #591)
-#597 := (and #235 #594)
-#600 := (exists (vars (?v1 S2)) #597)
-#606 := (not #234)
-#607 := (or #606 #600)
-#612 := (forall (vars (?v0 S2)) #607)
-#576 := (<= #224 #573)
-#582 := (not #228)
-#583 := (or #582 #576)
-#588 := (forall (vars (?v0 S2) (?v1 S2)) #583)
-#625 := (not #588)
-#626 := (or #625 #612)
-#631 := (and #588 #626)
-#567 := (or #566 #225)
-#570 := (forall (vars (?v0 S2) (?v1 S2)) #567)
-#637 := (not #570)
-#638 := (or #637 #631)
-#643 := (and #570 #638)
-#649 := (not #219)
-#650 := (or #649 #643)
-#655 := (and #219 #650)
-#662 := (or #661 #655)
-#667 := (and #217 #662)
-#674 := (or #673 #667)
-#679 := (and #563 #674)
-#685 := (not #211)
-#686 := (or #685 #679)
-#691 := (and #211 #686)
-#554 := (or #201 #207)
-#557 := (forall (vars (?v0 S2)) #554)
-#697 := (not #557)
-#698 := (or #697 #691)
-#548 := (or #206 #203)
-#551 := (forall (vars (?v0 S2)) #548)
-#706 := (not #551)
-#707 := (or #706 #698)
-#716 := (or #715 #707)
-#542 := (or #117 #187)
-#545 := (forall (vars (?v0 S2)) #542)
-#724 := (not #545)
-#725 := (or #724 #716)
-#733 := (not #186)
-#734 := (or #733 #725)
-#742 := (or #183 #734)
-#750 := (or #142 #742)
-#426 := (+ #93 #153)
-#450 := (<= #156 #426)
-#456 := (not #164)
-#457 := (or #456 #450)
-#462 := (forall (vars (?v0 S2) (?v1 S2)) #457)
-#470 := (not #462)
-#471 := (or #470 #169)
-#476 := (and #462 #471)
-#429 := (= #156 #426)
-#432 := (and #157 #429)
-#435 := (exists (vars (?v1 S2)) #432)
-#441 := (not #155)
-#442 := (or #441 #435)
-#447 := (forall (vars (?v0 S2)) #442)
-#482 := (not #447)
-#483 := (or #482 #476)
-#488 := (and #447 #483)
-#495 := (or #494 #488)
-#504 := (or #503 #495)
-#513 := (or #512 #504)
-#522 := (or #521 #513)
-#530 := (or #141 #522)
-#762 := (and #530 #750)
-#384 := (+ #93 #112)
-#402 := (= #122 #384)
-#405 := (and #117 #402)
-#408 := (and #133 #405)
-#411 := (exists (vars (?v1 S2)) #408)
-#417 := (not #132)
-#418 := (or #417 #411)
-#423 := (forall (vars (?v0 S2)) #418)
-#768 := (not #423)
-#769 := (or #768 #762)
-#387 := (<= #122 #384)
-#393 := (not #126)
-#394 := (or #393 #387)
-#399 := (forall (vars (?v0 S2) (?v1 S2)) #394)
-#777 := (not #399)
-#778 := (or #777 #769)
-#378 := (or #377 #123)
-#381 := (forall (vars (?v0 S2) (?v1 S2)) #378)
-#786 := (not #381)
-#787 := (or #786 #778)
-#795 := (not #114)
-#796 := (or #795 #787)
-#805 := (or #804 #796)
-#370 := (not #101)
-#371 := (or #370 #106)
-#374 := (forall (vars (?v0 S2)) #371)
-#820 := (not #374)
-#821 := (or #820 #805)
-#826 := (and #374 #821)
-#363 := (not #95)
-#364 := (or #363 #97)
-#367 := (forall (vars (?v0 S2) (?v1 S2)) #364)
-#832 := (not #367)
-#833 := (or #832 #826)
-#838 := (and #367 #833)
-#357 := (or #356 #88)
-#360 := (forall (vars (?v0 S2) (?v1 S2)) #357)
-#844 := (not #360)
-#845 := (or #844 #838)
-#850 := (and #360 #845)
-#856 := (not #83)
-#857 := (or #856 #850)
-#862 := (and #83 #857)
-#869 := (or #868 #862)
-#874 := (and #81 #869)
-#881 := (or #880 #874)
-#890 := (or #889 #881)
-#899 := (or #898 #890)
-#911 := (not #899)
-#1471 := (iff #911 #1470)
-#1468 := (iff #899 #1465)
-#1456 := (or #880 #1453)
-#1459 := (or #889 #1456)
-#1462 := (or #898 #1459)
-#1466 := (iff #1462 #1465)
-#1467 := [rewrite]: #1466
-#1463 := (iff #899 #1462)
-#1460 := (iff #890 #1459)
-#1457 := (iff #881 #1456)
-#1454 := (iff #874 #1453)
-#1451 := (iff #869 #1450)
-#1448 := (iff #862 #1447)
-#1445 := (iff #857 #1444)
-#1442 := (iff #850 #1441)
-#1439 := (iff #845 #1438)
-#1436 := (iff #838 #1435)
-#1433 := (iff #833 #1432)
-#1430 := (iff #826 #1429)
-#1427 := (iff #821 #1424)
-#1406 := (or #1044 #1362)
-#1409 := (or #1380 #1406)
-#1412 := (or #1391 #1409)
-#1415 := (or #1400 #1412)
-#1418 := (or #804 #1415)
-#1421 := (or #1403 #1418)
-#1425 := (iff #1421 #1424)
-#1426 := [rewrite]: #1425
-#1422 := (iff #821 #1421)
-#1419 := (iff #805 #1418)
-#1416 := (iff #796 #1415)
-#1413 := (iff #787 #1412)
-#1410 := (iff #778 #1409)
-#1407 := (iff #769 #1406)
-#1363 := (iff #762 #1362)
-#1360 := (iff #750 #1357)
-#1336 := (or #1162 #1288)
-#1339 := (or #1305 #1336)
-#1342 := (or #715 #1339)
-#1345 := (or #1318 #1342)
-#1348 := (or #1322 #1345)
-#1351 := (or #183 #1348)
-#1354 := (or #1333 #1351)
-#1358 := (iff #1354 #1357)
-#1359 := [rewrite]: #1358
-#1355 := (iff #750 #1354)
-#1352 := (iff #742 #1351)
-#1349 := (iff #734 #1348)
-#1346 := (iff #725 #1345)
-#1343 := (iff #716 #1342)
-#1340 := (iff #707 #1339)
-#1337 := (iff #698 #1336)
-#1289 := (iff #691 #1288)
-#1286 := (iff #686 #1285)
-#1283 := (iff #679 #1282)
-#1280 := (iff #674 #1279)
-#1277 := (iff #667 #1276)
-#1274 := (iff #662 #1273)
-#1271 := (iff #655 #1270)
-#1268 := (iff #650 #1267)
-#1265 := (iff #643 #1264)
-#1262 := (iff #638 #1261)
-#1259 := (iff #631 #1258)
-#1256 := (iff #626 #1255)
-#1253 := (iff #612 #1252)
-#1250 := (iff #607 #1249)
-#1247 := (iff #600 #1246)
-#1244 := (iff #597 #1241)
-#1235 := (and #213 #1229)
-#1238 := (and #1232 #1235)
-#1242 := (iff #1238 #1241)
-#1243 := [rewrite]: #1242
-#1239 := (iff #597 #1238)
-#1236 := (iff #594 #1235)
-#1230 := (iff #591 #1229)
-#1231 := [rewrite]: #1230
-#1237 := [monotonicity #1231]: #1236
-#1233 := (iff #235 #1232)
-#1234 := [rewrite]: #1233
-#1240 := [monotonicity #1234 #1237]: #1239
-#1245 := [trans #1240 #1243]: #1244
-#1248 := [quant-intro #1245]: #1247
-#1227 := (iff #606 #1226)
-#1224 := (iff #234 #1223)
-#1221 := (iff #233 #1220)
-#1222 := [rewrite]: #1221
-#1225 := [monotonicity #1222]: #1224
-#1228 := [monotonicity #1225]: #1227
-#1251 := [monotonicity #1228 #1248]: #1250
-#1254 := [quant-intro #1251]: #1253
-#1216 := (iff #625 #1215)
-#1213 := (iff #588 #1212)
-#1210 := (iff #583 #1209)
-#1206 := (iff #576 #1205)
-#1208 := [rewrite]: #1206
-#1203 := (iff #582 #1202)
-#1200 := (iff #228 #1199)
-#942 := (iff #94 #941)
-#943 := [rewrite]: #942
-#1201 := [monotonicity #943]: #1200
-#1204 := [monotonicity #1201]: #1203
-#1211 := [monotonicity #1204 #1208]: #1210
-#1214 := [quant-intro #1211]: #1213
-#1217 := [monotonicity #1214]: #1216
-#1257 := [monotonicity #1217 #1254]: #1256
-#1260 := [monotonicity #1214 #1257]: #1259
-#1197 := (iff #637 #1196)
-#1194 := (iff #570 #1193)
-#1191 := (iff #567 #1190)
-#1188 := (iff #225 #1185)
-#1189 := [rewrite]: #1188
-#1192 := [monotonicity #1189]: #1191
-#1195 := [quant-intro #1192]: #1194
-#1198 := [monotonicity #1195]: #1197
-#1263 := [monotonicity #1198 #1260]: #1262
-#1266 := [monotonicity #1195 #1263]: #1265
-#1183 := (iff #649 #1182)
-#1180 := (iff #219 #1179)
-#1177 := (iff #218 #1176)
-#1178 := [rewrite]: #1177
-#1181 := [quant-intro #1178]: #1180
-#1184 := [monotonicity #1181]: #1183
-#1269 := [monotonicity #1184 #1266]: #1268
-#1272 := [monotonicity #1181 #1269]: #1271
-#1275 := [monotonicity #1272]: #1274
-#1278 := [monotonicity #1275]: #1277
-#1281 := [monotonicity #1278]: #1280
-#1284 := [monotonicity #1281]: #1283
-#1174 := (iff #685 #1173)
-#1171 := (iff #211 #1170)
-#1168 := (iff #210 #1165)
-#1169 := [rewrite]: #1168
-#1172 := [quant-intro #1169]: #1171
-#1175 := [monotonicity #1172]: #1174
-#1287 := [monotonicity #1175 #1284]: #1286
-#1290 := [monotonicity #1172 #1287]: #1289
-#1163 := (iff #697 #1162)
-#1160 := (iff #557 #1159)
-#1157 := (iff #554 #1156)
-#1154 := (iff #201 #1153)
-#1151 := (iff #200 #1150)
-#1152 := [rewrite]: #1151
-#1144 := (iff #198 #1143)
-#1145 := [rewrite]: #1144
-#1155 := [monotonicity #1145 #1152]: #1154
-#1158 := [monotonicity #1155]: #1157
-#1161 := [quant-intro #1158]: #1160
-#1164 := [monotonicity #1161]: #1163
-#1338 := [monotonicity #1164 #1290]: #1337
-#1306 := (iff #706 #1305)
-#1303 := (iff #551 #1302)
-#1300 := (iff #548 #1299)
-#1295 := (iff #203 #1294)
-#1298 := [rewrite]: #1295
-#1292 := (iff #206 #1291)
-#1293 := [monotonicity #1155]: #1292
-#1301 := [monotonicity #1293 #1298]: #1300
-#1304 := [quant-intro #1301]: #1303
-#1307 := [monotonicity #1304]: #1306
-#1341 := [monotonicity #1307 #1338]: #1340
-#1344 := [monotonicity #1341]: #1343
-#1319 := (iff #724 #1318)
-#1316 := (iff #545 #1315)
-#1313 := (iff #542 #1312)
-#1310 := (iff #187 #1308)
-#1311 := [rewrite]: #1310
-#1314 := [monotonicity #1311]: #1313
-#1317 := [quant-intro #1314]: #1316
-#1320 := [monotonicity #1317]: #1319
-#1347 := [monotonicity #1320 #1344]: #1346
-#1331 := (iff #733 #1322)
-#1323 := (not #1322)
-#1326 := (not #1323)
-#1329 := (iff #1326 #1322)
-#1330 := [rewrite]: #1329
-#1327 := (iff #733 #1326)
-#1324 := (iff #186 #1323)
-#1325 := [rewrite]: #1324
-#1328 := [monotonicity #1325]: #1327
-#1332 := [trans #1328 #1330]: #1331
-#1350 := [monotonicity #1332 #1347]: #1349
-#1353 := [monotonicity #1350]: #1352
-#1334 := (iff #142 #1333)
-#1118 := (iff #141 #1117)
-#1115 := (iff #140 #1114)
-#1006 := (iff #131 #1005)
-#1007 := [rewrite]: #1006
-#1116 := [monotonicity #1007]: #1115
-#1119 := [quant-intro #1116]: #1118
-#1335 := [monotonicity #1119]: #1334
-#1356 := [monotonicity #1335 #1353]: #1355
-#1361 := [trans #1356 #1359]: #1360
-#1138 := (iff #530 #1135)
-#1120 := (or #494 #1111)
-#1123 := (or #503 #1120)
-#1126 := (or #512 #1123)
-#1129 := (or #521 #1126)
-#1132 := (or #1117 #1129)
-#1136 := (iff #1132 #1135)
-#1137 := [rewrite]: #1136
-#1133 := (iff #530 #1132)
-#1130 := (iff #522 #1129)
-#1127 := (iff #513 #1126)
-#1124 := (iff #504 #1123)
-#1121 := (iff #495 #1120)
-#1112 := (iff #488 #1111)
-#1109 := (iff #483 #1108)
-#1106 := (iff #476 #1105)
-#1103 := (iff #471 #1102)
-#1100 := (iff #470 #1099)
-#1097 := (iff #462 #1096)
-#1094 := (iff #457 #1093)
-#1091 := (iff #450 #1090)
-#1092 := [rewrite]: #1091
-#1088 := (iff #456 #1087)
-#1085 := (iff #164 #1084)
-#1051 := (iff #154 #1050)
-#1052 := [rewrite]: #1051
-#1086 := [monotonicity #1052 #943]: #1085
-#1089 := [monotonicity #1086]: #1088
-#1095 := [monotonicity #1089 #1092]: #1094
-#1098 := [quant-intro #1095]: #1097
-#1101 := [monotonicity #1098]: #1100
-#1104 := [monotonicity #1101]: #1103
-#1107 := [monotonicity #1098 #1104]: #1106
-#1082 := (iff #482 #1081)
-#1079 := (iff #447 #1078)
-#1076 := (iff #442 #1075)
-#1073 := (iff #435 #1072)
-#1070 := (iff #432 #1069)
-#1066 := (iff #429 #1065)
-#1068 := [rewrite]: #1066
-#1063 := (iff #157 #1062)
-#1064 := [rewrite]: #1063
-#1071 := [monotonicity #1064 #1068]: #1070
-#1074 := [quant-intro #1071]: #1073
-#1057 := (iff #441 #1056)
-#1054 := (iff #155 #1053)
-#1055 := [monotonicity #1052]: #1054
-#1058 := [monotonicity #1055]: #1057
-#1077 := [monotonicity #1058 #1074]: #1076
-#1080 := [quant-intro #1077]: #1079
-#1083 := [monotonicity #1080]: #1082
-#1110 := [monotonicity #1083 #1107]: #1109
-#1113 := [monotonicity #1080 #1110]: #1112
-#1122 := [monotonicity #1113]: #1121
-#1125 := [monotonicity #1122]: #1124
-#1128 := [monotonicity #1125]: #1127
-#1131 := [monotonicity #1128]: #1130
-#1134 := [monotonicity #1119 #1131]: #1133
-#1139 := [trans #1134 #1137]: #1138
-#1364 := [monotonicity #1139 #1361]: #1363
-#1045 := (iff #768 #1044)
-#1042 := (iff #423 #1041)
-#1039 := (iff #418 #1038)
-#1036 := (iff #411 #1035)
-#1033 := (iff #408 #1030)
-#1024 := (and #117 #1014)
-#1027 := (and #1021 #1024)
-#1031 := (iff #1027 #1030)
-#1032 := [rewrite]: #1031
-#1028 := (iff #408 #1027)
-#1025 := (iff #405 #1024)
-#1015 := (iff #402 #1014)
-#1019 := [rewrite]: #1015
-#1026 := [monotonicity #1019]: #1025
-#1022 := (iff #133 #1021)
-#1023 := [rewrite]: #1022
-#1029 := [monotonicity #1023 #1026]: #1028
-#1034 := [trans #1029 #1032]: #1033
-#1037 := [quant-intro #1034]: #1036
-#1012 := (iff #417 #1011)
-#1009 := (iff #132 #1008)
-#1010 := [monotonicity #1007]: #1009
-#1013 := [monotonicity #1010]: #1012
-#1040 := [monotonicity #1013 #1037]: #1039
-#1043 := [quant-intro #1040]: #1042
-#1046 := [monotonicity #1043]: #1045
-#1408 := [monotonicity #1046 #1364]: #1407
-#1381 := (iff #777 #1380)
-#1378 := (iff #399 #1377)
-#1375 := (iff #394 #1374)
-#1372 := (iff #387 #1371)
-#1373 := [rewrite]: #1372
-#1369 := (iff #393 #1368)
-#1366 := (iff #126 #1365)
-#1367 := [monotonicity #943]: #1366
-#1370 := [monotonicity #1367]: #1369
-#1376 := [monotonicity #1370 #1373]: #1375
-#1379 := [quant-intro #1376]: #1378
-#1382 := [monotonicity #1379]: #1381
-#1411 := [monotonicity #1382 #1408]: #1410
-#1392 := (iff #786 #1391)
-#1389 := (iff #381 #1388)
-#1386 := (iff #378 #1385)
-#1383 := (iff #123 #1020)
-#1384 := [rewrite]: #1383
-#1387 := [monotonicity #1384]: #1386
-#1390 := [quant-intro #1387]: #1389
-#1393 := [monotonicity #1390]: #1392
-#1414 := [monotonicity #1393 #1411]: #1413
-#1401 := (iff #795 #1400)
-#1398 := (iff #114 #1397)
-#1395 := (iff #113 #1394)
-#1396 := [rewrite]: #1395
-#1399 := [quant-intro #1396]: #1398
-#1402 := [monotonicity #1399]: #1401
-#1417 := [monotonicity #1402 #1414]: #1416
-#1420 := [monotonicity #1417]: #1419
-#1404 := (iff #820 #1403)
-#1000 := (iff #374 #999)
-#997 := (iff #371 #996)
-#994 := (iff #106 #993)
-#991 := (iff #105 #988)
-#982 := (and #77 #976)
-#985 := (and #979 #982)
-#989 := (iff #985 #988)
-#990 := [rewrite]: #989
-#986 := (iff #105 #985)
-#983 := (iff #104 #982)
-#977 := (iff #103 #976)
-#978 := [rewrite]: #977
-#984 := [monotonicity #978]: #983
-#980 := (iff #102 #979)
-#981 := [rewrite]: #980
-#987 := [monotonicity #981 #984]: #986
-#992 := [trans #987 #990]: #991
-#995 := [quant-intro #992]: #994
-#974 := (iff #370 #973)
-#971 := (iff #101 #970)
-#968 := (iff #100 #967)
-#969 := [rewrite]: #968
-#972 := [monotonicity #969]: #971
-#975 := [monotonicity #972]: #974
-#998 := [monotonicity #975 #995]: #997
-#1001 := [quant-intro #998]: #1000
-#1405 := [monotonicity #1001]: #1404
-#1423 := [monotonicity #1405 #1420]: #1422
-#1428 := [trans #1423 #1426]: #1427
-#1431 := [monotonicity #1001 #1428]: #1430
-#962 := (iff #832 #961)
-#959 := (iff #367 #958)
-#956 := (iff #364 #955)
-#950 := (iff #97 #951)
-#952 := [rewrite]: #950
-#948 := (iff #363 #947)
-#945 := (iff #95 #944)
-#946 := [monotonicity #943]: #945
-#949 := [monotonicity #946]: #948
-#957 := [monotonicity #949 #952]: #956
-#960 := [quant-intro #957]: #959
-#963 := [monotonicity #960]: #962
-#1434 := [monotonicity #963 #1431]: #1433
-#1437 := [monotonicity #960 #1434]: #1436
-#936 := (iff #844 #935)
-#933 := (iff #360 #932)
-#930 := (iff #357 #929)
-#924 := (iff #88 #925)
-#928 := [rewrite]: #924
-#931 := [monotonicity #928]: #930
-#934 := [quant-intro #931]: #933
-#937 := [monotonicity #934]: #936
-#1440 := [monotonicity #937 #1437]: #1439
-#1443 := [monotonicity #934 #1440]: #1442
-#921 := (iff #856 #920)
-#918 := (iff #83 #917)
-#914 := (iff #82 #916)
-#915 := [rewrite]: #914
-#919 := [quant-intro #915]: #918
-#922 := [monotonicity #919]: #921
-#1446 := [monotonicity #922 #1443]: #1445
-#1449 := [monotonicity #919 #1446]: #1448
-#1452 := [monotonicity #1449]: #1451
-#1455 := [monotonicity #1452]: #1454
-#1458 := [monotonicity #1455]: #1457
-#1461 := [monotonicity #1458]: #1460
-#1464 := [monotonicity #1461]: #1463
-#1469 := [trans #1464 #1467]: #1468
-#1472 := [monotonicity #1469]: #1471
-#912 := (iff #284 #911)
-#909 := (iff #283 #899)
-#904 := (implies true #899)
-#907 := (iff #904 #899)
-#908 := [rewrite]: #907
-#905 := (iff #283 #904)
-#902 := (iff #282 #899)
-#895 := (implies #347 #890)
-#900 := (iff #895 #899)
-#901 := [rewrite]: #900
-#896 := (iff #282 #895)
-#893 := (iff #281 #890)
-#886 := (implies #353 #881)
-#891 := (iff #886 #890)
-#892 := [rewrite]: #891
-#887 := (iff #281 #886)
-#884 := (iff #280 #881)
-#877 := (implies #79 #874)
-#882 := (iff #877 #881)
-#883 := [rewrite]: #882
-#878 := (iff #280 #877)
-#875 := (iff #279 #874)
-#872 := (iff #278 #869)
-#865 := (implies #81 #862)
-#870 := (iff #865 #869)
-#871 := [rewrite]: #870
-#866 := (iff #278 #865)
-#863 := (iff #277 #862)
-#860 := (iff #276 #857)
-#853 := (implies #83 #850)
-#858 := (iff #853 #857)
-#859 := [rewrite]: #858
-#854 := (iff #276 #853)
-#851 := (iff #275 #850)
-#848 := (iff #274 #845)
-#841 := (implies #360 #838)
-#846 := (iff #841 #845)
-#847 := [rewrite]: #846
-#842 := (iff #274 #841)
-#839 := (iff #273 #838)
-#836 := (iff #272 #833)
-#829 := (implies #367 #826)
-#834 := (iff #829 #833)
-#835 := [rewrite]: #834
-#830 := (iff #272 #829)
-#827 := (iff #271 #826)
-#824 := (iff #270 #821)
-#817 := (implies #374 #805)
-#822 := (iff #817 #821)
-#823 := [rewrite]: #822
-#818 := (iff #270 #817)
-#815 := (iff #269 #805)
-#810 := (implies true #805)
-#813 := (iff #810 #805)
-#814 := [rewrite]: #813
-#811 := (iff #269 #810)
-#808 := (iff #268 #805)
-#801 := (implies #111 #796)
-#806 := (iff #801 #805)
-#807 := [rewrite]: #806
-#802 := (iff #268 #801)
-#799 := (iff #267 #796)
-#792 := (implies #114 #787)
-#797 := (iff #792 #796)
-#798 := [rewrite]: #797
-#793 := (iff #267 #792)
-#790 := (iff #266 #787)
-#783 := (implies #381 #778)
-#788 := (iff #783 #787)
-#789 := [rewrite]: #788
-#784 := (iff #266 #783)
-#781 := (iff #265 #778)
-#774 := (implies #399 #769)
-#779 := (iff #774 #778)
-#780 := [rewrite]: #779
-#775 := (iff #265 #774)
-#772 := (iff #264 #769)
-#765 := (implies #423 #762)
-#770 := (iff #765 #769)
-#771 := [rewrite]: #770
-#766 := (iff #264 #765)
-#763 := (iff #263 #762)
-#760 := (iff #262 #750)
-#755 := (implies true #750)
-#758 := (iff #755 #750)
-#759 := [rewrite]: #758
-#756 := (iff #262 #755)
-#753 := (iff #261 #750)
-#747 := (implies #141 #742)
-#751 := (iff #747 #750)
-#752 := [rewrite]: #751
-#748 := (iff #261 #747)
-#745 := (iff #260 #742)
-#739 := (implies #184 #734)
-#743 := (iff #739 #742)
-#744 := [rewrite]: #743
-#740 := (iff #260 #739)
-#737 := (iff #259 #734)
-#730 := (implies #186 #725)
-#735 := (iff #730 #734)
-#736 := [rewrite]: #735
-#731 := (iff #259 #730)
-#728 := (iff #258 #725)
-#721 := (implies #545 #716)
-#726 := (iff #721 #725)
-#727 := [rewrite]: #726
-#722 := (iff #258 #721)
-#719 := (iff #257 #716)
-#712 := (implies #194 #707)
-#717 := (iff #712 #716)
-#718 := [rewrite]: #717
-#713 := (iff #257 #712)
-#710 := (iff #256 #707)
-#703 := (implies #551 #698)
-#708 := (iff #703 #707)
-#709 := [rewrite]: #708
-#704 := (iff #256 #703)
-#701 := (iff #255 #698)
-#694 := (implies #557 #691)
-#699 := (iff #694 #698)
-#700 := [rewrite]: #699
-#695 := (iff #255 #694)
-#692 := (iff #254 #691)
-#689 := (iff #253 #686)
-#682 := (implies #211 #679)
-#687 := (iff #682 #686)
-#688 := [rewrite]: #687
-#683 := (iff #253 #682)
-#680 := (iff #252 #679)
-#677 := (iff #251 #674)
-#670 := (implies #563 #667)
-#675 := (iff #670 #674)
-#676 := [rewrite]: #675
-#671 := (iff #251 #670)
-#668 := (iff #250 #667)
-#665 := (iff #249 #662)
-#658 := (implies #217 #655)
-#663 := (iff #658 #662)
-#664 := [rewrite]: #663
-#659 := (iff #249 #658)
-#656 := (iff #248 #655)
-#653 := (iff #247 #650)
-#646 := (implies #219 #643)
-#651 := (iff #646 #650)
-#652 := [rewrite]: #651
-#647 := (iff #247 #646)
-#644 := (iff #246 #643)
-#641 := (iff #245 #638)
-#634 := (implies #570 #631)
-#639 := (iff #634 #638)
-#640 := [rewrite]: #639
-#635 := (iff #245 #634)
-#632 := (iff #244 #631)
-#629 := (iff #243 #626)
-#622 := (implies #588 #612)
-#627 := (iff #622 #626)
-#628 := [rewrite]: #627
-#623 := (iff #243 #622)
-#620 := (iff #242 #612)
-#615 := (and #612 true)
-#618 := (iff #615 #612)
-#619 := [rewrite]: #618
-#616 := (iff #242 #615)
-#613 := (iff #241 #612)
-#610 := (iff #240 #607)
-#603 := (implies #234 #600)
-#608 := (iff #603 #607)
-#609 := [rewrite]: #608
-#604 := (iff #240 #603)
-#601 := (iff #239 #600)
-#598 := (iff #238 #597)
-#595 := (iff #237 #594)
-#592 := (iff #236 #591)
-#574 := (= #229 #573)
-#575 := [rewrite]: #574
-#593 := [monotonicity #575]: #592
-#596 := [monotonicity #593]: #595
-#599 := [monotonicity #596]: #598
-#602 := [quant-intro #599]: #601
-#605 := [monotonicity #602]: #604
-#611 := [trans #605 #609]: #610
-#614 := [quant-intro #611]: #613
-#617 := [monotonicity #614]: #616
-#621 := [trans #617 #619]: #620
-#589 := (iff #232 #588)
-#586 := (iff #231 #583)
-#579 := (implies #228 #576)
-#584 := (iff #579 #583)
-#585 := [rewrite]: #584
-#580 := (iff #231 #579)
-#577 := (iff #230 #576)
-#578 := [monotonicity #575]: #577
-#581 := [monotonicity #578]: #580
-#587 := [trans #581 #585]: #586
-#590 := [quant-intro #587]: #589
-#624 := [monotonicity #590 #621]: #623
-#630 := [trans #624 #628]: #629
-#633 := [monotonicity #590 #630]: #632
-#571 := (iff #227 #570)
-#568 := (iff #226 #567)
-#569 := [rewrite]: #568
-#572 := [quant-intro #569]: #571
-#636 := [monotonicity #572 #633]: #635
-#642 := [trans #636 #640]: #641
-#645 := [monotonicity #572 #642]: #644
-#648 := [monotonicity #645]: #647
-#654 := [trans #648 #652]: #653
-#657 := [monotonicity #654]: #656
-#660 := [monotonicity #657]: #659
-#666 := [trans #660 #664]: #665
-#669 := [monotonicity #666]: #668
-#564 := (iff #215 #563)
-#561 := (iff #214 #560)
-#562 := [rewrite]: #561
-#565 := [quant-intro #562]: #564
-#672 := [monotonicity #565 #669]: #671
-#678 := [trans #672 #676]: #677
-#681 := [monotonicity #565 #678]: #680
-#684 := [monotonicity #681]: #683
-#690 := [trans #684 #688]: #689
-#693 := [monotonicity #690]: #692
-#558 := (iff #209 #557)
-#555 := (iff #208 #554)
-#556 := [rewrite]: #555
-#559 := [quant-intro #556]: #558
-#696 := [monotonicity #559 #693]: #695
-#702 := [trans #696 #700]: #701
-#552 := (iff #205 #551)
-#549 := (iff #204 #548)
-#550 := [rewrite]: #549
-#553 := [quant-intro #550]: #552
-#705 := [monotonicity #553 #702]: #704
-#711 := [trans #705 #709]: #710
-#714 := [monotonicity #711]: #713
-#720 := [trans #714 #718]: #719
-#546 := (iff #189 #545)
-#543 := (iff #188 #542)
-#544 := [rewrite]: #543
-#547 := [quant-intro #544]: #546
-#723 := [monotonicity #547 #720]: #722
-#729 := [trans #723 #727]: #728
-#732 := [monotonicity #729]: #731
-#738 := [trans #732 #736]: #737
-#741 := [monotonicity #738]: #740
-#746 := [trans #741 #744]: #745
-#749 := [monotonicity #746]: #748
-#754 := [trans #749 #752]: #753
-#757 := [monotonicity #754]: #756
-#761 := [trans #757 #759]: #760
-#540 := (iff #180 #530)
-#535 := (implies true #530)
-#538 := (iff #535 #530)
-#539 := [rewrite]: #538
-#536 := (iff #180 #535)
-#533 := (iff #179 #530)
-#527 := (implies #142 #522)
-#531 := (iff #527 #530)
-#532 := [rewrite]: #531
-#528 := (iff #179 #527)
-#525 := (iff #178 #522)
-#518 := (implies #144 #513)
-#523 := (iff #518 #522)
-#524 := [rewrite]: #523
-#519 := (iff #178 #518)
-#516 := (iff #177 #513)
-#509 := (implies #147 #504)
-#514 := (iff #509 #513)
-#515 := [rewrite]: #514
-#510 := (iff #177 #509)
-#507 := (iff #176 #504)
-#500 := (implies #149 #495)
-#505 := (iff #500 #504)
-#506 := [rewrite]: #505
-#501 := (iff #176 #500)
-#498 := (iff #175 #495)
-#491 := (implies #152 #488)
-#496 := (iff #491 #495)
-#497 := [rewrite]: #496
-#492 := (iff #175 #491)
-#489 := (iff #174 #488)
-#486 := (iff #173 #483)
-#479 := (implies #447 #476)
-#484 := (iff #479 #483)
-#485 := [rewrite]: #484
-#480 := (iff #173 #479)
-#477 := (iff #172 #476)
-#474 := (iff #171 #471)
-#467 := (implies #462 #169)
-#472 := (iff #467 #471)
-#473 := [rewrite]: #472
-#468 := (iff #171 #467)
-#465 := (iff #170 #169)
-#466 := [rewrite]: #465
-#463 := (iff #167 #462)
-#460 := (iff #166 #457)
-#453 := (implies #164 #450)
-#458 := (iff #453 #457)
-#459 := [rewrite]: #458
-#454 := (iff #166 #453)
-#451 := (iff #165 #450)
-#427 := (= #158 #426)
-#428 := [rewrite]: #427
-#452 := [monotonicity #428]: #451
-#455 := [monotonicity #452]: #454
-#461 := [trans #455 #459]: #460
-#464 := [quant-intro #461]: #463
-#469 := [monotonicity #464 #466]: #468
-#475 := [trans #469 #473]: #474
-#478 := [monotonicity #464 #475]: #477
-#448 := (iff #163 #447)
-#445 := (iff #162 #442)
-#438 := (implies #155 #435)
-#443 := (iff #438 #442)
-#444 := [rewrite]: #443
-#439 := (iff #162 #438)
-#436 := (iff #161 #435)
-#433 := (iff #160 #432)
-#430 := (iff #159 #429)
-#431 := [monotonicity #428]: #430
-#434 := [monotonicity #431]: #433
-#437 := [quant-intro #434]: #436
-#440 := [monotonicity #437]: #439
-#446 := [trans #440 #444]: #445
-#449 := [quant-intro #446]: #448
-#481 := [monotonicity #449 #478]: #480
-#487 := [trans #481 #485]: #486
-#490 := [monotonicity #449 #487]: #489
-#493 := [monotonicity #490]: #492
-#499 := [trans #493 #497]: #498
-#502 := [monotonicity #499]: #501
-#508 := [trans #502 #506]: #507
-#511 := [monotonicity #508]: #510
-#517 := [trans #511 #515]: #516
-#520 := [monotonicity #517]: #519
-#526 := [trans #520 #524]: #525
-#529 := [monotonicity #526]: #528
-#534 := [trans #529 #532]: #533
-#537 := [monotonicity #534]: #536
-#541 := [trans #537 #539]: #540
-#764 := [monotonicity #541 #761]: #763
-#424 := (iff #139 #423)
-#421 := (iff #138 #418)
-#414 := (implies #132 #411)
-#419 := (iff #414 #418)
-#420 := [rewrite]: #419
-#415 := (iff #138 #414)
-#412 := (iff #137 #411)
-#409 := (iff #136 #408)
-#406 := (iff #135 #405)
-#403 := (iff #134 #402)
-#385 := (= #127 #384)
-#386 := [rewrite]: #385
-#404 := [monotonicity #386]: #403
-#407 := [monotonicity #404]: #406
-#410 := [monotonicity #407]: #409
-#413 := [quant-intro #410]: #412
-#416 := [monotonicity #413]: #415
-#422 := [trans #416 #420]: #421
-#425 := [quant-intro #422]: #424
-#767 := [monotonicity #425 #764]: #766
-#773 := [trans #767 #771]: #772
-#400 := (iff #130 #399)
-#397 := (iff #129 #394)
-#390 := (implies #126 #387)
-#395 := (iff #390 #394)
-#396 := [rewrite]: #395
-#391 := (iff #129 #390)
-#388 := (iff #128 #387)
-#389 := [monotonicity #386]: #388
-#392 := [monotonicity #389]: #391
-#398 := [trans #392 #396]: #397
-#401 := [quant-intro #398]: #400
-#776 := [monotonicity #401 #773]: #775
-#782 := [trans #776 #780]: #781
-#382 := (iff #125 #381)
-#379 := (iff #124 #378)
-#380 := [rewrite]: #379
-#383 := [quant-intro #380]: #382
-#785 := [monotonicity #383 #782]: #784
-#791 := [trans #785 #789]: #790
-#794 := [monotonicity #791]: #793
-#800 := [trans #794 #798]: #799
-#803 := [monotonicity #800]: #802
-#809 := [trans #803 #807]: #808
-#812 := [monotonicity #809]: #811
-#816 := [trans #812 #814]: #815
-#375 := (iff #108 #374)
-#372 := (iff #107 #371)
-#373 := [rewrite]: #372
-#376 := [quant-intro #373]: #375
-#819 := [monotonicity #376 #816]: #818
-#825 := [trans #819 #823]: #824
-#828 := [monotonicity #376 #825]: #827
-#368 := (iff #99 #367)
-#365 := (iff #98 #364)
-#366 := [rewrite]: #365
-#369 := [quant-intro #366]: #368
-#831 := [monotonicity #369 #828]: #830
-#837 := [trans #831 #835]: #836
-#840 := [monotonicity #369 #837]: #839
-#361 := (iff #90 #360)
-#358 := (iff #89 #357)
-#359 := [rewrite]: #358
-#362 := [quant-intro #359]: #361
-#843 := [monotonicity #362 #840]: #842
-#849 := [trans #843 #847]: #848
-#852 := [monotonicity #362 #849]: #851
-#855 := [monotonicity #852]: #854
-#861 := [trans #855 #859]: #860
-#864 := [monotonicity #861]: #863
-#867 := [monotonicity #864]: #866
-#873 := [trans #867 #871]: #872
-#876 := [monotonicity #873]: #875
-#879 := [monotonicity #876]: #878
-#885 := [trans #879 #883]: #884
-#354 := (iff #74 #353)
-#351 := (iff #73 #350)
-#352 := [rewrite]: #351
-#355 := [quant-intro #352]: #354
-#888 := [monotonicity #355 #885]: #887
-#894 := [trans #888 #892]: #893
-#348 := (iff #70 #347)
-#345 := (iff #69 #344)
-#346 := [rewrite]: #345
-#349 := [quant-intro #346]: #348
-#897 := [monotonicity #349 #894]: #896
-#903 := [trans #897 #901]: #902
-#906 := [monotonicity #903]: #905
-#910 := [trans #906 #908]: #909
-#913 := [monotonicity #910]: #912
-#1474 := [trans #913 #1472]: #1473
-#343 := [asserted]: #284
-#1475 := [mp #343 #1474]: #1470
-#1477 := [not-or-elim #1475]: #353
-#1532 := [mp~ #1477 #1501]: #353
-#3659 := [mp #1532 #3658]: #3654
-#3289 := (not #3654)
-#5853 := (or #3289 #1614 #6574)
-#6575 := (or #1614 #6574)
-#6155 := (or #3289 #6575)
-#6156 := (iff #6155 #5853)
-#6165 := [rewrite]: #6156
-#5702 := [quant-inst #1613]: #6155
-#6166 := [mp #5702 #6165]: #5853
-#6626 := [unit-resolution #6166 #3659 #6625]: #6574
-#6631 := [mp #6626 #6630]: #6176
-#6615 := (not #6176)
-#1620 := (not #1619)
-#3595 := (or #3702 #1620)
-#3596 := [def-axiom]: #3595
-#5852 := [unit-resolution #3596 #5790]: #1620
-#6616 := (or #6615 #1619)
-#6622 := [th-lemma arith triangle-eq]: #6616
-#6624 := [unit-resolution #6622 #5852]: #6615
-#6632 := [unit-resolution #6624 #6631]: false
-#6633 := [lemma #6632]: #3702
-#3986 := (or #3705 #3983)
-#3989 := (not #3986)
-#2540 := (or #78 #940 #951)
-#3685 := (forall (vars (?v0 S2) (?v1 S2)) (:pat #3684) #2540)
-#3690 := (not #3685)
-#3992 := (or #3690 #3989)
-#3995 := (not #3992)
-decl ?v0!4 :: S2
-#1582 := ?v0!4
-#1595 := (f17 ?v0!4)
-#1596 := (* -1::Int #1595)
-decl ?v1!3 :: S2
-#1581 := ?v1!3
-#1594 := (f17 ?v1!3)
-#2173 := (+ #1594 #1596)
-#1585 := (f6 f7 ?v1!3)
-#1586 := (f5 #1585 ?v0!4)
-#1587 := (f4 #1586)
-#2174 := (+ #1587 #2173)
-#2177 := (>= #2174 0::Int)
-#1588 := (* -1::Int #1587)
-#1589 := (+ f3 #1588)
-#1590 := (<= #1589 0::Int)
-#1583 := (f11 f18 ?v1!3)
-#1584 := (= #1583 f1)
-#2503 := (not #1584)
-#2518 := (or #2503 #1590 #2177)
-#2523 := (not #2518)
-#3998 := (or #2523 #3995)
-#4001 := (not #3998)
-#3675 := (pattern #67 #87)
-#1605 := (not #85)
-#2495 := (or #77 #1605 #925)
-#3676 := (forall (vars (?v0 S2) (?v1 S2)) (:pat #3675) #2495)
-#3681 := (not #3676)
-#4004 := (or #3681 #4001)
-#4007 := (not #4004)
-decl ?v0!2 :: S2
-#1555 := ?v0!2
-#1564 := (f17 ?v0!2)
-#1565 := (* -1::Int #1564)
-decl ?v1!1 :: S2
-#1554 := ?v1!1
-#1563 := (f17 ?v1!1)
-#1566 := (+ #1563 #1565)
-#1567 := (>= #1566 0::Int)
-#1559 := (f11 f18 ?v0!2)
-#1560 := (= #1559 f1)
-#1884 := (not #1560)
-#1556 := (f11 f18 ?v1!1)
-#1557 := (= #1556 f1)
-#1841 := (or #1557 #1884 #1567)
-#1946 := (not #1841)
-#4010 := (or #1946 #4007)
-#4013 := (not #4010)
-#3667 := (forall (vars (?v0 S2)) (:pat #3647) #916)
-#3672 := (not #3667)
-#4016 := (or #3672 #4013)
-#4019 := (not #4016)
-decl ?v0!0 :: S2
-#1539 := ?v0!0
-#1540 := (f17 ?v0!0)
-#1541 := (>= #1540 0::Int)
-#3259 := (= f3 #1540)
-#3321 := (= #1540 f3)
-#3223 := (iff #3321 #3259)
-#3224 := (iff #3259 #3321)
-#3227 := [commutativity]: #3224
-#3222 := [symm #3227]: #3223
-#3345 := (= ?v0!0 f16)
-#3353 := (not #3345)
-#3306 := (= #1540 0::Int)
-#3281 := (not #3306)
-#1542 := (not #1541)
-#3290 := [hypothesis]: #1542
-#3246 := (or #3281 #1541)
-#3244 := [th-lemma arith triangle-eq]: #3246
-#3247 := [unit-resolution #3244 #3290]: #3281
-#3648 := (forall (vars (?v0 S2)) (:pat #3647) #344)
-#3651 := (iff #347 #3648)
-#3649 := (iff #344 #344)
-#3650 := [refl]: #3649
-#3652 := [quant-intro #3650]: #3651
-#1498 := (~ #347 #347)
-#1527 := (~ #344 #344)
-#1528 := [refl]: #1527
-#1499 := [nnf-pos #1528]: #1498
-#1476 := [not-or-elim #1475]: #347
-#1529 := [mp~ #1476 #1499]: #347
-#3653 := [mp #1529 #3652]: #3648
-#3310 := (not #3648)
-#3309 := (or #3310 #3353 #3306)
-#3307 := (or #3353 #3306)
-#3303 := (or #3310 #3307)
-#3294 := (iff #3303 #3309)
-#3295 := [rewrite]: #3294
-#3304 := [quant-inst #1539]: #3303
-#3305 := [mp #3304 #3295]: #3309
-#3236 := [unit-resolution #3305 #3653 #3247]: #3353
-#3291 := (or #3289 #3345 #3321)
-#3308 := (or #3345 #3321)
-#3258 := (or #3289 #3308)
-#3268 := (iff #3258 #3291)
-#3269 := [rewrite]: #3268
-#3260 := [quant-inst #1539]: #3258
-#3267 := [mp #3260 #3269]: #3291
-#3248 := [unit-resolution #3267 #3659 #3236]: #3321
-#3228 := [mp #3248 #3222]: #3259
-#3316 := (* -1::Int #1540)
-#3270 := (+ f3 #3316)
-#3253 := (<= #3270 0::Int)
-#4045 := (not #3253)
-#307 := (<= f3 0::Int)
-#308 := (not #307)
-#9 := (< 0::Int f3)
-#309 := (iff #9 #308)
-#310 := [rewrite]: #309
-#304 := [asserted]: #9
-#311 := [mp #304 #310]: #308
-#3286 := (<= #1540 0::Int)
-#3203 := (or #3286 #1541)
-#3208 := [th-lemma arith farkas 1 1]: #3203
-#3213 := [unit-resolution #3208 #3290]: #3286
-#4046 := (not #3286)
-#4047 := (or #4045 #307 #4046)
-#4048 := [th-lemma arith assign-bounds 1 1]: #4047
-#4049 := [unit-resolution #4048 #3213 #311]: #4045
-#4044 := (not #3259)
-#4050 := (or #4044 #3253)
-#4051 := [th-lemma arith triangle-eq]: #4050
-#4052 := [unit-resolution #4051 #4049 #3228]: false
-#4053 := [lemma #4052]: #1541
-#4022 := (or #1542 #4019)
-#4025 := (not #4022)
-#4028 := (or #868 #4025)
-#4031 := (not #4028)
-#4121 := [hypothesis]: #868
-#4075 := (or #3310 #81)
-#6826 := (= f16 f16)
-#6905 := (not #6826)
-#4054 := (or #6905 #81)
-#4076 := (or #3310 #4054)
-#4084 := (iff #4076 #4075)
-#4117 := (iff #4075 #4075)
-#4118 := [rewrite]: #4117
-#4073 := (iff #4054 #81)
-#4068 := (or false #81)
-#4069 := (iff #4068 #81)
-#4072 := [rewrite]: #4069
-#4070 := (iff #4054 #4068)
-#6915 := (iff #6905 false)
-#6910 := (not true)
-#6913 := (iff #6910 false)
-#6914 := [rewrite]: #6913
-#6911 := (iff #6905 #6910)
-#6829 := (iff #6826 true)
-#6830 := [rewrite]: #6829
-#6912 := [monotonicity #6830]: #6911
-#6916 := [trans #6912 #6914]: #6915
-#4071 := [monotonicity #6916]: #4070
-#4074 := [trans #4071 #4072]: #4073
-#4085 := [monotonicity #4074]: #4084
-#4119 := [trans #4085 #4118]: #4084
-#4077 := [quant-inst #65]: #4076
-#4120 := [mp #4077 #4119]: #4075
-#4116 := [unit-resolution #4120 #3653 #4121]: false
-#4122 := [lemma #4116]: #81
-#4034 := (or #868 #4031)
-#2953 := (forall (vars (?v1 S2)) #2942)
-#2960 := (not #2953)
-#2938 := (forall (vars (?v0 S2) (?v1 S2)) #2933)
-#2959 := (not #2938)
-#2961 := (or #2959 #2035 #2040 #2960)
-#2962 := (not #2961)
-#2967 := (or #2916 #2962)
-#2974 := (not #2967)
-#2893 := (forall (vars (?v0 S2) (?v1 S2)) #2888)
-#2973 := (not #2893)
-#2975 := (or #2973 #2974)
-#2976 := (not #2975)
-#2981 := (or #2870 #2976)
-#2987 := (not #2981)
-#2988 := (or #1182 #2987)
-#2989 := (not #2988)
-#2994 := (or #1963 #2989)
-#3000 := (not #2994)
-#3001 := (or #661 #3000)
-#3002 := (not #3001)
-#3007 := (or #661 #3002)
-#3013 := (not #3007)
-#3014 := (or #673 #3013)
-#3015 := (not #3014)
-#3020 := (or #1943 #3015)
-#3026 := (not #3020)
-#3027 := (or #1173 #3026)
-#3028 := (not #3027)
-#3033 := (or #1923 #3028)
-#3041 := (not #3033)
-#2847 := (forall (vars (?v0 S2)) #2844)
-#3040 := (not #2847)
-#2841 := (forall (vars (?v0 S2)) #2836)
-#3039 := (not #2841)
-#3042 := (or #1874 #1879 #183 #1322 #1318 #715 #3039 #3040 #3041)
-#3043 := (not #3042)
-#2781 := (forall (vars (?v0 S2) (?v1 S2)) #2776)
-#2787 := (not #2781)
-#2788 := (or #2787 #169)
-#2789 := (not #2788)
-#2794 := (or #2759 #2789)
-#2801 := (not #2794)
-#2737 := (forall (vars (?v0 S2)) #2732)
-#2800 := (not #2737)
-#2802 := (or #2800 #2801)
-#2803 := (not #2802)
-#2700 := (forall (vars (?v1 S2)) #2689)
-#2706 := (not #2700)
-#2707 := (or #1752 #1757 #2706)
-#2708 := (not #2707)
-#2808 := (or #2708 #2803)
-#2815 := (not #2808)
-#2685 := (forall (vars (?v0 S2)) #2674)
-#2814 := (not #2685)
-#2816 := (or #2814 #521 #512 #503 #494 #2815)
-#2817 := (not #2816)
-#3048 := (or #2817 #3043)
-#3058 := (not #3048)
-#2671 := (forall (vars (?v0 S2)) #2666)
-#3057 := (not #2671)
-#2643 := (forall (vars (?v0 S2) (?v1 S2)) #2638)
-#3056 := (not #2643)
-#2621 := (forall (vars (?v0 S2) (?v1 S2)) #2616)
-#3055 := (not #2621)
-#2598 := (forall (vars (?v0 S2)) #2593)
-#3054 := (not #2598)
-#3059 := (or #3054 #804 #1400 #3055 #3056 #3057 #3058)
-#3060 := (not #3059)
-#2560 := (forall (vars (?v1 S2)) #2549)
-#2566 := (not #2560)
-#2567 := (or #1614 #1619 #2566)
-#2568 := (not #2567)
-#3065 := (or #2568 #3060)
-#3072 := (not #3065)
-#2545 := (forall (vars (?v0 S2) (?v1 S2)) #2540)
-#3071 := (not #2545)
-#3073 := (or #3071 #3072)
-#3074 := (not #3073)
-#3079 := (or #2523 #3074)
-#3086 := (not #3079)
-#2500 := (forall (vars (?v0 S2) (?v1 S2)) #2495)
-#3085 := (not #2500)
-#3087 := (or #3085 #3086)
-#3088 := (not #3087)
-#3093 := (or #1946 #3088)
-#3099 := (not #3093)
-#3100 := (or #920 #3099)
-#3101 := (not #3100)
-#3106 := (or #1542 #3101)
-#3112 := (not #3106)
-#3113 := (or #868 #3112)
-#3114 := (not #3113)
-#3119 := (or #868 #3114)
-#4035 := (iff #3119 #4034)
-#4032 := (iff #3114 #4031)
-#4029 := (iff #3113 #4028)
-#4026 := (iff #3112 #4025)
-#4023 := (iff #3106 #4022)
-#4020 := (iff #3101 #4019)
-#4017 := (iff #3100 #4016)
-#4014 := (iff #3099 #4013)
-#4011 := (iff #3093 #4010)
-#4008 := (iff #3088 #4007)
-#4005 := (iff #3087 #4004)
-#4002 := (iff #3086 #4001)
-#3999 := (iff #3079 #3998)
-#3996 := (iff #3074 #3995)
-#3993 := (iff #3073 #3992)
-#3990 := (iff #3072 #3989)
-#3987 := (iff #3065 #3986)
-#3984 := (iff #3060 #3983)
-#3981 := (iff #3059 #3980)
-#3978 := (iff #3058 #3977)
-#3975 := (iff #3048 #3974)
-#3972 := (iff #3043 #3971)
-#3969 := (iff #3042 #3968)
-#3966 := (iff #3041 #3965)
-#3963 := (iff #3033 #3962)
-#3960 := (iff #3028 #3959)
-#3957 := (iff #3027 #3956)
-#3954 := (iff #3026 #3953)
-#3951 := (iff #3020 #3950)
-#3948 := (iff #3015 #3947)
-#3945 := (iff #3014 #3944)
-#3942 := (iff #3013 #3941)
-#3939 := (iff #3007 #3938)
-#3936 := (iff #3002 #3935)
-#3933 := (iff #3001 #3932)
-#3930 := (iff #3000 #3929)
-#3927 := (iff #2994 #3926)
-#3924 := (iff #2989 #3923)
-#3921 := (iff #2988 #3920)
-#3918 := (iff #2987 #3917)
-#3915 := (iff #2981 #3914)
-#3912 := (iff #2976 #3911)
-#3909 := (iff #2975 #3908)
-#3906 := (iff #2974 #3905)
-#3903 := (iff #2967 #3902)
-#3900 := (iff #2962 #3899)
-#3897 := (iff #2961 #3896)
-#3894 := (iff #2960 #3893)
-#3891 := (iff #2953 #3888)
-#3889 := (iff #2942 #2942)
-#3890 := [refl]: #3889
-#3892 := [quant-intro #3890]: #3891
-#3895 := [monotonicity #3892]: #3894
-#3886 := (iff #2959 #3885)
-#3883 := (iff #2938 #3880)
-#3881 := (iff #2933 #2933)
-#3882 := [refl]: #3881
-#3884 := [quant-intro #3882]: #3883
-#3887 := [monotonicity #3884]: #3886
-#3898 := [monotonicity #3887 #3895]: #3897
-#3901 := [monotonicity #3898]: #3900
-#3904 := [monotonicity #3901]: #3903
-#3907 := [monotonicity #3904]: #3906
-#3878 := (iff #2973 #3877)
-#3875 := (iff #2893 #3872)
-#3873 := (iff #2888 #2888)
-#3874 := [refl]: #3873
-#3876 := [quant-intro #3874]: #3875
-#3879 := [monotonicity #3876]: #3878
-#3910 := [monotonicity #3879 #3907]: #3909
-#3913 := [monotonicity #3910]: #3912
-#3916 := [monotonicity #3913]: #3915
-#3919 := [monotonicity #3916]: #3918
-#3869 := (iff #1182 #3868)
-#3866 := (iff #1179 #3863)
-#3864 := (iff #1176 #1176)
-#3865 := [refl]: #3864
-#3867 := [quant-intro #3865]: #3866
-#3870 := [monotonicity #3867]: #3869
-#3922 := [monotonicity #3870 #3919]: #3921
-#3925 := [monotonicity #3922]: #3924
-#3928 := [monotonicity #3925]: #3927
-#3931 := [monotonicity #3928]: #3930
-#3934 := [monotonicity #3931]: #3933
-#3937 := [monotonicity #3934]: #3936
-#3940 := [monotonicity #3937]: #3939
-#3943 := [monotonicity #3940]: #3942
-#3861 := (iff #673 #3860)
-#3858 := (iff #563 #3855)
-#3856 := (iff #560 #560)
-#3857 := [refl]: #3856
-#3859 := [quant-intro #3857]: #3858
-#3862 := [monotonicity #3859]: #3861
-#3946 := [monotonicity #3862 #3943]: #3945
-#3949 := [monotonicity #3946]: #3948
-#3952 := [monotonicity #3949]: #3951
-#3955 := [monotonicity #3952]: #3954
-#3852 := (iff #1173 #3851)
-#3849 := (iff #1170 #3846)
-#3847 := (iff #1165 #1165)
-#3848 := [refl]: #3847
-#3850 := [quant-intro #3848]: #3849
-#3853 := [monotonicity #3850]: #3852
-#3958 := [monotonicity #3853 #3955]: #3957
-#3961 := [monotonicity #3958]: #3960
-#3964 := [monotonicity #3961]: #3963
-#3967 := [monotonicity #3964]: #3966
-#3844 := (iff #3040 #3843)
-#3841 := (iff #2847 #3838)
-#3839 := (iff #2844 #2844)
-#3840 := [refl]: #3839
-#3842 := [quant-intro #3840]: #3841
-#3845 := [monotonicity #3842]: #3844
-#3836 := (iff #3039 #3835)
-#3833 := (iff #2841 #3830)
-#3831 := (iff #2836 #2836)
-#3832 := [refl]: #3831
-#3834 := [quant-intro #3832]: #3833
-#3837 := [monotonicity #3834]: #3836
-#3826 := (iff #1318 #3825)
-#3823 := (iff #1315 #3820)
-#3821 := (iff #1312 #1312)
-#3822 := [refl]: #3821
-#3824 := [quant-intro #3822]: #3823
-#3827 := [monotonicity #3824]: #3826
-#3970 := [monotonicity #3827 #3837 #3845 #3967]: #3969
-#3973 := [monotonicity #3970]: #3972
-#3818 := (iff #2817 #3817)
-#3815 := (iff #2816 #3814)
-#3812 := (iff #2815 #3811)
-#3809 := (iff #2808 #3808)
-#3806 := (iff #2803 #3805)
-#3803 := (iff #2802 #3802)
-#3800 := (iff #2801 #3799)
-#3797 := (iff #2794 #3796)
-#3794 := (iff #2789 #3793)
-#3791 := (iff #2788 #3790)
-#3788 := (iff #2787 #3787)
-#3785 := (iff #2781 #3782)
-#3783 := (iff #2776 #2776)
-#3784 := [refl]: #3783
-#3786 := [quant-intro #3784]: #3785
-#3789 := [monotonicity #3786]: #3788
-#3792 := [monotonicity #3789]: #3791
-#3795 := [monotonicity #3792]: #3794
-#3798 := [monotonicity #3795]: #3797
-#3801 := [monotonicity #3798]: #3800
-#3780 := (iff #2800 #3779)
-#3777 := (iff #2737 #3774)
-#3775 := (iff #2732 #2732)
-#3776 := [refl]: #3775
-#3778 := [quant-intro #3776]: #3777
-#3781 := [monotonicity #3778]: #3780
-#3804 := [monotonicity #3781 #3801]: #3803
-#3807 := [monotonicity #3804]: #3806
-#3772 := (iff #2708 #3771)
-#3769 := (iff #2707 #3768)
-#3766 := (iff #2706 #3765)
-#3763 := (iff #2700 #3760)
-#3761 := (iff #2689 #2689)
-#3762 := [refl]: #3761
-#3764 := [quant-intro #3762]: #3763
-#3767 := [monotonicity #3764]: #3766
-#3770 := [monotonicity #3767]: #3769
-#3773 := [monotonicity #3770]: #3772
-#3810 := [monotonicity #3773 #3807]: #3809
-#3813 := [monotonicity #3810]: #3812
-#3757 := (iff #2814 #3756)
-#3754 := (iff #2685 #3751)
-#3752 := (iff #2674 #2674)
-#3753 := [refl]: #3752
-#3755 := [quant-intro #3753]: #3754
-#3758 := [monotonicity #3755]: #3757
-#3816 := [monotonicity #3758 #3813]: #3815
-#3819 := [monotonicity #3816]: #3818
-#3976 := [monotonicity #3819 #3973]: #3975
-#3979 := [monotonicity #3976]: #3978
-#3748 := (iff #3057 #3747)
-#3745 := (iff #2671 #3742)
-#3743 := (iff #2666 #2666)
-#3744 := [refl]: #3743
-#3746 := [quant-intro #3744]: #3745
-#3749 := [monotonicity #3746]: #3748
-#3740 := (iff #3056 #3739)
-#3737 := (iff #2643 #3734)
-#3735 := (iff #2638 #2638)
-#3736 := [refl]: #3735
-#3738 := [quant-intro #3736]: #3737
-#3741 := [monotonicity #3738]: #3740
-#3732 := (iff #3055 #3731)
-#3729 := (iff #2621 #3726)
-#3727 := (iff #2616 #2616)
-#3728 := [refl]: #3727
-#3730 := [quant-intro #3728]: #3729
-#3733 := [monotonicity #3730]: #3732
-#3723 := (iff #1400 #3722)
-#3720 := (iff #1397 #3717)
-#3718 := (iff #1394 #1394)
-#3719 := [refl]: #3718
-#3721 := [quant-intro #3719]: #3720
-#3724 := [monotonicity #3721]: #3723
-#3714 := (iff #3054 #3713)
-#3711 := (iff #2598 #3708)
-#3709 := (iff #2593 #2593)
-#3710 := [refl]: #3709
-#3712 := [quant-intro #3710]: #3711
-#3715 := [monotonicity #3712]: #3714
-#3982 := [monotonicity #3715 #3724 #3733 #3741 #3749 #3979]: #3981
-#3985 := [monotonicity #3982]: #3984
-#3706 := (iff #2568 #3705)
-#3703 := (iff #2567 #3702)
-#3700 := (iff #2566 #3699)
-#3697 := (iff #2560 #3694)
-#3695 := (iff #2549 #2549)
-#3696 := [refl]: #3695
-#3698 := [quant-intro #3696]: #3697
-#3701 := [monotonicity #3698]: #3700
-#3704 := [monotonicity #3701]: #3703
-#3707 := [monotonicity #3704]: #3706
-#3988 := [monotonicity #3707 #3985]: #3987
-#3991 := [monotonicity #3988]: #3990
-#3691 := (iff #3071 #3690)
-#3688 := (iff #2545 #3685)
-#3686 := (iff #2540 #2540)
-#3687 := [refl]: #3686
-#3689 := [quant-intro #3687]: #3688
-#3692 := [monotonicity #3689]: #3691
-#3994 := [monotonicity #3692 #3991]: #3993
-#3997 := [monotonicity #3994]: #3996
-#4000 := [monotonicity #3997]: #3999
-#4003 := [monotonicity #4000]: #4002
-#3682 := (iff #3085 #3681)
-#3679 := (iff #2500 #3676)
-#3677 := (iff #2495 #2495)
-#3678 := [refl]: #3677
-#3680 := [quant-intro #3678]: #3679
-#3683 := [monotonicity #3680]: #3682
-#4006 := [monotonicity #3683 #4003]: #4005
-#4009 := [monotonicity #4006]: #4008
-#4012 := [monotonicity #4009]: #4011
-#4015 := [monotonicity #4012]: #4014
-#3673 := (iff #920 #3672)
-#3670 := (iff #917 #3667)
-#3668 := (iff #916 #916)
-#3669 := [refl]: #3668
-#3671 := [quant-intro #3669]: #3670
-#3674 := [monotonicity #3671]: #3673
-#4018 := [monotonicity #3674 #4015]: #4017
-#4021 := [monotonicity #4018]: #4020
-#4024 := [monotonicity #4021]: #4023
-#4027 := [monotonicity #4024]: #4026
-#4030 := [monotonicity #4027]: #4029
-#4033 := [monotonicity #4030]: #4032
-#4036 := [monotonicity #4033]: #4035
-#2046 := (not #2045)
-#2390 := (and #2046 #213 #2387)
-#2393 := (not #2390)
-#2396 := (forall (vars (?v1 S2)) #2393)
-#2041 := (not #2040)
-#2036 := (not #2035)
-#2405 := (and #1212 #2036 #2041 #2396)
-#2012 := (not #2011)
-#2013 := (and #2005 #2012)
-#2014 := (not #2013)
-#2021 := (or #2014 #2020)
-#2022 := (not #2021)
-#2410 := (or #2022 #2405)
-#2413 := (and #1193 #2410)
-#1979 := (not #1978)
-#1982 := (and #1979 #1981)
-#1983 := (not #1982)
-#1989 := (or #1983 #1988)
-#1990 := (not #1989)
-#2416 := (or #1990 #2413)
-#2419 := (and #1179 #2416)
-#2422 := (or #1963 #2419)
-#2425 := (and #217 #2422)
-#2428 := (or #661 #2425)
-#2431 := (and #563 #2428)
-#2434 := (or #1943 #2431)
-#2437 := (and #1170 #2434)
-#2440 := (or #1923 #2437)
-#1880 := (not #1879)
-#1875 := (not #1874)
-#2446 := (and #1875 #1880 #184 #1323 #1315 #194 #1302 #1159 #2440)
-#1849 := (not #169)
-#1852 := (and #1096 #1849)
-#1828 := (not #1827)
-#1821 := (not #1820)
-#1829 := (and #1821 #1828)
-#1830 := (not #1829)
-#2359 := (or #1830 #2356)
-#2362 := (not #2359)
-#2365 := (or #2362 #1852)
-#2325 := (not #2320)
-#2343 := (and #2325 #2338)
-#2346 := (or #1056 #2343)
-#2349 := (forall (vars (?v0 S2)) #2346)
-#2368 := (and #2349 #2365)
-#1763 := (not #1762)
-#2295 := (and #1763 #2292)
-#2298 := (not #2295)
-#2301 := (forall (vars (?v1 S2)) #2298)
-#1758 := (not #1757)
-#1753 := (not #1752)
-#2307 := (and #1753 #1758 #2301)
-#2371 := (or #2307 #2368)
-#1733 := (not #1114)
-#1736 := (forall (vars (?v0 S2)) #1733)
-#2374 := (and #1736 #144 #147 #149 #152 #2371)
-#2451 := (or #2374 #2446)
-#2251 := (not #2246)
-#2269 := (and #2251 #1713 #2264)
-#2272 := (or #1011 #2269)
-#2275 := (forall (vars (?v0 S2)) #2272)
-#2209 := (not #2204)
-#2227 := (and #2209 #1660 #2222)
-#2230 := (or #973 #2227)
-#2233 := (forall (vars (?v0 S2)) #2230)
-#2454 := (and #2233 #111 #1397 #1388 #1377 #2275 #2451)
-#1625 := (not #1624)
-#1631 := (and #1625 #77 #1630)
-#1640 := (not #1631)
-#1643 := (forall (vars (?v1 S2)) #1640)
-#2191 := (and #1615 #1620 #1643)
-#2457 := (or #2191 #2454)
-#2460 := (and #958 #2457)
-#1591 := (not #1590)
-#1592 := (and #1584 #1591)
-#1593 := (not #1592)
-#2180 := (or #1593 #2177)
-#2183 := (not #2180)
-#2463 := (or #2183 #2460)
-#2466 := (and #932 #2463)
-#1558 := (not #1557)
-#1561 := (and #1558 #1560)
-#1562 := (not #1561)
-#1568 := (or #1562 #1567)
-#1569 := (not #1568)
-#2469 := (or #1569 #2466)
-#2472 := (and #917 #2469)
-#2475 := (or #1542 #2472)
-#2478 := (and #81 #2475)
-#2481 := (or #868 #2478)
-#3120 := (iff #2481 #3119)
-#3117 := (iff #2478 #3114)
-#3109 := (and #81 #3106)
-#3115 := (iff #3109 #3114)
-#3116 := [rewrite]: #3115
-#3110 := (iff #2478 #3109)
-#3107 := (iff #2475 #3106)
-#3104 := (iff #2472 #3101)
-#3096 := (and #917 #3093)
-#3102 := (iff #3096 #3101)
-#3103 := [rewrite]: #3102
-#3097 := (iff #2472 #3096)
-#3094 := (iff #2469 #3093)
-#3091 := (iff #2466 #3088)
-#3082 := (and #2500 #3079)
-#3089 := (iff #3082 #3088)
-#3090 := [rewrite]: #3089
-#3083 := (iff #2466 #3082)
-#3080 := (iff #2463 #3079)
-#3077 := (iff #2460 #3074)
-#3068 := (and #2545 #3065)
-#3075 := (iff #3068 #3074)
-#3076 := [rewrite]: #3075
-#3069 := (iff #2460 #3068)
-#3066 := (iff #2457 #3065)
-#3063 := (iff #2454 #3060)
-#3051 := (and #2598 #111 #1397 #2621 #2643 #2671 #3048)
-#3061 := (iff #3051 #3060)
-#3062 := [rewrite]: #3061
-#3052 := (iff #2454 #3051)
-#3049 := (iff #2451 #3048)
-#3046 := (iff #2446 #3043)
-#3036 := (and #1875 #1880 #184 #1323 #1315 #194 #2841 #2847 #3033)
-#3044 := (iff #3036 #3043)
-#3045 := [rewrite]: #3044
-#3037 := (iff #2446 #3036)
-#3034 := (iff #2440 #3033)
-#3031 := (iff #2437 #3028)
-#3023 := (and #1170 #3020)
-#3029 := (iff #3023 #3028)
-#3030 := [rewrite]: #3029
-#3024 := (iff #2437 #3023)
-#3021 := (iff #2434 #3020)
-#3018 := (iff #2431 #3015)
-#3010 := (and #563 #3007)
-#3016 := (iff #3010 #3015)
-#3017 := [rewrite]: #3016
-#3011 := (iff #2431 #3010)
-#3008 := (iff #2428 #3007)
-#3005 := (iff #2425 #3002)
-#2997 := (and #217 #2994)
-#3003 := (iff #2997 #3002)
-#3004 := [rewrite]: #3003
-#2998 := (iff #2425 #2997)
-#2995 := (iff #2422 #2994)
-#2992 := (iff #2419 #2989)
-#2984 := (and #1179 #2981)
-#2990 := (iff #2984 #2989)
-#2991 := [rewrite]: #2990
-#2985 := (iff #2419 #2984)
-#2982 := (iff #2416 #2981)
-#2979 := (iff #2413 #2976)
-#2970 := (and #2893 #2967)
-#2977 := (iff #2970 #2976)
-#2978 := [rewrite]: #2977
-#2971 := (iff #2413 #2970)
-#2968 := (iff #2410 #2967)
-#2965 := (iff #2405 #2962)
-#2956 := (and #2938 #2036 #2041 #2953)
-#2963 := (iff #2956 #2962)
-#2964 := [rewrite]: #2963
-#2957 := (iff #2405 #2956)
-#2954 := (iff #2396 #2953)
-#2951 := (iff #2393 #2942)
-#2943 := (not #2942)
-#2946 := (not #2943)
-#2949 := (iff #2946 #2942)
-#2950 := [rewrite]: #2949
-#2947 := (iff #2393 #2946)
-#2944 := (iff #2390 #2943)
-#2945 := [rewrite]: #2944
-#2948 := [monotonicity #2945]: #2947
-#2952 := [trans #2948 #2950]: #2951
-#2955 := [quant-intro #2952]: #2954
-#2939 := (iff #1212 #2938)
-#2936 := (iff #1209 #2933)
-#2919 := (or #220 #940)
-#2930 := (or #2919 #1205)
-#2934 := (iff #2930 #2933)
-#2935 := [rewrite]: #2934
-#2931 := (iff #1209 #2930)
-#2928 := (iff #1202 #2919)
-#2920 := (not #2919)
-#2923 := (not #2920)
-#2926 := (iff #2923 #2919)
-#2927 := [rewrite]: #2926
-#2924 := (iff #1202 #2923)
-#2921 := (iff #1199 #2920)
-#2922 := [rewrite]: #2921
-#2925 := [monotonicity #2922]: #2924
-#2929 := [trans #2925 #2927]: #2928
-#2932 := [monotonicity #2929]: #2931
-#2937 := [trans #2932 #2935]: #2936
-#2940 := [quant-intro #2937]: #2939
-#2958 := [monotonicity #2940 #2955]: #2957
-#2966 := [trans #2958 #2964]: #2965
-#2917 := (iff #2022 #2916)
-#2914 := (iff #2021 #2911)
-#2897 := (or #2896 #2011)
-#2908 := (or #2897 #2020)
-#2912 := (iff #2908 #2911)
-#2913 := [rewrite]: #2912
-#2909 := (iff #2021 #2908)
-#2906 := (iff #2014 #2897)
-#2898 := (not #2897)
-#2901 := (not #2898)
-#2904 := (iff #2901 #2897)
-#2905 := [rewrite]: #2904
-#2902 := (iff #2014 #2901)
-#2899 := (iff #2013 #2898)
-#2900 := [rewrite]: #2899
-#2903 := [monotonicity #2900]: #2902
-#2907 := [trans #2903 #2905]: #2906
-#2910 := [monotonicity #2907]: #2909
-#2915 := [trans #2910 #2913]: #2914
-#2918 := [monotonicity #2915]: #2917
-#2969 := [monotonicity #2918 #2966]: #2968
-#2894 := (iff #1193 #2893)
-#2891 := (iff #1190 #2888)
-#2874 := (or #213 #2873)
-#2885 := (or #2874 #1185)
-#2889 := (iff #2885 #2888)
-#2890 := [rewrite]: #2889
-#2886 := (iff #1190 #2885)
-#2883 := (iff #566 #2874)
-#2875 := (not #2874)
-#2878 := (not #2875)
-#2881 := (iff #2878 #2874)
-#2882 := [rewrite]: #2881
-#2879 := (iff #566 #2878)
-#2876 := (iff #223 #2875)
-#2877 := [rewrite]: #2876
-#2880 := [monotonicity #2877]: #2879
-#2884 := [trans #2880 #2882]: #2883
-#2887 := [monotonicity #2884]: #2886
-#2892 := [trans #2887 #2890]: #2891
-#2895 := [quant-intro #2892]: #2894
-#2972 := [monotonicity #2895 #2969]: #2971
-#2980 := [trans #2972 #2978]: #2979
-#2871 := (iff #1990 #2870)
-#2868 := (iff #1989 #2865)
-#2851 := (or #1978 #2850)
-#2862 := (or #2851 #1988)
-#2866 := (iff #2862 #2865)
-#2867 := [rewrite]: #2866
-#2863 := (iff #1989 #2862)
-#2860 := (iff #1983 #2851)
-#2852 := (not #2851)
-#2855 := (not #2852)
-#2858 := (iff #2855 #2851)
-#2859 := [rewrite]: #2858
-#2856 := (iff #1983 #2855)
-#2853 := (iff #1982 #2852)
-#2854 := [rewrite]: #2853
-#2857 := [monotonicity #2854]: #2856
-#2861 := [trans #2857 #2859]: #2860
-#2864 := [monotonicity #2861]: #2863
-#2869 := [trans #2864 #2867]: #2868
-#2872 := [monotonicity #2869]: #2871
-#2983 := [monotonicity #2872 #2980]: #2982
-#2986 := [monotonicity #2983]: #2985
-#2993 := [trans #2986 #2991]: #2992
-#2996 := [monotonicity #2993]: #2995
-#2999 := [monotonicity #2996]: #2998
-#3006 := [trans #2999 #3004]: #3005
-#3009 := [monotonicity #3006]: #3008
-#3012 := [monotonicity #3009]: #3011
-#3019 := [trans #3012 #3017]: #3018
-#3022 := [monotonicity #3019]: #3021
-#3025 := [monotonicity #3022]: #3024
-#3032 := [trans #3025 #3030]: #3031
-#3035 := [monotonicity #3032]: #3034
-#2848 := (iff #1159 #2847)
-#2845 := (iff #1156 #2844)
-#2824 := (iff #1153 #2823)
-#2825 := [rewrite]: #2824
-#2846 := [monotonicity #2825]: #2845
-#2849 := [quant-intro #2846]: #2848
-#2842 := (iff #1302 #2841)
-#2839 := (iff #1299 #2836)
-#2833 := (or #2822 #1294)
-#2837 := (iff #2833 #2836)
-#2838 := [rewrite]: #2837
-#2834 := (iff #1299 #2833)
-#2831 := (iff #1291 #2822)
-#2826 := (not #2823)
-#2829 := (iff #2826 #2822)
-#2830 := [rewrite]: #2829
-#2827 := (iff #1291 #2826)
-#2828 := [monotonicity #2825]: #2827
-#2832 := [trans #2828 #2830]: #2831
-#2835 := [monotonicity #2832]: #2834
-#2840 := [trans #2835 #2838]: #2839
-#2843 := [quant-intro #2840]: #2842
-#3038 := [monotonicity #2843 #2849 #3035]: #3037
-#3047 := [trans #3038 #3045]: #3046
-#2820 := (iff #2374 #2817)
-#2811 := (and #2685 #144 #147 #149 #152 #2808)
-#2818 := (iff #2811 #2817)
-#2819 := [rewrite]: #2818
-#2812 := (iff #2374 #2811)
-#2809 := (iff #2371 #2808)
-#2806 := (iff #2368 #2803)
-#2797 := (and #2737 #2794)
-#2804 := (iff #2797 #2803)
-#2805 := [rewrite]: #2804
-#2798 := (iff #2368 #2797)
-#2795 := (iff #2365 #2794)
-#2792 := (iff #1852 #2789)
-#2784 := (and #2781 #1849)
-#2790 := (iff #2784 #2789)
-#2791 := [rewrite]: #2790
-#2785 := (iff #1852 #2784)
-#2782 := (iff #1096 #2781)
-#2779 := (iff #1093 #2776)
-#2762 := (or #1049 #940)
-#2773 := (or #2762 #1090)
-#2777 := (iff #2773 #2776)
-#2778 := [rewrite]: #2777
-#2774 := (iff #1093 #2773)
-#2771 := (iff #1087 #2762)
-#2763 := (not #2762)
-#2766 := (not #2763)
-#2769 := (iff #2766 #2762)
-#2770 := [rewrite]: #2769
-#2767 := (iff #1087 #2766)
-#2764 := (iff #1084 #2763)
-#2765 := [rewrite]: #2764
-#2768 := [monotonicity #2765]: #2767
-#2772 := [trans #2768 #2770]: #2771
-#2775 := [monotonicity #2772]: #2774
-#2780 := [trans #2775 #2778]: #2779
-#2783 := [quant-intro #2780]: #2782
-#2786 := [monotonicity #2783]: #2785
-#2793 := [trans #2786 #2791]: #2792
-#2760 := (iff #2362 #2759)
-#2757 := (iff #2359 #2754)
-#2740 := (or #1820 #1827)
-#2751 := (or #2740 #2356)
-#2755 := (iff #2751 #2754)
-#2756 := [rewrite]: #2755
-#2752 := (iff #2359 #2751)
-#2749 := (iff #1830 #2740)
-#2741 := (not #2740)
-#2744 := (not #2741)
-#2747 := (iff #2744 #2740)
-#2748 := [rewrite]: #2747
-#2745 := (iff #1830 #2744)
-#2742 := (iff #1829 #2741)
-#2743 := [rewrite]: #2742
-#2746 := [monotonicity #2743]: #2745
-#2750 := [trans #2746 #2748]: #2749
-#2753 := [monotonicity #2750]: #2752
-#2758 := [trans #2753 #2756]: #2757
-#2761 := [monotonicity #2758]: #2760
-#2796 := [monotonicity #2761 #2793]: #2795
-#2738 := (iff #2349 #2737)
-#2735 := (iff #2346 #2732)
-#2713 := (or #66 #1049)
-#2729 := (or #2713 #2726)
-#2733 := (iff #2729 #2732)
-#2734 := [rewrite]: #2733
-#2730 := (iff #2346 #2729)
-#2727 := (iff #2343 #2726)
-#2728 := [rewrite]: #2727
-#2722 := (iff #1056 #2713)
-#2714 := (not #2713)
-#2717 := (not #2714)
-#2720 := (iff #2717 #2713)
-#2721 := [rewrite]: #2720
-#2718 := (iff #1056 #2717)
-#2715 := (iff #1053 #2714)
-#2716 := [rewrite]: #2715
-#2719 := [monotonicity #2716]: #2718
-#2723 := [trans #2719 #2721]: #2722
-#2731 := [monotonicity #2723 #2728]: #2730
-#2736 := [trans #2731 #2734]: #2735
-#2739 := [quant-intro #2736]: #2738
-#2799 := [monotonicity #2739 #2796]: #2798
-#2807 := [trans #2799 #2805]: #2806
-#2711 := (iff #2307 #2708)
-#2703 := (and #1753 #1758 #2700)
-#2709 := (iff #2703 #2708)
-#2710 := [rewrite]: #2709
-#2704 := (iff #2307 #2703)
-#2701 := (iff #2301 #2700)
-#2698 := (iff #2298 #2689)
-#2690 := (not #2689)
-#2693 := (not #2690)
-#2696 := (iff #2693 #2689)
-#2697 := [rewrite]: #2696
-#2694 := (iff #2298 #2693)
-#2691 := (iff #2295 #2690)
-#2692 := [rewrite]: #2691
-#2695 := [monotonicity #2692]: #2694
-#2699 := [trans #2695 #2697]: #2698
-#2702 := [quant-intro #2699]: #2701
-#2705 := [monotonicity #2702]: #2704
-#2712 := [trans #2705 #2710]: #2711
-#2810 := [monotonicity #2712 #2807]: #2809
-#2686 := (iff #1736 #2685)
-#2683 := (iff #1733 #2674)
-#2675 := (not #2674)
-#2678 := (not #2675)
-#2681 := (iff #2678 #2674)
-#2682 := [rewrite]: #2681
-#2679 := (iff #1733 #2678)
-#2676 := (iff #1114 #2675)
-#2677 := [rewrite]: #2676
-#2680 := [monotonicity #2677]: #2679
-#2684 := [trans #2680 #2682]: #2683
-#2687 := [quant-intro #2684]: #2686
-#2813 := [monotonicity #2687 #2810]: #2812
-#2821 := [trans #2813 #2819]: #2820
-#3050 := [monotonicity #2821 #3047]: #3049
-#2672 := (iff #2275 #2671)
-#2669 := (iff #2272 #2666)
-#2646 := (or #66 #1004)
-#2663 := (or #2646 #2660)
-#2667 := (iff #2663 #2666)
-#2668 := [rewrite]: #2667
-#2664 := (iff #2272 #2663)
-#2661 := (iff #2269 #2660)
-#2662 := [rewrite]: #2661
-#2655 := (iff #1011 #2646)
-#2647 := (not #2646)
-#2650 := (not #2647)
-#2653 := (iff #2650 #2646)
-#2654 := [rewrite]: #2653
-#2651 := (iff #1011 #2650)
-#2648 := (iff #1008 #2647)
-#2649 := [rewrite]: #2648
-#2652 := [monotonicity #2649]: #2651
-#2656 := [trans #2652 #2654]: #2655
-#2665 := [monotonicity #2656 #2662]: #2664
-#2670 := [trans #2665 #2668]: #2669
-#2673 := [quant-intro #2670]: #2672
-#2644 := (iff #1377 #2643)
-#2641 := (iff #1374 #2638)
-#2624 := (or #118 #940)
-#2635 := (or #2624 #1371)
-#2639 := (iff #2635 #2638)
-#2640 := [rewrite]: #2639
-#2636 := (iff #1374 #2635)
-#2633 := (iff #1368 #2624)
-#2625 := (not #2624)
-#2628 := (not #2625)
-#2631 := (iff #2628 #2624)
-#2632 := [rewrite]: #2631
-#2629 := (iff #1368 #2628)
-#2626 := (iff #1365 #2625)
-#2627 := [rewrite]: #2626
-#2630 := [monotonicity #2627]: #2629
-#2634 := [trans #2630 #2632]: #2633
-#2637 := [monotonicity #2634]: #2636
-#2642 := [trans #2637 #2640]: #2641
-#2645 := [quant-intro #2642]: #2644
-#2622 := (iff #1388 #2621)
-#2619 := (iff #1385 #2616)
-#2602 := (or #117 #2601)
-#2613 := (or #2602 #1020)
-#2617 := (iff #2613 #2616)
-#2618 := [rewrite]: #2617
-#2614 := (iff #1385 #2613)
-#2611 := (iff #377 #2602)
-#2603 := (not #2602)
-#2606 := (not #2603)
-#2609 := (iff #2606 #2602)
-#2610 := [rewrite]: #2609
-#2607 := (iff #377 #2606)
-#2604 := (iff #121 #2603)
-#2605 := [rewrite]: #2604
-#2608 := [monotonicity #2605]: #2607
-#2612 := [trans #2608 #2610]: #2611
-#2615 := [monotonicity #2612]: #2614
-#2620 := [trans #2615 #2618]: #2619
-#2623 := [quant-intro #2620]: #2622
-#2599 := (iff #2233 #2598)
-#2596 := (iff #2230 #2593)
-#2573 := (or #66 #966)
-#2590 := (or #2573 #2587)
-#2594 := (iff #2590 #2593)
-#2595 := [rewrite]: #2594
-#2591 := (iff #2230 #2590)
-#2588 := (iff #2227 #2587)
-#2589 := [rewrite]: #2588
-#2582 := (iff #973 #2573)
-#2574 := (not #2573)
-#2577 := (not #2574)
-#2580 := (iff #2577 #2573)
-#2581 := [rewrite]: #2580
-#2578 := (iff #973 #2577)
-#2575 := (iff #970 #2574)
-#2576 := [rewrite]: #2575
-#2579 := [monotonicity #2576]: #2578
-#2583 := [trans #2579 #2581]: #2582
-#2592 := [monotonicity #2583 #2589]: #2591
-#2597 := [trans #2592 #2595]: #2596
-#2600 := [quant-intro #2597]: #2599
-#3053 := [monotonicity #2600 #2623 #2645 #2673 #3050]: #3052
-#3064 := [trans #3053 #3062]: #3063
-#2571 := (iff #2191 #2568)
-#2563 := (and #1615 #1620 #2560)
-#2569 := (iff #2563 #2568)
-#2570 := [rewrite]: #2569
-#2564 := (iff #2191 #2563)
-#2561 := (iff #1643 #2560)
-#2558 := (iff #1640 #2549)
-#2550 := (not #2549)
-#2553 := (not #2550)
-#2556 := (iff #2553 #2549)
-#2557 := [rewrite]: #2556
-#2554 := (iff #1640 #2553)
-#2551 := (iff #1631 #2550)
-#2552 := [rewrite]: #2551
-#2555 := [monotonicity #2552]: #2554
-#2559 := [trans #2555 #2557]: #2558
-#2562 := [quant-intro #2559]: #2561
-#2565 := [monotonicity #2562]: #2564
-#2572 := [trans #2565 #2570]: #2571
-#3067 := [monotonicity #2572 #3064]: #3066
-#2546 := (iff #958 #2545)
-#2543 := (iff #955 #2540)
-#2526 := (or #78 #940)
-#2537 := (or #2526 #951)
-#2541 := (iff #2537 #2540)
-#2542 := [rewrite]: #2541
-#2538 := (iff #955 #2537)
-#2535 := (iff #947 #2526)
-#2527 := (not #2526)
-#2530 := (not #2527)
-#2533 := (iff #2530 #2526)
-#2534 := [rewrite]: #2533
-#2531 := (iff #947 #2530)
-#2528 := (iff #944 #2527)
-#2529 := [rewrite]: #2528
-#2532 := [monotonicity #2529]: #2531
-#2536 := [trans #2532 #2534]: #2535
-#2539 := [monotonicity #2536]: #2538
-#2544 := [trans #2539 #2542]: #2543
-#2547 := [quant-intro #2544]: #2546
-#3070 := [monotonicity #2547 #3067]: #3069
-#3078 := [trans #3070 #3076]: #3077
-#2524 := (iff #2183 #2523)
-#2521 := (iff #2180 #2518)
-#2504 := (or #2503 #1590)
-#2515 := (or #2504 #2177)
-#2519 := (iff #2515 #2518)
-#2520 := [rewrite]: #2519
-#2516 := (iff #2180 #2515)
-#2513 := (iff #1593 #2504)
-#2505 := (not #2504)
-#2508 := (not #2505)
-#2511 := (iff #2508 #2504)
-#2512 := [rewrite]: #2511
-#2509 := (iff #1593 #2508)
-#2506 := (iff #1592 #2505)
-#2507 := [rewrite]: #2506
-#2510 := [monotonicity #2507]: #2509
-#2514 := [trans #2510 #2512]: #2513
-#2517 := [monotonicity #2514]: #2516
-#2522 := [trans #2517 #2520]: #2521
-#2525 := [monotonicity #2522]: #2524
-#3081 := [monotonicity #2525 #3078]: #3080
-#2501 := (iff #932 #2500)
-#2498 := (iff #929 #2495)
-#1670 := (or #77 #1605)
-#2492 := (or #1670 #925)
-#2496 := (iff #2492 #2495)
-#2497 := [rewrite]: #2496
-#2493 := (iff #929 #2492)
-#2490 := (iff #356 #1670)
-#1671 := (not #1670)
-#2485 := (not #1671)
-#2488 := (iff #2485 #1670)
-#2489 := [rewrite]: #2488
-#2486 := (iff #356 #2485)
-#2170 := (iff #86 #1671)
-#2484 := [rewrite]: #2170
-#2487 := [monotonicity #2484]: #2486
-#2491 := [trans #2487 #2489]: #2490
-#2494 := [monotonicity #2491]: #2493
-#2499 := [trans #2494 #2497]: #2498
-#2502 := [quant-intro #2499]: #2501
-#3084 := [monotonicity #2502 #3081]: #3083
-#3092 := [trans #3084 #3090]: #3091
-#1947 := (iff #1569 #1946)
-#1966 := (iff #1568 #1841)
-#1885 := (or #1557 #1884)
-#1545 := (or #1885 #1567)
-#1926 := (iff #1545 #1841)
-#1927 := [rewrite]: #1926
-#1546 := (iff #1568 #1545)
-#1572 := (iff #1562 #1885)
-#1804 := (not #1885)
-#2026 := (not #1804)
-#1993 := (iff #2026 #1885)
-#1994 := [rewrite]: #1993
-#1722 := (iff #1562 #2026)
-#1805 := (iff #1561 #1804)
-#2025 := [rewrite]: #1805
-#1723 := [monotonicity #2025]: #1722
-#1573 := [trans #1723 #1994]: #1572
-#1840 := [monotonicity #1573]: #1546
-#1967 := [trans #1840 #1927]: #1966
-#1604 := [monotonicity #1967]: #1947
-#3095 := [monotonicity #1604 #3092]: #3094
-#3098 := [monotonicity #3095]: #3097
-#3105 := [trans #3098 #3103]: #3104
-#3108 := [monotonicity #3105]: #3107
-#3111 := [monotonicity #3108]: #3110
-#3118 := [trans #3111 #3116]: #3117
-#3121 := [monotonicity #3118]: #3120
-#2049 := (+ #2048 #2044)
-#2050 := (= #2049 0::Int)
-#2051 := (and #2046 #213 #2050)
-#2061 := (not #2051)
-#2064 := (forall (vars (?v1 S2)) #2061)
-#2042 := (and #2036 #2041)
-#2043 := (not #2042)
-#2058 := (not #2043)
-#2068 := (and #2058 #2064)
-#2073 := (and #1212 #2068)
-#2077 := (or #2022 #2073)
-#2081 := (and #1193 #2077)
-#2085 := (or #1990 #2081)
-#2089 := (and #1179 #2085)
-#2093 := (or #1963 #2089)
-#1957 := (not #661)
-#2097 := (and #1957 #2093)
-#2101 := (or #661 #2097)
-#2105 := (and #563 #2101)
-#2109 := (or #1943 #2105)
-#2113 := (and #1170 #2109)
-#2117 := (or #1923 #2113)
-#1900 := (not #715)
-#1881 := (and #1875 #1880)
-#2121 := (and #1881 #184 #1323 #1315 #1900 #1302 #1159 #2117)
-#1833 := (+ #1817 #1832)
-#1834 := (+ #1824 #1833)
-#1835 := (>= #1834 0::Int)
-#1836 := (or #1830 #1835)
-#1837 := (not #1836)
-#1856 := (or #1837 #1852)
-#1793 := (+ #1792 #1047)
-#1799 := (+ #1798 #1793)
-#1800 := (= #1799 0::Int)
-#1794 := (>= #1793 0::Int)
-#1795 := (not #1794)
-#1801 := (and #1795 #1800)
-#1806 := (or #1056 #1801)
-#1809 := (forall (vars (?v0 S2)) #1806)
-#1860 := (and #1809 #1856)
-#1766 := (+ #1765 #1761)
-#1767 := (= #1766 0::Int)
-#1768 := (and #1763 #1767)
-#1777 := (not #1768)
-#1780 := (forall (vars (?v1 S2)) #1777)
-#1759 := (and #1753 #1758)
-#1760 := (not #1759)
-#1774 := (not #1760)
-#1784 := (and #1774 #1780)
-#1864 := (or #1784 #1860)
-#1748 := (not #494)
-#1745 := (not #503)
-#1742 := (not #512)
-#1739 := (not #521)
-#1868 := (and #1736 #1739 #1742 #1745 #1748 #1864)
-#2125 := (or #1868 #2121)
-#1709 := (+ #1708 #1002)
-#1717 := (+ #1716 #1709)
-#1718 := (= #1717 0::Int)
-#1710 := (>= #1709 0::Int)
-#1711 := (not #1710)
-#1719 := (and #1711 #1713 #1718)
-#1724 := (or #1011 #1719)
-#1727 := (forall (vars (?v0 S2)) #1724)
-#1681 := (not #804)
-#1664 := (+ #964 #1663)
-#1665 := (+ #1655 #1664)
-#1666 := (= #1665 0::Int)
-#1656 := (+ #1655 #964)
-#1657 := (>= #1656 0::Int)
-#1658 := (not #1657)
-#1667 := (and #1658 #1660 #1666)
-#1672 := (or #973 #1667)
-#1675 := (forall (vars (?v0 S2)) #1672)
-#2129 := (and #1675 #1681 #1397 #1388 #1377 #1727 #2125)
-#1621 := (and #1615 #1620)
-#1622 := (not #1621)
-#1637 := (not #1622)
-#1647 := (and #1637 #1643)
-#2133 := (or #1647 #2129)
-#2137 := (and #958 #2133)
-#1597 := (+ #1596 #1587)
-#1598 := (+ #1594 #1597)
-#1599 := (>= #1598 0::Int)
-#1600 := (or #1593 #1599)
-#1601 := (not #1600)
-#2141 := (or #1601 #2137)
-#2145 := (and #932 #2141)
-#2149 := (or #1569 #2145)
-#2153 := (and #917 #2149)
-#2157 := (or #1542 #2153)
-#1504 := (not #868)
-#2161 := (and #1504 #2157)
-#2165 := (or #868 #2161)
-#2482 := (iff #2165 #2481)
-#2479 := (iff #2161 #2478)
-#2476 := (iff #2157 #2475)
-#2473 := (iff #2153 #2472)
-#2470 := (iff #2149 #2469)
-#2467 := (iff #2145 #2466)
-#2464 := (iff #2141 #2463)
-#2461 := (iff #2137 #2460)
-#2458 := (iff #2133 #2457)
-#2455 := (iff #2129 #2454)
-#2452 := (iff #2125 #2451)
-#2449 := (iff #2121 #2446)
-#2443 := (and #1881 #184 #1323 #1315 #194 #1302 #1159 #2440)
-#2447 := (iff #2443 #2446)
-#2448 := [rewrite]: #2447
-#2444 := (iff #2121 #2443)
-#2441 := (iff #2117 #2440)
-#2438 := (iff #2113 #2437)
-#2435 := (iff #2109 #2434)
-#2432 := (iff #2105 #2431)
-#2429 := (iff #2101 #2428)
-#2426 := (iff #2097 #2425)
-#2423 := (iff #2093 #2422)
-#2420 := (iff #2089 #2419)
-#2417 := (iff #2085 #2416)
-#2414 := (iff #2081 #2413)
-#2411 := (iff #2077 #2410)
-#2408 := (iff #2073 #2405)
-#2399 := (and #2042 #2396)
-#2402 := (and #1212 #2399)
-#2406 := (iff #2402 #2405)
-#2407 := [rewrite]: #2406
-#2403 := (iff #2073 #2402)
-#2400 := (iff #2068 #2399)
-#2397 := (iff #2064 #2396)
-#2394 := (iff #2061 #2393)
-#2391 := (iff #2051 #2390)
-#2388 := (iff #2050 #2387)
-#2385 := (= #2049 #2384)
-#2386 := [rewrite]: #2385
-#2389 := [monotonicity #2386]: #2388
-#2392 := [monotonicity #2389]: #2391
-#2395 := [monotonicity #2392]: #2394
-#2398 := [quant-intro #2395]: #2397
-#2381 := (iff #2058 #2042)
-#2382 := [rewrite]: #2381
-#2401 := [monotonicity #2382 #2398]: #2400
-#2404 := [monotonicity #2401]: #2403
-#2409 := [trans #2404 #2407]: #2408
-#2412 := [monotonicity #2409]: #2411
-#2415 := [monotonicity #2412]: #2414
-#2418 := [monotonicity #2415]: #2417
-#2421 := [monotonicity #2418]: #2420
-#2424 := [monotonicity #2421]: #2423
-#2379 := (iff #1957 #217)
-#2380 := [rewrite]: #2379
-#2427 := [monotonicity #2380 #2424]: #2426
-#2430 := [monotonicity #2427]: #2429
-#2433 := [monotonicity #2430]: #2432
-#2436 := [monotonicity #2433]: #2435
-#2439 := [monotonicity #2436]: #2438
-#2442 := [monotonicity #2439]: #2441
-#2377 := (iff #1900 #194)
-#2378 := [rewrite]: #2377
-#2445 := [monotonicity #2378 #2442]: #2444
-#2450 := [trans #2445 #2448]: #2449
-#2375 := (iff #1868 #2374)
-#2372 := (iff #1864 #2371)
-#2369 := (iff #1860 #2368)
-#2366 := (iff #1856 #2365)
-#2363 := (iff #1837 #2362)
-#2360 := (iff #1836 #2359)
-#2357 := (iff #1835 #2356)
-#2354 := (= #1834 #2353)
-#2355 := [rewrite]: #2354
-#2358 := [monotonicity #2355]: #2357
-#2361 := [monotonicity #2358]: #2360
-#2364 := [monotonicity #2361]: #2363
-#2367 := [monotonicity #2364]: #2366
-#2350 := (iff #1809 #2349)
-#2347 := (iff #1806 #2346)
-#2344 := (iff #1801 #2343)
-#2341 := (iff #1800 #2338)
-#2328 := (+ #1792 #1798)
-#2329 := (+ #1047 #2328)
-#2332 := (= #2329 0::Int)
-#2339 := (iff #2332 #2338)
-#2340 := [rewrite]: #2339
-#2333 := (iff #1800 #2332)
-#2330 := (= #1799 #2329)
-#2331 := [rewrite]: #2330
-#2334 := [monotonicity #2331]: #2333
-#2342 := [trans #2334 #2340]: #2341
-#2326 := (iff #1795 #2325)
-#2323 := (iff #1794 #2320)
-#2312 := (+ #1047 #1792)
-#2315 := (>= #2312 0::Int)
-#2321 := (iff #2315 #2320)
-#2322 := [rewrite]: #2321
-#2316 := (iff #1794 #2315)
-#2313 := (= #1793 #2312)
-#2314 := [rewrite]: #2313
-#2317 := [monotonicity #2314]: #2316
-#2324 := [trans #2317 #2322]: #2323
-#2327 := [monotonicity #2324]: #2326
-#2345 := [monotonicity #2327 #2342]: #2344
-#2348 := [monotonicity #2345]: #2347
-#2351 := [quant-intro #2348]: #2350
-#2370 := [monotonicity #2351 #2367]: #2369
-#2310 := (iff #1784 #2307)
-#2304 := (and #1759 #2301)
-#2308 := (iff #2304 #2307)
-#2309 := [rewrite]: #2308
-#2305 := (iff #1784 #2304)
-#2302 := (iff #1780 #2301)
-#2299 := (iff #1777 #2298)
-#2296 := (iff #1768 #2295)
-#2293 := (iff #1767 #2292)
-#2290 := (= #1766 #2289)
-#2291 := [rewrite]: #2290
-#2294 := [monotonicity #2291]: #2293
-#2297 := [monotonicity #2294]: #2296
-#2300 := [monotonicity #2297]: #2299
-#2303 := [quant-intro #2300]: #2302
-#2286 := (iff #1774 #1759)
-#2287 := [rewrite]: #2286
-#2306 := [monotonicity #2287 #2303]: #2305
-#2311 := [trans #2306 #2309]: #2310
-#2373 := [monotonicity #2311 #2370]: #2372
-#2284 := (iff #1748 #152)
-#2285 := [rewrite]: #2284
-#2282 := (iff #1745 #149)
-#2283 := [rewrite]: #2282
-#2280 := (iff #1742 #147)
-#2281 := [rewrite]: #2280
-#2278 := (iff #1739 #144)
-#2279 := [rewrite]: #2278
-#2376 := [monotonicity #2279 #2281 #2283 #2285 #2373]: #2375
-#2453 := [monotonicity #2376 #2450]: #2452
-#2276 := (iff #1727 #2275)
-#2273 := (iff #1724 #2272)
-#2270 := (iff #1719 #2269)
-#2267 := (iff #1718 #2264)
-#2254 := (+ #1708 #1716)
-#2255 := (+ #1002 #2254)
-#2258 := (= #2255 0::Int)
-#2265 := (iff #2258 #2264)
-#2266 := [rewrite]: #2265
-#2259 := (iff #1718 #2258)
-#2256 := (= #1717 #2255)
-#2257 := [rewrite]: #2256
-#2260 := [monotonicity #2257]: #2259
-#2268 := [trans #2260 #2266]: #2267
-#2252 := (iff #1711 #2251)
-#2249 := (iff #1710 #2246)
-#2238 := (+ #1002 #1708)
-#2241 := (>= #2238 0::Int)
-#2247 := (iff #2241 #2246)
-#2248 := [rewrite]: #2247
-#2242 := (iff #1710 #2241)
-#2239 := (= #1709 #2238)
-#2240 := [rewrite]: #2239
-#2243 := [monotonicity #2240]: #2242
-#2250 := [trans #2243 #2248]: #2249
-#2253 := [monotonicity #2250]: #2252
-#2271 := [monotonicity #2253 #2268]: #2270
-#2274 := [monotonicity #2271]: #2273
-#2277 := [quant-intro #2274]: #2276
-#2236 := (iff #1681 #111)
-#2237 := [rewrite]: #2236
-#2234 := (iff #1675 #2233)
-#2231 := (iff #1672 #2230)
-#2228 := (iff #1667 #2227)
-#2225 := (iff #1666 #2222)
-#2212 := (+ #1655 #1663)
-#2213 := (+ #964 #2212)
-#2216 := (= #2213 0::Int)
-#2223 := (iff #2216 #2222)
-#2224 := [rewrite]: #2223
-#2217 := (iff #1666 #2216)
-#2214 := (= #1665 #2213)
-#2215 := [rewrite]: #2214
-#2218 := [monotonicity #2215]: #2217
-#2226 := [trans #2218 #2224]: #2225
-#2210 := (iff #1658 #2209)
-#2207 := (iff #1657 #2204)
-#2196 := (+ #964 #1655)
-#2199 := (>= #2196 0::Int)
-#2205 := (iff #2199 #2204)
-#2206 := [rewrite]: #2205
-#2200 := (iff #1657 #2199)
-#2197 := (= #1656 #2196)
-#2198 := [rewrite]: #2197
-#2201 := [monotonicity #2198]: #2200
-#2208 := [trans #2201 #2206]: #2207
-#2211 := [monotonicity #2208]: #2210
-#2229 := [monotonicity #2211 #2226]: #2228
-#2232 := [monotonicity #2229]: #2231
-#2235 := [quant-intro #2232]: #2234
-#2456 := [monotonicity #2235 #2237 #2277 #2453]: #2455
-#2194 := (iff #1647 #2191)
-#2188 := (and #1621 #1643)
-#2192 := (iff #2188 #2191)
-#2193 := [rewrite]: #2192
-#2189 := (iff #1647 #2188)
-#2186 := (iff #1637 #1621)
-#2187 := [rewrite]: #2186
-#2190 := [monotonicity #2187]: #2189
-#2195 := [trans #2190 #2193]: #2194
-#2459 := [monotonicity #2195 #2456]: #2458
-#2462 := [monotonicity #2459]: #2461
-#2184 := (iff #1601 #2183)
-#2181 := (iff #1600 #2180)
-#2178 := (iff #1599 #2177)
-#2175 := (= #1598 #2174)
-#2176 := [rewrite]: #2175
-#2179 := [monotonicity #2176]: #2178
-#2182 := [monotonicity #2179]: #2181
-#2185 := [monotonicity #2182]: #2184
-#2465 := [monotonicity #2185 #2462]: #2464
-#2468 := [monotonicity #2465]: #2467
-#2471 := [monotonicity #2468]: #2470
-#2474 := [monotonicity #2471]: #2473
-#2477 := [monotonicity #2474]: #2476
-#2171 := (iff #1504 #81)
-#2172 := [rewrite]: #2171
-#2480 := [monotonicity #2172 #2477]: #2479
-#2483 := [monotonicity #2480]: #2482
-#1479 := (not #1453)
-#2166 := (~ #1479 #2165)
-#2162 := (not #1450)
-#2163 := (~ #2162 #2161)
-#2158 := (not #1447)
-#2159 := (~ #2158 #2157)
-#2154 := (not #1444)
-#2155 := (~ #2154 #2153)
-#2150 := (not #1441)
-#2151 := (~ #2150 #2149)
-#2146 := (not #1438)
-#2147 := (~ #2146 #2145)
-#2142 := (not #1435)
-#2143 := (~ #2142 #2141)
-#2138 := (not #1432)
-#2139 := (~ #2138 #2137)
-#2134 := (not #1429)
-#2135 := (~ #2134 #2133)
-#2130 := (not #1424)
-#2131 := (~ #2130 #2129)
-#2126 := (not #1362)
-#2127 := (~ #2126 #2125)
-#2122 := (not #1357)
-#2123 := (~ #2122 #2121)
-#2118 := (not #1288)
-#2119 := (~ #2118 #2117)
-#2114 := (not #1285)
-#2115 := (~ #2114 #2113)
-#2110 := (not #1282)
-#2111 := (~ #2110 #2109)
-#2106 := (not #1279)
-#2107 := (~ #2106 #2105)
-#2102 := (not #1276)
-#2103 := (~ #2102 #2101)
-#2098 := (not #1273)
-#2099 := (~ #2098 #2097)
-#2094 := (not #1270)
-#2095 := (~ #2094 #2093)
-#2090 := (not #1267)
-#2091 := (~ #2090 #2089)
-#2086 := (not #1264)
-#2087 := (~ #2086 #2085)
-#2082 := (not #1261)
-#2083 := (~ #2082 #2081)
-#2078 := (not #1258)
-#2079 := (~ #2078 #2077)
-#2074 := (not #1255)
-#2075 := (~ #2074 #2073)
-#2055 := (not #1252)
-#2071 := (~ #2055 #2068)
-#2052 := (exists (vars (?v1 S2)) #2051)
-#2053 := (or #2043 #2052)
-#2054 := (not #2053)
-#2069 := (~ #2054 #2068)
-#2065 := (not #2052)
-#2066 := (~ #2065 #2064)
-#2062 := (~ #2061 #2061)
-#2063 := [refl]: #2062
-#2067 := [nnf-neg #2063]: #2066
-#2059 := (~ #2058 #2058)
-#2060 := [refl]: #2059
-#2070 := [nnf-neg #2060 #2067]: #2069
-#2056 := (~ #2055 #2054)
-#2057 := [sk]: #2056
-#2072 := [trans #2057 #2070]: #2071
-#2031 := (not #1215)
-#2032 := (~ #2031 #1212)
-#2029 := (~ #1212 #1212)
-#2027 := (~ #1209 #1209)
-#2028 := [refl]: #2027
-#2030 := [nnf-pos #2028]: #2029
-#2033 := [nnf-neg #2030]: #2032
-#2076 := [nnf-neg #2033 #2072]: #2075
-#2023 := (~ #1215 #2022)
-#2024 := [sk]: #2023
-#2080 := [nnf-neg #2024 #2076]: #2079
-#1999 := (not #1196)
-#2000 := (~ #1999 #1193)
-#1997 := (~ #1193 #1193)
-#1995 := (~ #1190 #1190)
-#1996 := [refl]: #1995
-#1998 := [nnf-pos #1996]: #1997
-#2001 := [nnf-neg #1998]: #2000
-#2084 := [nnf-neg #2001 #2080]: #2083
-#1991 := (~ #1196 #1990)
-#1992 := [sk]: #1991
-#2088 := [nnf-neg #1992 #2084]: #2087
-#1972 := (not #1182)
-#1973 := (~ #1972 #1179)
-#1970 := (~ #1179 #1179)
-#1968 := (~ #1176 #1176)
-#1969 := [refl]: #1968
-#1971 := [nnf-pos #1969]: #1970
-#1974 := [nnf-neg #1971]: #1973
-#2092 := [nnf-neg #1974 #2088]: #2091
-#1964 := (~ #1182 #1963)
-#1965 := [sk]: #1964
-#2096 := [nnf-neg #1965 #2092]: #2095
-#1958 := (~ #1957 #1957)
-#1959 := [refl]: #1958
-#2100 := [nnf-neg #1959 #2096]: #2099
-#1955 := (~ #661 #661)
-#1956 := [refl]: #1955
-#2104 := [nnf-neg #1956 #2100]: #2103
-#1952 := (not #673)
-#1953 := (~ #1952 #563)
-#1950 := (~ #563 #563)
-#1948 := (~ #560 #560)
-#1949 := [refl]: #1948
-#1951 := [nnf-pos #1949]: #1950
-#1954 := [nnf-neg #1951]: #1953
-#2108 := [nnf-neg #1954 #2104]: #2107
-#1944 := (~ #673 #1943)
-#1945 := [sk]: #1944
-#2112 := [nnf-neg #1945 #2108]: #2111
-#1932 := (not #1173)
-#1933 := (~ #1932 #1170)
-#1930 := (~ #1170 #1170)
-#1928 := (~ #1165 #1165)
-#1929 := [refl]: #1928
-#1931 := [nnf-pos #1929]: #1930
-#1934 := [nnf-neg #1931]: #1933
-#2116 := [nnf-neg #1934 #2112]: #2115
-#1924 := (~ #1173 #1923)
-#1925 := [sk]: #1924
-#2120 := [nnf-neg #1925 #2116]: #2119
-#1914 := (not #1162)
-#1915 := (~ #1914 #1159)
-#1912 := (~ #1159 #1159)
-#1910 := (~ #1156 #1156)
-#1911 := [refl]: #1910
-#1913 := [nnf-pos #1911]: #1912
-#1916 := [nnf-neg #1913]: #1915
-#1907 := (not #1305)
-#1908 := (~ #1907 #1302)
-#1905 := (~ #1302 #1302)
-#1903 := (~ #1299 #1299)
-#1904 := [refl]: #1903
-#1906 := [nnf-pos #1904]: #1905
-#1909 := [nnf-neg #1906]: #1908
-#1901 := (~ #1900 #1900)
-#1902 := [refl]: #1901
-#1897 := (not #1318)
-#1898 := (~ #1897 #1315)
-#1895 := (~ #1315 #1315)
-#1893 := (~ #1312 #1312)
-#1894 := [refl]: #1893
-#1896 := [nnf-pos #1894]: #1895
-#1899 := [nnf-neg #1896]: #1898
-#1891 := (~ #1323 #1323)
-#1892 := [refl]: #1891
-#1889 := (~ #184 #184)
-#1890 := [refl]: #1889
-#1886 := (not #1333)
-#1887 := (~ #1886 #1881)
-#1882 := (~ #1117 #1881)
-#1883 := [sk]: #1882
-#1888 := [nnf-neg #1883]: #1887
-#2124 := [nnf-neg #1888 #1890 #1892 #1899 #1902 #1909 #1916 #2120]: #2123
-#1869 := (not #1135)
-#1870 := (~ #1869 #1868)
-#1865 := (not #1111)
-#1866 := (~ #1865 #1864)
-#1861 := (not #1108)
-#1862 := (~ #1861 #1860)
-#1857 := (not #1105)
-#1858 := (~ #1857 #1856)
-#1853 := (not #1102)
-#1854 := (~ #1853 #1852)
-#1850 := (~ #1849 #1849)
-#1851 := [refl]: #1850
-#1846 := (not #1099)
-#1847 := (~ #1846 #1096)
-#1844 := (~ #1096 #1096)
-#1842 := (~ #1093 #1093)
-#1843 := [refl]: #1842
-#1845 := [nnf-pos #1843]: #1844
-#1848 := [nnf-neg #1845]: #1847
-#1855 := [nnf-neg #1848 #1851]: #1854
-#1838 := (~ #1099 #1837)
-#1839 := [sk]: #1838
-#1859 := [nnf-neg #1839 #1855]: #1858
-#1812 := (not #1081)
-#1813 := (~ #1812 #1809)
-#1810 := (~ #1078 #1809)
-#1807 := (~ #1075 #1806)
-#1802 := (~ #1072 #1801)
-#1803 := [sk]: #1802
-#1789 := (~ #1056 #1056)
-#1790 := [refl]: #1789
-#1808 := [monotonicity #1790 #1803]: #1807
-#1811 := [nnf-pos #1808]: #1810
-#1814 := [nnf-neg #1811]: #1813
-#1863 := [nnf-neg #1814 #1859]: #1862
-#1787 := (~ #1081 #1784)
-#1769 := (exists (vars (?v1 S2)) #1768)
-#1770 := (or #1760 #1769)
-#1771 := (not #1770)
-#1785 := (~ #1771 #1784)
-#1781 := (not #1769)
-#1782 := (~ #1781 #1780)
-#1778 := (~ #1777 #1777)
-#1779 := [refl]: #1778
-#1783 := [nnf-neg #1779]: #1782
-#1775 := (~ #1774 #1774)
-#1776 := [refl]: #1775
-#1786 := [nnf-neg #1776 #1783]: #1785
-#1772 := (~ #1081 #1771)
-#1773 := [sk]: #1772
-#1788 := [trans #1773 #1786]: #1787
-#1867 := [nnf-neg #1788 #1863]: #1866
-#1749 := (~ #1748 #1748)
-#1750 := [refl]: #1749
-#1746 := (~ #1745 #1745)
-#1747 := [refl]: #1746
-#1743 := (~ #1742 #1742)
-#1744 := [refl]: #1743
-#1740 := (~ #1739 #1739)
-#1741 := [refl]: #1740
-#1737 := (~ #1333 #1736)
-#1734 := (~ #1733 #1733)
-#1735 := [refl]: #1734
-#1738 := [nnf-neg #1735]: #1737
-#1871 := [nnf-neg #1738 #1741 #1744 #1747 #1750 #1867]: #1870
-#2128 := [nnf-neg #1871 #2124]: #2127
-#1730 := (not #1044)
-#1731 := (~ #1730 #1727)
-#1728 := (~ #1041 #1727)
-#1725 := (~ #1038 #1724)
-#1720 := (~ #1035 #1719)
-#1721 := [sk]: #1720
-#1705 := (~ #1011 #1011)
-#1706 := [refl]: #1705
-#1726 := [monotonicity #1706 #1721]: #1725
-#1729 := [nnf-pos #1726]: #1728
-#1732 := [nnf-neg #1729]: #1731
-#1702 := (not #1380)
-#1703 := (~ #1702 #1377)
-#1700 := (~ #1377 #1377)
-#1698 := (~ #1374 #1374)
-#1699 := [refl]: #1698
-#1701 := [nnf-pos #1699]: #1700
-#1704 := [nnf-neg #1701]: #1703
-#1695 := (not #1391)
-#1696 := (~ #1695 #1388)
-#1693 := (~ #1388 #1388)
-#1691 := (~ #1385 #1385)
-#1692 := [refl]: #1691
-#1694 := [nnf-pos #1692]: #1693
-#1697 := [nnf-neg #1694]: #1696
-#1688 := (not #1400)
-#1689 := (~ #1688 #1397)
-#1686 := (~ #1397 #1397)
-#1684 := (~ #1394 #1394)
-#1685 := [refl]: #1684
-#1687 := [nnf-pos #1685]: #1686
-#1690 := [nnf-neg #1687]: #1689
-#1682 := (~ #1681 #1681)
-#1683 := [refl]: #1682
-#1678 := (not #1403)
-#1679 := (~ #1678 #1675)
-#1676 := (~ #999 #1675)
-#1673 := (~ #996 #1672)
-#1668 := (~ #993 #1667)
-#1669 := [sk]: #1668
-#1652 := (~ #973 #973)
-#1653 := [refl]: #1652
-#1674 := [monotonicity #1653 #1669]: #1673
-#1677 := [nnf-pos #1674]: #1676
-#1680 := [nnf-neg #1677]: #1679
-#2132 := [nnf-neg #1680 #1683 #1690 #1697 #1704 #1732 #2128]: #2131
-#1650 := (~ #1403 #1647)
-#1632 := (exists (vars (?v1 S2)) #1631)
-#1633 := (or #1622 #1632)
-#1634 := (not #1633)
-#1648 := (~ #1634 #1647)
-#1644 := (not #1632)
-#1645 := (~ #1644 #1643)
-#1641 := (~ #1640 #1640)
-#1642 := [refl]: #1641
-#1646 := [nnf-neg #1642]: #1645
-#1638 := (~ #1637 #1637)
-#1639 := [refl]: #1638
-#1649 := [nnf-neg #1639 #1646]: #1648
-#1635 := (~ #1403 #1634)
-#1636 := [sk]: #1635
-#1651 := [trans #1636 #1649]: #1650
-#2136 := [nnf-neg #1651 #2132]: #2135
-#1610 := (not #961)
-#1611 := (~ #1610 #958)
-#1608 := (~ #958 #958)
-#1606 := (~ #955 #955)
-#1607 := [refl]: #1606
-#1609 := [nnf-pos #1607]: #1608
-#1612 := [nnf-neg #1609]: #1611
-#2140 := [nnf-neg #1612 #2136]: #2139
-#1602 := (~ #961 #1601)
-#1603 := [sk]: #1602
-#2144 := [nnf-neg #1603 #2140]: #2143
-#1578 := (not #935)
-#1579 := (~ #1578 #932)
-#1576 := (~ #932 #932)
-#1574 := (~ #929 #929)
-#1575 := [refl]: #1574
-#1577 := [nnf-pos #1575]: #1576
-#1580 := [nnf-neg #1577]: #1579
-#2148 := [nnf-neg #1580 #2144]: #2147
-#1570 := (~ #935 #1569)
-#1571 := [sk]: #1570
-#2152 := [nnf-neg #1571 #2148]: #2151
-#1551 := (not #920)
-#1552 := (~ #1551 #917)
-#1549 := (~ #917 #917)
-#1547 := (~ #916 #916)
-#1548 := [refl]: #1547
-#1550 := [nnf-pos #1548]: #1549
-#1553 := [nnf-neg #1550]: #1552
-#2156 := [nnf-neg #1553 #2152]: #2155
-#1543 := (~ #920 #1542)
-#1544 := [sk]: #1543
-#2160 := [nnf-neg #1544 #2156]: #2159
-#1505 := (~ #1504 #1504)
-#1538 := [refl]: #1505
-#2164 := [nnf-neg #1538 #2160]: #2163
-#1536 := (~ #868 #868)
-#1537 := [refl]: #1536
-#2167 := [nnf-neg #1537 #2164]: #2166
-#1480 := [not-or-elim #1475]: #1479
-#2168 := [mp~ #1480 #2167]: #2165
-#2169 := [mp #2168 #2483]: #2481
-#3122 := [mp #2169 #3121]: #3119
-#4037 := [mp #3122 #4036]: #4034
-#7144 := [unit-resolution #4037 #4122]: #4031
-#3361 := (or #4028 #4022)
-#3351 := [def-axiom]: #3361
-#7145 := [unit-resolution #3351 #7144]: #4022
-#3357 := (or #4025 #1542 #4019)
-#3359 := [def-axiom]: #3357
-#7146 := [unit-resolution #3359 #7145 #4053]: #4019
-#3355 := (or #4016 #4010)
-#3358 := [def-axiom]: #3355
-#7147 := [unit-resolution #3358 #7146]: #4010
-#4245 := [hypothesis]: #1560
-#3661 := (forall (vars (?v0 S2)) (:pat #3660) #78)
-#3664 := (iff #79 #3661)
-#3662 := (iff #78 #78)
-#3663 := [refl]: #3662
-#3665 := [quant-intro #3663]: #3664
-#1502 := (~ #79 #79)
-#1533 := (~ #78 #78)
-#1534 := [refl]: #1533
-#1503 := [nnf-pos #1534]: #1502
-#1478 := [not-or-elim #1475]: #79
-#1535 := [mp~ #1478 #1503]: #79
-#3666 := [mp #1535 #3665]: #3661
-#6940 := (not #3661)
-#4154 := (or #6940 #1884)
-#4155 := [quant-inst #1555]: #4154
-#4251 := [unit-resolution #4155 #3666 #4245]: false
-#4288 := [lemma #4251]: #1884
-#3186 := (or #1841 #1560)
-#3272 := [def-axiom]: #3186
-#7148 := [unit-resolution #3272 #4288]: #1841
-#3217 := (or #4013 #1946 #4007)
-#3375 := [def-axiom]: #3217
-#7149 := [unit-resolution #3375 #7148 #7147]: #4007
-#3397 := (or #4004 #3998)
-#3367 := [def-axiom]: #3397
-#7150 := [unit-resolution #3367 #7149]: #3998
-#6160 := [hypothesis]: #1584
-#6331 := (or #6940 #2503)
-#6332 := [quant-inst #1581]: #6331
-#6161 := [unit-resolution #6332 #3666 #6160]: false
-#6581 := [lemma #6161]: #2503
-#3277 := (or #2518 #1584)
-#3274 := [def-axiom]: #3277
-#7151 := [unit-resolution #3274 #6581]: #2518
-#3394 := (or #4001 #2523 #3995)
-#3395 := [def-axiom]: #3394
-#7152 := [unit-resolution #3395 #7151 #7150]: #3995
-#3378 := (or #3992 #3986)
-#3385 := [def-axiom]: #3378
-#7153 := [unit-resolution #3385 #7152]: #3986
-#3415 := (or #3989 #3705 #3983)
-#3400 := [def-axiom]: #3415
-#7154 := [unit-resolution #3400 #7153]: #3986
-#7155 := [unit-resolution #7154 #6633]: #3983
-#3407 := (or #3980 #3974)
-#3408 := [def-axiom]: #3407
-#7801 := [unit-resolution #3408 #7155]: #3974
-#6893 := [hypothesis]: #3817
-#3559 := (or #3814 #149)
-#3554 := [def-axiom]: #3559
-#6888 := [unit-resolution #3554 #6893]: #149
-#3555 := (or #3814 #3751)
-#3556 := [def-axiom]: #3555
-#6808 := [unit-resolution #3556 #6893]: #3751
-#3431 := (or #3980 #111)
-#3432 := [def-axiom]: #3431
-#6894 := [unit-resolution #3432 #7155]: #111
-#4278 := (or #503 #169 #804)
-#4273 := [hypothesis]: #111
-#4275 := (= #168 #110)
-#4274 := [hypothesis]: #149
-#4276 := [monotonicity #4274]: #4275
-#4271 := [trans #4276 #4273]: #169
-#4158 := [hypothesis]: #1849
-#4277 := [unit-resolution #4158 #4271]: false
-#4279 := [lemma #4277]: #4278
-#6895 := [unit-resolution #4279 #6888 #6894]: #169
-#3298 := (or #3790 #1849)
-#3299 := [def-axiom]: #3298
-#6896 := [unit-resolution #3299 #6895]: #3790
-#3402 := (or #3814 #3808)
-#3403 := [def-axiom]: #3402
-#6861 := [unit-resolution #3403 #6893]: #3808
-#7454 := (or #3768 #503)
-#6806 := (?v1!7 ?v0!8)
-#7010 := (f19 f25 #6806)
-#7056 := (* -1::Int #7010)
-#6807 := (f19 f20 #6806)
-#7278 := (+ #6807 #7056)
-#7328 := (>= #7278 0::Int)
-#7277 := (= #6807 #7010)
-#7409 := (= #7010 #6807)
-#7410 := [monotonicity #4274]: #7409
-#7411 := [symm #7410]: #7277
-#7412 := (not #7277)
-#7413 := (or #7412 #7328)
-#7414 := [th-lemma arith triangle-eq]: #7413
-#7415 := [unit-resolution #7414 #7411]: #7328
-#6897 := (* -1::Int #6807)
-#6741 := (f19 f20 ?v0!8)
-#6898 := (+ #6741 #6897)
-#6903 := (<= #6898 0::Int)
-#6821 := (not #6903)
-#6850 := (f6 f7 #6806)
-#6853 := (f5 #6850 ?v0!8)
-#6889 := (f4 #6853)
-#6890 := (* -1::Int #6889)
-#6908 := (+ #6897 #6890)
-#6925 := (+ #6741 #6908)
-#5298 := (= #6925 0::Int)
-#6077 := (not #5298)
-#6904 := (f11 f21 #6806)
-#6851 := (= #6904 f1)
-#6852 := (not #6851)
-#6078 := (or #6903 #6852 #6077)
-#6102 := (not #6078)
-#6742 := (* -1::Int #6741)
-#6772 := (+ f3 #6742)
-#6773 := (<= #6772 0::Int)
-#7378 := (not #6773)
-#7416 := [hypothesis]: #3771
-#3249 := (or #3768 #1758)
-#3591 := [def-axiom]: #3249
-#7417 := [unit-resolution #3591 #7416]: #1758
-#7026 := (+ #1754 #6742)
-#7012 := (>= #7026 0::Int)
-#7025 := (= #1754 #6741)
-#7418 := (= #6741 #1754)
-#6982 := (= f20 f25)
-#6983 := [symm #4274]: #6982
-#7419 := [monotonicity #6983]: #7418
-#7420 := [symm #7419]: #7025
-#7421 := (not #7025)
-#7422 := (or #7421 #7012)
-#7423 := [th-lemma arith triangle-eq]: #7422
-#7424 := [unit-resolution #7423 #7420]: #7012
-#7379 := (not #7012)
-#7380 := (or #7378 #7379 #1757)
-#7368 := [hypothesis]: #1758
-#7369 := [hypothesis]: #6773
-#7376 := [hypothesis]: #7012
-#7377 := [th-lemma arith farkas 1 -1 1 #7376 #7369 #7368]: false
-#7381 := [lemma #7377]: #7380
-#7425 := [unit-resolution #7381 #7424 #7417]: #7378
-#7428 := (or #6773 #6102)
-#3589 := (or #3768 #1753)
-#3254 := [def-axiom]: #3589
-#7426 := [unit-resolution #3254 #7416]: #1753
-#3405 := (or #3980 #3742)
-#3406 := [def-axiom]: #3405
-#7427 := [unit-resolution #3406 #7155]: #3742
-#6227 := (or #3747 #1752 #6773 #6102)
-#6103 := (or #1752 #6773 #6102)
-#6189 := (or #3747 #6103)
-#6704 := (iff #6189 #6227)
-#6805 := [rewrite]: #6704
-#6190 := [quant-inst #1751]: #6189
-#6797 := [mp #6190 #6805]: #6227
-#7429 := [unit-resolution #6797 #7427 #7426]: #7428
-#7430 := [unit-resolution #7429 #7425]: #6102
-#6854 := (or #6078 #6821)
-#6860 := [def-axiom]: #6854
-#7431 := [unit-resolution #6860 #7430]: #6821
-#7057 := (+ #1754 #7056)
-#7058 := (<= #7057 0::Int)
-#7101 := (+ #6890 #7056)
-#7102 := (+ #1754 #7101)
-#7105 := (= #7102 0::Int)
-#7195 := (<= #7102 0::Int)
-#7327 := (<= #7278 0::Int)
-#7432 := (or #7412 #7327)
-#7433 := [th-lemma arith triangle-eq]: #7432
-#7434 := [unit-resolution #7433 #7411]: #7327
-#6815 := (<= #6925 0::Int)
-#6862 := (or #6078 #5298)
-#6863 := [def-axiom]: #6862
-#7435 := [unit-resolution #6863 #7430]: #5298
-#7436 := (or #6077 #6815)
-#7437 := [th-lemma arith triangle-eq]: #7436
-#7438 := [unit-resolution #7437 #7435]: #6815
-#7011 := (<= #7026 0::Int)
-#7439 := (or #7421 #7011)
-#7440 := [th-lemma arith triangle-eq]: #7439
-#7441 := [unit-resolution #7440 #7420]: #7011
-#7390 := (not #7327)
-#7389 := (not #6815)
-#7388 := (not #7011)
-#7391 := (or #7195 #7388 #7389 #7390)
-#7375 := [hypothesis]: #7327
-#7128 := [hypothesis]: #6815
-#7384 := [hypothesis]: #7011
-#7385 := (not #7195)
-#7386 := [hypothesis]: #7385
-#7387 := [th-lemma arith farkas -1 1 1 1 #7386 #7384 #7128 #7375]: false
-#7392 := [lemma #7387]: #7391
-#7442 := [unit-resolution #7392 #7441 #7438 #7434]: #7195
-#7196 := (>= #7102 0::Int)
-#6820 := (>= #6925 0::Int)
-#7443 := (or #6077 #6820)
-#7444 := [th-lemma arith triangle-eq]: #7443
-#7445 := [unit-resolution #7444 #7435]: #6820
-#7406 := (not #7328)
-#7405 := (not #6820)
-#7407 := (or #7196 #7379 #7405 #7406)
-#7400 := [hypothesis]: #7328
-#7401 := [hypothesis]: #6820
-#7402 := (not #7196)
-#7403 := [hypothesis]: #7402
-#7404 := [th-lemma arith farkas -1 1 1 1 #7403 #7376 #7401 #7400]: false
-#7408 := [lemma #7404]: #7407
-#7446 := [unit-resolution #7408 #7424 #7445 #7415]: #7196
-#7447 := (or #7105 #7385 #7402)
-#7448 := [th-lemma arith triangle-eq]: #7447
-#7449 := [unit-resolution #7448 #7446 #7442]: #7105
-#7114 := (not #7105)
-#7122 := (or #7058 #7114)
-#3234 := (or #3768 #3760)
-#3575 := [def-axiom]: #3234
-#7450 := [unit-resolution #3575 #7416]: #3760
-#7125 := (or #3765 #7058 #7114)
-#7015 := (+ #1755 #6889)
-#7028 := (+ #7010 #7015)
-#7029 := (= #7028 0::Int)
-#7049 := (not #7029)
-#7013 := (+ #7010 #1755)
-#7014 := (>= #7013 0::Int)
-#7050 := (or #7014 #7049)
-#7126 := (or #3765 #7050)
-#7119 := (iff #7126 #7125)
-#6981 := (or #3765 #7122)
-#7017 := (iff #6981 #7125)
-#7118 := [rewrite]: #7017
-#7016 := (iff #7126 #6981)
-#7123 := (iff #7050 #7122)
-#7115 := (iff #7049 #7114)
-#7113 := (iff #7029 #7105)
-#7091 := (+ #6889 #7010)
-#7086 := (+ #1755 #7091)
-#7094 := (= #7086 0::Int)
-#7106 := (iff #7094 #7105)
-#7112 := [rewrite]: #7106
-#7095 := (iff #7029 #7094)
-#7092 := (= #7028 #7086)
-#7093 := [rewrite]: #7092
-#7096 := [monotonicity #7093]: #7095
-#7111 := [trans #7096 #7112]: #7113
-#7116 := [monotonicity #7111]: #7115
-#7089 := (iff #7014 #7058)
-#7051 := (+ #1755 #7010)
-#7048 := (>= #7051 0::Int)
-#7087 := (iff #7048 #7058)
-#7088 := [rewrite]: #7087
-#7054 := (iff #7014 #7048)
-#7052 := (= #7013 #7051)
-#7053 := [rewrite]: #7052
-#7055 := [monotonicity #7053]: #7054
-#7090 := [trans #7055 #7088]: #7089
-#7121 := [monotonicity #7090 #7116]: #7123
-#7124 := [monotonicity #7121]: #7016
-#7120 := [trans #7124 #7118]: #7119
-#6980 := [quant-inst #6806]: #7126
-#7127 := [mp #6980 #7120]: #7125
-#7451 := [unit-resolution #7127 #7450]: #7122
-#7452 := [unit-resolution #7451 #7449]: #7058
-#7453 := [th-lemma arith farkas 1 -1 1 1 #7424 #7452 #7431 #7415]: false
-#7455 := [lemma #7453]: #7454
-#6859 := [unit-resolution #7455 #6888]: #3768
-#3562 := (or #3811 #3771 #3805)
-#3566 := [def-axiom]: #3562
-#6864 := [unit-resolution #3566 #6859 #6861]: #3805
-#3283 := (or #3802 #3796)
-#3284 := [def-axiom]: #3283
-#6869 := [unit-resolution #3284 #6864]: #3796
-#3571 := (or #3799 #2759 #3793)
-#3568 := [def-axiom]: #3571
-#6935 := [unit-resolution #3568 #6869 #6896]: #2759
-#3225 := (or #2754 #1828)
-#3226 := [def-axiom]: #3225
-#6870 := [unit-resolution #3226 #6935]: #1828
-#6589 := (f19 f20 ?v1!10)
-#6584 := (* -1::Int #6589)
-#6868 := (+ #1817 #6584)
-#6946 := (>= #6868 0::Int)
-#4646 := (or #6946 #503)
-#6329 := (= #1817 #6589)
-#6605 := (= #6589 #1817)
-#6606 := [monotonicity #6983]: #6605
-#6585 := [symm #6606]: #6329
-#7134 := (not #6946)
-#6586 := [hypothesis]: #7134
-#6330 := (not #6329)
-#6322 := (or #6330 #6946)
-#6328 := [th-lemma arith triangle-eq]: #6322
-#6580 := [unit-resolution #6328 #6586 #6585]: false
-#4567 := [lemma #6580]: #4646
-#6867 := [unit-resolution #4567 #6888]: #6946
-#3583 := (or #2754 #1821)
-#3585 := [def-axiom]: #3583
-#6945 := [unit-resolution #3585 #6935]: #1821
-#3586 := (not #2356)
-#3584 := (or #2754 #3586)
-#3587 := [def-axiom]: #3584
-#6866 := [unit-resolution #3587 #6935]: #3586
-#7019 := (or #7134 #3756 #1827 #2356 #1820 #503)
-#6765 := (f19 f20 ?v0!11)
-#6766 := (* -1::Int #6765)
-#6948 := (+ #1831 #6766)
-#6949 := (<= #6948 0::Int)
-#6947 := (= #1831 #6765)
-#6984 := (= #6765 #1831)
-#6985 := [monotonicity #6983]: #6984
-#6986 := [symm #6985]: #6947
-#6987 := (not #6947)
-#6988 := (or #6987 #6949)
-#6975 := [th-lemma arith triangle-eq]: #6988
-#6976 := [unit-resolution #6975 #6986]: #6949
-#7130 := [hypothesis]: #3586
-#7143 := [hypothesis]: #1828
-#7140 := [hypothesis]: #3751
-#7131 := [hypothesis]: #6946
-#6572 := (+ f3 #6584)
-#6576 := (<= #6572 0::Int)
-#7138 := (not #6576)
-#6974 := [hypothesis]: #1821
-#6977 := (or #7138 #1820 #7134)
-#6978 := [th-lemma arith assign-bounds -1 -1]: #6977
-#6979 := [unit-resolution #6978 #7131 #6974]: #7138
-#7133 := (not #6949)
-#7160 := (or #6576 #1827 #3756 #7134 #2356 #7133)
-#6907 := (+ #6589 #6766)
-#6926 := (+ #1824 #6907)
-#6924 := (>= #6926 0::Int)
-#7132 := (not #6924)
-#7129 := [hypothesis]: #6949
-#7135 := (or #7132 #7133 #2356 #7134)
-#7136 := [th-lemma arith assign-bounds -1 -1 1]: #7135
-#7137 := [unit-resolution #7136 #7131 #7130 #7129]: #7132
-#6587 := (f11 f21 ?v1!10)
-#6588 := (= #6587 f1)
-#7139 := [hypothesis]: #7138
-#5703 := (or #6588 #6576)
-#6635 := (or #3756 #6588 #6576)
-#6636 := (or #3756 #5703)
-#6638 := (iff #6636 #6635)
-#6211 := [rewrite]: #6638
-#6637 := [quant-inst #1815]: #6636
-#6252 := [mp #6637 #6211]: #6635
-#7141 := [unit-resolution #6252 #7140]: #5703
-#7142 := [unit-resolution #7141 #7139]: #6588
-#6849 := (not #6588)
-#7157 := (or #6849 #6924)
-#3399 := (or #3980 #3734)
-#3404 := [def-axiom]: #3399
-#7156 := [unit-resolution #3404 #7155]: #3734
-#6930 := (or #3739 #6849 #1827 #6924)
-#6927 := (or #6849 #1827 #6924)
-#6931 := (or #3739 #6927)
-#3348 := (iff #6931 #6930)
-#3221 := [rewrite]: #3348
-#6932 := [quant-inst #1816 #1815]: #6931
-#4127 := [mp #6932 #3221]: #6930
-#7158 := [unit-resolution #4127 #7156 #7143]: #7157
-#7159 := [unit-resolution #7158 #7142 #7137]: false
-#7161 := [lemma #7159]: #7160
-#7018 := [unit-resolution #7161 #6979 #7131 #7140 #7143 #7130 #6976]: false
-#7020 := [lemma #7018]: #7019
-#6950 := [unit-resolution #7020 #6866 #6945 #6867 #6870 #6808 #6888]: false
-#7021 := [lemma #6950]: #3814
-#3419 := (or #3977 #3817 #3971)
-#3421 := [def-axiom]: #3419
-#7802 := [unit-resolution #3421 #7021 #7801]: #3971
-#3459 := (or #3968 #194)
-#3464 := [def-axiom]: #3459
-#8144 := [unit-resolution #3464 #7802]: #194
-#8146 := [symm #8144]: #8145
-#16380 := [monotonicity #8146]: #16289
-#16388 := [monotonicity #16380]: #16359
-#16389 := [symm #16388]: #16382
-#16391 := [monotonicity #16389]: #16390
-#19086 := [hypothesis]: #4534
-#5090 := (or #4541 #5087)
-#8395 := [hypothesis]: #1943
-#3547 := (or #1942 #1937)
-#3543 := [def-axiom]: #3547
-#8389 := [unit-resolution #3543 #8395]: #1937
-#3548 := (not #1941)
-#3542 := (or #1942 #3548)
-#3549 := [def-axiom]: #3542
-#8396 := [unit-resolution #3549 #8395]: #3548
-#5879 := (* -1::Int #1940)
-#5980 := (+ #185 #5879)
-#7760 := (>= #5980 0::Int)
-#4329 := (f11 f21 ?v0!14)
-#4316 := (= #4329 f1)
-#4315 := (= ?v0!14 f28)
-#4345 := (or #4315 #4316)
-#4328 := (f11 #193 ?v0!14)
-#4330 := (= #4328 f1)
-#4348 := (iff #4330 #4345)
-#8073 := (or #7518 #4348)
-#4314 := (if #4315 #4057 #4316)
-#4317 := (iff #4330 #4314)
-#8074 := (or #7518 #4317)
-#8078 := (iff #8074 #8073)
-#8080 := (iff #8073 #8073)
-#8081 := [rewrite]: #8080
-#4337 := (iff #4317 #4348)
-#4344 := (iff #4314 #4345)
-#4318 := (if #4315 true #4316)
-#4346 := (iff #4318 #4345)
-#4338 := [rewrite]: #4346
-#4319 := (iff #4314 #4318)
-#4332 := [monotonicity #4060]: #4319
-#4347 := [trans #4332 #4338]: #4344
-#4349 := [monotonicity #4347]: #4337
-#8079 := [monotonicity #4349]: #8078
-#8082 := [trans #8079 #8081]: #8078
-#8077 := [quant-inst #115 #181 #3 #1935]: #8074
-#8083 := [mp #8077 #8082]: #8073
-#8113 := [unit-resolution #8083 #3646]: #4348
-#8114 := [hypothesis]: #1937
-#8147 := (= #4328 #1936)
-#8148 := [monotonicity #8146]: #8147
-#8149 := [trans #8148 #8114]: #4330
-#8094 := (not #4330)
-#8091 := (not #4348)
-#8095 := (or #8091 #8094 #4345)
-#8096 := [def-axiom]: #8095
-#8150 := [unit-resolution #8096 #8149 #8113]: #4345
-#8088 := (not #4345)
-#8152 := (or #8088 #4316)
-#7690 := (not #4315)
-#7795 := [hypothesis]: #3548
-#7815 := (or #7690 #1941)
-#7809 := (= #185 #1940)
-#7807 := (= #1940 #185)
-#7796 := [hypothesis]: #4315
-#7808 := [monotonicity #7796]: #7807
-#7810 := [symm #7808]: #7809
-#7811 := (= #1939 #185)
-#4131 := (= #4078 #185)
-#4082 := (f5 #195 f28)
-#4083 := (f4 #4082)
-#4101 := (>= #4083 0::Int)
-#4086 := (* -1::Int #4083)
-#4087 := (+ f3 #4086)
-#4088 := (<= #4087 0::Int)
-#4133 := (or #4088 #4101)
-#6365 := (= #4083 0::Int)
-#6474 := (not #6365)
-#6475 := [hypothesis]: #6474
-#14 := (f6 f7 #10)
-#15 := (f5 #14 #11)
-#3600 := (pattern #15)
-#16 := (f4 #15)
-#17 := (= #16 0::Int)
-#12 := (= #10 #11)
-#20 := (not #12)
-#306 := (or #20 #17)
-#3601 := (forall (vars (?v0 S2) (?v1 S2)) (:pat #3600) #306)
-#314 := (forall (vars (?v0 S2) (?v1 S2)) #306)
-#3604 := (iff #314 #3601)
-#3602 := (iff #306 #306)
-#3603 := [refl]: #3602
-#3605 := [quant-intro #3603]: #3604
-#1510 := (~ #314 #314)
-#1508 := (~ #306 #306)
-#1509 := [refl]: #1508
-#1511 := [nnf-pos #1509]: #1510
-#18 := (implies #12 #17)
-#19 := (forall (vars (?v0 S2) (?v1 S2)) #18)
-#315 := (iff #19 #314)
-#312 := (iff #18 #306)
-#313 := [rewrite]: #312
-#316 := [quant-intro #313]: #315
-#305 := [asserted]: #19
-#319 := [mp #305 #316]: #314
-#1484 := [mp~ #319 #1511]: #314
-#3606 := [mp #1484 #3605]: #3601
-#6226 := (not #3601)
-#6401 := (or #6226 #6365)
-#4403 := (= f28 f28)
-#6358 := (not #4403)
-#6369 := (or #6358 #6365)
-#6400 := (or #6226 #6369)
-#6403 := (iff #6400 #6401)
-#6471 := (iff #6401 #6401)
-#6466 := [rewrite]: #6471
-#6387 := (iff #6369 #6365)
-#6377 := (or false #6365)
-#6382 := (iff #6377 #6365)
-#6386 := [rewrite]: #6382
-#6383 := (iff #6369 #6377)
-#6378 := (iff #6358 false)
-#6370 := (iff #6358 #6910)
-#4406 := (iff #4403 true)
-#4407 := [rewrite]: #4406
-#6375 := [monotonicity #4407]: #6370
-#6379 := [trans #6375 #6914]: #6378
-#6384 := [monotonicity #6379]: #6383
-#6398 := [trans #6384 #6386]: #6387
-#6439 := [monotonicity #6398]: #6403
-#6470 := [trans #6439 #6466]: #6403
-#6402 := [quant-inst #181 #181]: #6400
-#6472 := [mp #6402 #6470]: #6401
-#6476 := [unit-resolution #6472 #3606 #6475]: false
-#6482 := [lemma #6476]: #6365
-#7797 := (or #6474 #4101)
-#7798 := [th-lemma arith triangle-eq]: #7797
-#7799 := [unit-resolution #7798 #6482]: #4101
-#7285 := (not #4101)
-#7283 := (or #4133 #7285)
-#7286 := [def-axiom]: #7283
-#7800 := [unit-resolution #7286 #7799]: #4133
-#3445 := (or #3968 #3838)
-#3426 := [def-axiom]: #3445
-#7803 := [unit-resolution #3426 #7802]: #3838
-#4136 := (not #4133)
-#7259 := (or #3843 #4136 #4131)
-#4089 := (+ #1146 #4086)
-#4090 := (+ #185 #4089)
-#4091 := (<= #4090 0::Int)
-#4129 := (or #4088 #4091)
-#4130 := (not #4129)
-#4132 := (or #4130 #4131)
-#7260 := (or #3843 #4132)
-#7266 := (iff #7260 #7259)
-#4139 := (or #4136 #4131)
-#7262 := (or #3843 #4139)
-#7265 := (iff #7262 #7259)
-#7261 := [rewrite]: #7265
-#7263 := (iff #7260 #7262)
-#4140 := (iff #4132 #4139)
-#4137 := (iff #4130 #4136)
-#4134 := (iff #4129 #4133)
-#4104 := (iff #4091 #4101)
-#4098 := (<= #4086 0::Int)
-#4102 := (iff #4098 #4101)
-#4103 := [rewrite]: #4102
-#4099 := (iff #4091 #4098)
-#4096 := (= #4090 #4086)
-#4097 := [rewrite]: #4096
-#4100 := [monotonicity #4097]: #4099
-#4105 := [trans #4100 #4103]: #4104
-#4135 := [monotonicity #4105]: #4134
-#4138 := [monotonicity #4135]: #4137
-#4141 := [monotonicity #4138]: #4140
-#7264 := [monotonicity #4141]: #7263
-#7271 := [trans #7264 #7261]: #7266
-#7256 := [quant-inst #181]: #7260
-#7272 := [mp #7256 #7271]: #7259
-#7804 := [unit-resolution #7272 #7803 #7800]: #4131
-#7805 := (= #1939 #4078)
-#7806 := [monotonicity #7796]: #7805
-#7812 := [trans #7806 #7804]: #7811
-#7813 := [trans #7812 #7810]: #1941
-#7814 := [unit-resolution #7795 #7813]: false
-#7816 := [lemma #7814]: #7815
-#8151 := [unit-resolution #7816 #7795]: #7690
-#8089 := (or #8088 #4315 #4316)
-#8090 := [def-axiom]: #8089
-#8153 := [unit-resolution #8090 #8151]: #8152
-#8154 := [unit-resolution #8153 #8150]: #4316
-#3456 := (or #3968 #184)
-#3457 := [def-axiom]: #3456
-#8155 := [unit-resolution #3457 #7802]: #184
-#3443 := (or #3980 #3726)
-#3398 := [def-axiom]: #3443
-#8156 := [unit-resolution #3398 #7155]: #3726
-#7693 := (not #4316)
-#8138 := (or #3731 #183 #7693 #7760)
-#7761 := (or #183 #7693 #7760)
-#8139 := (or #3731 #7761)
-#8141 := (iff #8139 #8138)
-#8142 := [rewrite]: #8141
-#8140 := [quant-inst #1935 #181]: #8139
-#8143 := [mp #8140 #8142]: #8138
-#8157 := [unit-resolution #8143 #8156 #8155 #8154]: #7760
-#6005 := (f5 #195 ?v0!14)
-#6011 := (f4 #6005)
-#6036 := (+ #5879 #6011)
-#6037 := (+ #185 #6036)
-#6038 := (>= #6037 0::Int)
-#8046 := (not #6038)
-#6012 := (* -1::Int #6011)
-#6010 := (+ f3 #6012)
-#6014 := (<= #6010 0::Int)
-#6049 := (or #6014 #6038)
-#6056 := (not #6049)
-#8032 := (or #3843 #6056 #1941)
-#6015 := (+ #1146 #6012)
-#6013 := (+ #1940 #6015)
-#6017 := (<= #6013 0::Int)
-#6051 := (or #6014 #6017)
-#6052 := (not #6051)
-#6054 := (or #6052 #1941)
-#8033 := (or #3843 #6054)
-#8040 := (iff #8033 #8032)
-#6067 := (or #6056 #1941)
-#8035 := (or #3843 #6067)
-#8038 := (iff #8035 #8032)
-#8039 := [rewrite]: #8038
-#8036 := (iff #8033 #8035)
-#6068 := (iff #6054 #6067)
-#6057 := (iff #6052 #6056)
-#6055 := (iff #6051 #6049)
-#6043 := (iff #6017 #6038)
-#6027 := (+ #1940 #6012)
-#6030 := (+ #1146 #6027)
-#6034 := (<= #6030 0::Int)
-#6039 := (iff #6034 #6038)
-#6040 := [rewrite]: #6039
-#6035 := (iff #6017 #6034)
-#6031 := (= #6013 #6030)
-#6032 := [rewrite]: #6031
-#6033 := [monotonicity #6032]: #6035
-#6044 := [trans #6033 #6040]: #6043
-#6053 := [monotonicity #6044]: #6055
-#6058 := [monotonicity #6053]: #6057
-#6045 := [monotonicity #6058]: #6068
-#8037 := [monotonicity #6045]: #8036
-#8041 := [trans #8037 #8039]: #8040
-#8034 := [quant-inst #1935]: #8033
-#8042 := [mp #8034 #8041]: #8032
-#8158 := [unit-resolution #8042 #7803 #7795]: #6056
-#8047 := (or #6049 #8046)
-#8048 := [def-axiom]: #8047
-#8159 := [unit-resolution #8048 #8158]: #8046
-#8125 := (>= #6011 0::Int)
-#6251 := (<= #6011 0::Int)
-#6224 := (not #6251)
-#6208 := (= f28 ?v0!14)
-#6209 := (not #6208)
-#8164 := (iff #7690 #6209)
-#8162 := (iff #4315 #6208)
-#8160 := (iff #6208 #4315)
-#8161 := [commutativity]: #8160
-#8163 := [symm #8161]: #8162
-#8165 := [monotonicity #8163]: #8164
-#8166 := [mp #8151 #8165]: #6209
-#6225 := (or #6208 #6224)
-#325 := (<= #16 0::Int)
-#326 := (not #325)
-#329 := (or #12 #326)
-#3607 := (forall (vars (?v0 S2) (?v1 S2)) (:pat #3600) #329)
-#332 := (forall (vars (?v0 S2) (?v1 S2)) #329)
-#3610 := (iff #332 #3607)
-#3608 := (iff #329 #329)
-#3609 := [refl]: #3608
-#3611 := [quant-intro #3609]: #3610
-#1486 := (~ #332 #332)
-#1485 := (~ #329 #329)
-#1482 := [refl]: #1485
-#1487 := [nnf-pos #1482]: #1486
-#21 := (< 0::Int #16)
-#22 := (implies #20 #21)
-#23 := (forall (vars (?v0 S2) (?v1 S2)) #22)
-#335 := (iff #23 #332)
-#318 := (or #12 #21)
-#322 := (forall (vars (?v0 S2) (?v1 S2)) #318)
-#333 := (iff #322 #332)
-#330 := (iff #318 #329)
-#327 := (iff #21 #326)
-#328 := [rewrite]: #327
-#331 := [monotonicity #328]: #330
-#334 := [quant-intro #331]: #333
-#323 := (iff #23 #322)
-#320 := (iff #22 #318)
-#321 := [rewrite]: #320
-#324 := [quant-intro #321]: #323
-#336 := [trans #324 #334]: #335
-#317 := [asserted]: #23
-#337 := [mp #317 #336]: #332
-#1483 := [mp~ #337 #1487]: #332
-#3612 := [mp #1483 #3611]: #3607
-#7600 := (not #3607)
-#8126 := (or #7600 #6208 #6224)
-#8127 := (or #7600 #6225)
-#8129 := (iff #8127 #8126)
-#8130 := [rewrite]: #8129
-#8128 := [quant-inst #181 #1935]: #8127
-#8131 := [mp #8128 #8130]: #8126
-#8167 := [unit-resolution #8131 #3612]: #6225
-#8168 := [unit-resolution #8167 #8166]: #6224
-#8169 := (or #8125 #6251)
-#8170 := [th-lemma arith farkas 1 1]: #8169
-#8171 := [unit-resolution #8170 #8168]: #8125
-#8172 := [th-lemma arith farkas 1 -1 1 #8171 #8159 #8157]: false
-#8173 := [lemma #8172]: #1942
-#8376 := [unit-resolution #8173 #8396 #8389]: false
-#8397 := [lemma #8376]: #1942
-#3450 := (or #3968 #3962)
-#3451 := [def-axiom]: #3450
-#7966 := [unit-resolution #3451 #7802]: #3962
-#3444 := (or #3968 #3830)
-#3390 := [def-axiom]: #3444
-#7967 := [unit-resolution #3390 #7802]: #3830
-#5748 := (or #1922 #3835 #3843)
-#4760 := (f5 #195 ?v0!13)
-#4761 := (f4 #4760)
-#4791 := (+ #1920 #4761)
-#4792 := (+ #185 #4791)
-#4951 := (>= #4792 0::Int)
-#4794 := (= #4792 0::Int)
-#4766 := (* -1::Int #4761)
-#4767 := (+ f3 #4766)
-#4768 := (<= #4767 0::Int)
-#5036 := (not #4768)
-#4703 := (* -1::Int #1918)
-#4881 := (+ #4703 #4761)
-#4879 := (+ #185 #4881)
-#4882 := (>= #4879 0::Int)
-#4975 := (or #4768 #4882)
-#4979 := (not #4975)
-#4945 := (= #1919 #1918)
-#5470 := (not #4945)
-#5031 := (= #1918 #1919)
-#5429 := (not #5031)
-#5471 := (iff #5429 #5470)
-#5469 := (iff #5031 #4945)
-#5468 := [commutativity]: #5469
-#5472 := [monotonicity #5468]: #5471
-#5428 := [hypothesis]: #1923
-#5431 := (or #5429 #1922)
-#5467 := [th-lemma arith triangle-eq]: #5431
-#5466 := [unit-resolution #5467 #5428]: #5429
-#5474 := [mp #5466 #5472]: #5470
-#4981 := (or #4979 #4945)
-#5494 := [hypothesis]: #3838
-#4983 := (or #3843 #4979 #4945)
-#4769 := (+ #1146 #4766)
-#4770 := (+ #1918 #4769)
-#4798 := (<= #4770 0::Int)
-#4954 := (or #4768 #4798)
-#4955 := (not #4954)
-#4974 := (or #4955 #4945)
-#5002 := (or #3843 #4974)
-#5074 := (iff #5002 #4983)
-#5004 := (or #3843 #4981)
-#5066 := (iff #5004 #4983)
-#5067 := [rewrite]: #5066
-#5064 := (iff #5002 #5004)
-#4977 := (iff #4974 #4981)
-#4980 := (iff #4955 #4979)
-#4976 := (iff #4954 #4975)
-#4889 := (iff #4798 #4882)
-#4871 := (+ #1918 #4766)
-#4872 := (+ #1146 #4871)
-#4874 := (<= #4872 0::Int)
-#4883 := (iff #4874 #4882)
-#4884 := [rewrite]: #4883
-#4875 := (iff #4798 #4874)
-#4870 := (= #4770 #4872)
-#4873 := [rewrite]: #4870
-#4880 := [monotonicity #4873]: #4875
-#4890 := [trans #4880 #4884]: #4889
-#4978 := [monotonicity #4890]: #4976
-#4973 := [monotonicity #4978]: #4980
-#4982 := [monotonicity #4973]: #4977
-#5065 := [monotonicity #4982]: #5064
-#5075 := [trans #5065 #5067]: #5074
-#5003 := [quant-inst #1917]: #5002
-#5035 := [mp #5003 #5075]: #4983
-#5473 := [unit-resolution #5035 #5494]: #4981
-#5495 := [unit-resolution #5473 #5474]: #4979
-#5008 := (or #4975 #5036)
-#5009 := [def-axiom]: #5008
-#5496 := [unit-resolution #5009 #5495]: #5036
-#5005 := (not #4882)
-#5010 := (or #4975 #5005)
-#5011 := [def-axiom]: #5010
-#5497 := [unit-resolution #5011 #5495]: #5005
-#4796 := (or #4768 #4882 #4794)
-#5498 := [hypothesis]: #3830
-#4854 := (or #3835 #4768 #4882 #4794)
-#4799 := (+ #4761 #1920)
-#4864 := (+ #185 #4799)
-#4865 := (= #4864 0::Int)
-#4866 := (or #4768 #4798 #4865)
-#4852 := (or #3835 #4866)
-#4863 := (iff #4852 #4854)
-#4855 := (or #3835 #4796)
-#4857 := (iff #4855 #4854)
-#4862 := [rewrite]: #4857
-#4888 := (iff #4852 #4855)
-#4797 := (iff #4866 #4796)
-#4789 := (iff #4865 #4794)
-#4790 := (= #4864 #4792)
-#4793 := [rewrite]: #4790
-#4795 := [monotonicity #4793]: #4789
-#4891 := [monotonicity #4890 #4795]: #4797
-#4856 := [monotonicity #4891]: #4888
-#4861 := [trans #4856 #4862]: #4863
-#4853 := [quant-inst #1917]: #4852
-#4952 := [mp #4853 #4861]: #4854
-#5493 := [unit-resolution #4952 #5498]: #4796
-#5450 := [unit-resolution #5493 #5497 #5496]: #4794
-#5424 := (not #4794)
-#5366 := (or #5424 #4951)
-#5425 := [th-lemma arith triangle-eq]: #5366
-#5423 := [unit-resolution #5425 #5450]: #4951
-#5032 := (<= #1921 0::Int)
-#5426 := (or #5032 #1922)
-#5743 := [th-lemma arith farkas 1 1]: #5426
-#5746 := [unit-resolution #5743 #5428]: #5032
-#5747 := [th-lemma arith farkas 1 -1 1 #5746 #5497 #5423]: false
-#5749 := [lemma #5747]: #5748
-#7968 := [unit-resolution #5749 #7803 #7967]: #1922
-#3449 := (or #3965 #1923 #3959)
-#3342 := [def-axiom]: #3449
-#7969 := [unit-resolution #3342 #7968 #7966]: #3959
-#3470 := (or #3956 #3950)
-#3471 := [def-axiom]: #3470
-#10050 := [unit-resolution #3471 #7969]: #3950
-#3469 := (or #3953 #1943 #3947)
-#3465 := [def-axiom]: #3469
-#10051 := [unit-resolution #3465 #10050]: #3950
-#10052 := [unit-resolution #10051 #8397]: #3947
-#3497 := (or #3944 #3855)
-#3475 := [def-axiom]: #3497
-#10053 := [unit-resolution #3475 #10052]: #3855
-#10974 := (or #3860 #4541 #5087)
-#11004 := (or #3860 #5090)
-#10996 := (iff #11004 #10974)
-#11018 := [rewrite]: #10996
-#10999 := [quant-inst #2034]: #11004
-#11019 := [mp #10999 #11018]: #10974
-#19095 := [unit-resolution #11019 #10053]: #5090
-#19096 := [unit-resolution #19095 #19086]: #5087
-#11867 := (not #5087)
-#11974 := (or #11867 #12022)
-#12636 := [th-lemma arith triangle-eq]: #11974
-#19097 := [unit-resolution #12636 #19096]: #12022
-#13344 := (+ #5343 #13327)
-#13366 := (>= #13344 0::Int)
-#3466 := (or #3956 #3846)
-#3467 := [def-axiom]: #3466
-#7970 := [unit-resolution #3467 #7969]: #3846
-#18080 := (or #3851 #13366)
-#18086 := [quant-inst #5342]: #18080
-#19087 := [unit-resolution #18086 #7970]: #13366
-#15969 := (not #12022)
-#19098 := (or #15968 #15969)
-#3538 := (not #1988)
-#5971 := [hypothesis]: #2870
-#3541 := (or #2865 #3538)
-#3425 := [def-axiom]: #3541
-#5972 := [unit-resolution #3425 #5971]: #3538
-#3536 := (or #2865 #1979)
-#3539 := [def-axiom]: #3536
-#5974 := [unit-resolution #3539 #5971]: #1979
-#3537 := (or #2865 #1981)
-#3540 := [def-axiom]: #3537
-#5973 := [unit-resolution #3540 #5971]: #1981
-#10629 := (or #2850 #1988 #1978)
-#5727 := (= ?v1!16 f28)
-#9301 := (= f28 ?v1!16)
-#5888 := (f5 #195 ?v1!16)
-#5889 := (f4 #5888)
-#9302 := (<= #5889 0::Int)
-#9711 := [hypothesis]: #3538
-#5985 := (+ #1985 #4079)
-#9236 := (<= #5985 0::Int)
-#5682 := (f11 f21 ?v0!17)
-#5683 := (= #5682 f1)
-#5681 := (= ?v0!17 f28)
-#5689 := (or #5681 #5683)
-#5671 := (f11 #193 ?v0!17)
-#5672 := (= #5671 f1)
-#5694 := (iff #5672 #5689)
-#9981 := (or #7518 #5694)
-#5684 := (if #5681 #4057 #5683)
-#5685 := (iff #5672 #5684)
-#9982 := (or #7518 #5685)
-#9550 := (iff #9982 #9981)
-#9683 := (iff #9981 #9981)
-#9737 := [rewrite]: #9683
-#5695 := (iff #5685 #5694)
-#5692 := (iff #5684 #5689)
-#5686 := (if #5681 true #5683)
-#5690 := (iff #5686 #5689)
-#5691 := [rewrite]: #5690
-#5687 := (iff #5684 #5686)
-#5688 := [monotonicity #4060]: #5687
-#5693 := [trans #5688 #5691]: #5692
-#5696 := [monotonicity #5693]: #5695
-#9267 := [monotonicity #5696]: #9550
-#9960 := [trans #9267 #9737]: #9550
-#9682 := [quant-inst #115 #181 #3 #1976]: #9982
-#9972 := [mp #9682 #9960]: #9981
-#10602 := [unit-resolution #9972 #3646]: #5694
-#10603 := [hypothesis]: #1981
-#10604 := (= #5671 #1980)
-#10599 := [monotonicity #8146]: #10604
-#10571 := [trans #10599 #10603]: #5672
-#9787 := (not #5672)
-#9971 := (not #5694)
-#9963 := (or #9971 #9787 #5689)
-#9964 := [def-axiom]: #9963
-#10605 := [unit-resolution #9964 #10571 #10602]: #5689
-#8065 := (not #5689)
-#10754 := (or #9236 #8065)
-#9774 := (not #9236)
-#10735 := [hypothesis]: #9774
-#5652 := (f19 f20 ?v0!17)
-#5545 := (* -1::Int #5652)
-#5605 := (+ #1985 #5545)
-#5610 := (<= #5605 0::Int)
-#8385 := (or #3851 #5610)
-#5582 := (+ #5652 #1986)
-#5595 := (>= #5582 0::Int)
-#8384 := (or #3851 #5595)
-#7710 := (iff #8384 #8385)
-#7667 := (iff #8385 #8385)
-#8392 := [rewrite]: #7667
-#5612 := (iff #5595 #5610)
-#5596 := (+ #1986 #5652)
-#5600 := (>= #5596 0::Int)
-#5611 := (iff #5600 #5610)
-#5609 := [rewrite]: #5611
-#5603 := (iff #5595 #5600)
-#5601 := (= #5582 #5596)
-#5602 := [rewrite]: #5601
-#5604 := [monotonicity #5602]: #5603
-#5613 := [trans #5604 #5609]: #5612
-#7791 := [monotonicity #5613]: #7710
-#7912 := [trans #7791 #8392]: #7710
-#8393 := [quant-inst #1976]: #8384
-#8060 := [mp #8393 #7912]: #8385
-#9712 := [unit-resolution #8060 #7970]: #5610
-#5540 := (+ #185 #5545)
-#9420 := (>= #5540 0::Int)
-#9400 := [hypothesis]: #5689
-#8193 := (not #5681)
-#9234 := (= #1985 #4078)
-#9689 := (not #9234)
-#9518 := (or #9689 #9236)
-#9690 := [th-lemma arith triangle-eq]: #9518
-#10715 := [unit-resolution #9690 #10735]: #9689
-#10725 := (or #8193 #9234)
-#9309 := [hypothesis]: #5681
-#10737 := [monotonicity #9309]: #9234
-#10736 := [hypothesis]: #9689
-#10724 := [unit-resolution #10736 #10737]: false
-#10726 := [lemma #10724]: #10725
-#10727 := [unit-resolution #10726 #10715]: #8193
-#8066 := (or #8065 #5681 #5683)
-#8067 := [def-axiom]: #8066
-#10723 := [unit-resolution #8067 #10727 #9400]: #5683
-#8196 := (not #5683)
-#9377 := (or #8196 #9420)
-#9723 := (not #9420)
-#9387 := [hypothesis]: #9723
-#9362 := [hypothesis]: #5683
-#9219 := (or #3731 #183 #8196 #9420)
-#9438 := (or #183 #8196 #9420)
-#9235 := (or #3731 #9438)
-#9246 := (iff #9235 #9219)
-#9220 := [rewrite]: #9246
-#9237 := [quant-inst #1976 #181]: #9235
-#8778 := [mp #9237 #9220]: #9219
-#9376 := [unit-resolution #8778 #8156 #8155 #9362 #9387]: false
-#9457 := [lemma #9376]: #9377
-#10721 := [unit-resolution #9457 #10723]: #9420
-#4080 := (+ #185 #4079)
-#4439 := (<= #4080 0::Int)
-#7295 := (= #185 #4078)
-#10741 := (iff #4131 #7295)
-#10728 := (iff #7295 #4131)
-#10729 := [commutativity]: #10728
-#10742 := [symm #10729]: #10741
-#10743 := [mp #7804 #10742]: #7295
-#10747 := (not #7295)
-#10748 := (or #10747 #4439)
-#10750 := [th-lemma arith triangle-eq]: #10748
-#10752 := [unit-resolution #10750 #10743]: #4439
-#10753 := [th-lemma arith farkas -1 1 -1 1 #10752 #10721 #9712 #10735]: false
-#10749 := [lemma #10753]: #10754
-#10606 := [unit-resolution #10749 #10605]: #9236
-#4081 := (>= #4080 0::Int)
-#7188 := (or #3851 #4081)
-#7230 := [quant-inst #181]: #7188
-#9459 := [unit-resolution #7230 #7970]: #4081
-#4492 := (* -1::Int #1984)
-#5947 := (+ #4492 #5889)
-#5948 := (+ #185 #5947)
-#7908 := (<= #5948 0::Int)
-#5951 := (= #5948 0::Int)
-#5855 := (f19 f20 ?v1!16)
-#5872 := (* -1::Int #5855)
-#5907 := (+ #5872 #5889)
-#5908 := (+ #185 #5907)
-#5909 := (>= #5908 0::Int)
-#5890 := (* -1::Int #5889)
-#5891 := (+ f3 #5890)
-#5892 := (<= #5891 0::Int)
-#5914 := (or #5892 #5909)
-#5917 := (not #5914)
-#5898 := (= #1984 #5855)
-#10613 := (not #5898)
-#5873 := (+ #1984 #5872)
-#8314 := (>= #5873 0::Int)
-#9718 := (not #8314)
-#4924 := (+ #185 #5872)
-#4925 := (<= #4924 0::Int)
-#5728 := (f11 f21 ?v1!16)
-#5729 := (= #5728 f1)
-#9937 := (not #5729)
-#5735 := (or #5727 #5729)
-#9563 := (not #5735)
-#5725 := (f11 #193 ?v1!16)
-#5726 := (= #5725 f1)
-#5740 := (iff #5726 #5735)
-#9884 := (or #7518 #5740)
-#5730 := (if #5727 #4057 #5729)
-#5731 := (iff #5726 #5730)
-#9851 := (or #7518 #5731)
-#9678 := (iff #9851 #9884)
-#9677 := (iff #9884 #9884)
-#9566 := [rewrite]: #9677
-#5741 := (iff #5731 #5740)
-#5738 := (iff #5730 #5735)
-#5732 := (if #5727 true #5729)
-#5736 := (iff #5732 #5735)
-#5737 := [rewrite]: #5736
-#5733 := (iff #5730 #5732)
-#5734 := [monotonicity #4060]: #5733
-#5739 := [trans #5734 #5737]: #5738
-#5742 := [monotonicity #5739]: #5741
-#9679 := [monotonicity #5742]: #9678
-#8478 := [trans #9679 #9566]: #9678
-#9932 := [quant-inst #115 #181 #3 #1975]: #9851
-#7394 := [mp #9932 #8478]: #9884
-#10476 := [unit-resolution #7394 #3646]: #5740
-#8333 := (not #5726)
-#10578 := (iff #1979 #8333)
-#10579 := (iff #1978 #5726)
-#10551 := (iff #5726 #1978)
-#10541 := (= #5725 #1977)
-#10544 := [monotonicity #8146]: #10541
-#10552 := [monotonicity #10544]: #10551
-#10580 := [symm #10552]: #10579
-#10581 := [monotonicity #10580]: #10578
-#10477 := [hypothesis]: #1979
-#10600 := [mp #10477 #10581]: #8333
-#8954 := (not #5740)
-#9175 := (or #8954 #5726 #9563)
-#7890 := [def-axiom]: #9175
-#10601 := [unit-resolution #7890 #10600 #10476]: #9563
-#8489 := (or #5735 #9937)
-#9914 := [def-axiom]: #8489
-#10607 := [unit-resolution #9914 #10601]: #9937
-#4948 := (or #5729 #4925)
-#3462 := (or #3968 #3820)
-#3463 := [def-axiom]: #3462
-#9789 := [unit-resolution #3463 #7802]: #3820
-#9716 := (or #3825 #5729 #4925)
-#4892 := (+ #5855 #1146)
-#4893 := (>= #4892 0::Int)
-#4918 := (or #5729 #4893)
-#9628 := (or #3825 #4918)
-#10139 := (iff #9628 #9716)
-#10115 := (or #3825 #4948)
-#10118 := (iff #10115 #9716)
-#10138 := [rewrite]: #10118
-#10117 := (iff #9628 #10115)
-#4949 := (iff #4918 #4948)
-#4946 := (iff #4893 #4925)
-#4919 := (+ #1146 #5855)
-#4922 := (>= #4919 0::Int)
-#4926 := (iff #4922 #4925)
-#4927 := [rewrite]: #4926
-#4917 := (iff #4893 #4922)
-#4920 := (= #4892 #4919)
-#4921 := [rewrite]: #4920
-#4923 := [monotonicity #4921]: #4917
-#4947 := [trans #4923 #4927]: #4946
-#4950 := [monotonicity #4947]: #4949
-#10112 := [monotonicity #4950]: #10117
-#10140 := [trans #10112 #10138]: #10139
-#10116 := [quant-inst #1975]: #9628
-#10141 := [mp #10116 #10140]: #9716
-#10608 := [unit-resolution #10141 #9789]: #4948
-#10609 := [unit-resolution #10608 #10607]: #4925
-#11010 := (not #4925)
-#11011 := (or #9774 #9718 #1988 #11010)
-#11006 := [hypothesis]: #4925
-#9759 := [hypothesis]: #9236
-#9707 := [hypothesis]: #8314
-#11009 := [th-lemma arith farkas -1 1 1 -1 1 #9707 #9759 #9711 #9459 #11006]: false
-#11012 := [lemma #11009]: #11011
-#10612 := [unit-resolution #11012 #10606 #9711 #10609]: #9718
-#10614 := (or #10613 #8314)
-#10615 := [th-lemma arith triangle-eq]: #10614
-#10610 := [unit-resolution #10615 #10612]: #10613
-#5920 := (or #5917 #5898)
-#10033 := (or #3843 #5917 #5898)
-#5893 := (+ #1146 #5890)
-#5894 := (+ #5855 #5893)
-#5895 := (<= #5894 0::Int)
-#5896 := (or #5892 #5895)
-#5897 := (not #5896)
-#5899 := (or #5897 #5898)
-#10016 := (or #3843 #5899)
-#9991 := (iff #10016 #10033)
-#10017 := (or #3843 #5920)
-#10008 := (iff #10017 #10033)
-#10004 := [rewrite]: #10008
-#9989 := (iff #10016 #10017)
-#5921 := (iff #5899 #5920)
-#5918 := (iff #5897 #5917)
-#5915 := (iff #5896 #5914)
-#5912 := (iff #5895 #5909)
-#5900 := (+ #5855 #5890)
-#5901 := (+ #1146 #5900)
-#5904 := (<= #5901 0::Int)
-#5910 := (iff #5904 #5909)
-#5911 := [rewrite]: #5910
-#5905 := (iff #5895 #5904)
-#5902 := (= #5894 #5901)
-#5903 := [rewrite]: #5902
-#5906 := [monotonicity #5903]: #5905
-#5913 := [trans #5906 #5911]: #5912
-#5916 := [monotonicity #5913]: #5915
-#5919 := [monotonicity #5916]: #5918
-#5922 := [monotonicity #5919]: #5921
-#10005 := [monotonicity #5922]: #9989
-#10002 := [trans #10005 #10004]: #9991
-#9962 := [quant-inst #1975]: #10016
-#10010 := [mp #9962 #10002]: #10033
-#10611 := [unit-resolution #10010 #7803]: #5920
-#10616 := [unit-resolution #10611 #10610]: #5917
-#9749 := (or #5914 #5951)
-#9732 := (not #5951)
-#9742 := [hypothesis]: #9732
-#8264 := (not #5892)
-#9745 := [hypothesis]: #5917
-#8000 := (or #5914 #8264)
-#8092 := [def-axiom]: #8000
-#9727 := [unit-resolution #8092 #9745]: #8264
-#8284 := (not #5909)
-#8311 := (or #5914 #8284)
-#8280 := [def-axiom]: #8311
-#9746 := [unit-resolution #8280 #9745]: #8284
-#5954 := (or #5892 #5909 #5951)
-#8310 := (or #3835 #5892 #5909 #5951)
-#5943 := (+ #5889 #4492)
-#5944 := (+ #185 #5943)
-#5945 := (= #5944 0::Int)
-#5946 := (or #5892 #5895 #5945)
-#8334 := (or #3835 #5946)
-#8377 := (iff #8334 #8310)
-#8335 := (or #3835 #5954)
-#8378 := (iff #8335 #8310)
-#8379 := [rewrite]: #8378
-#8336 := (iff #8334 #8335)
-#5955 := (iff #5946 #5954)
-#5952 := (iff #5945 #5951)
-#5949 := (= #5944 #5948)
-#5950 := [rewrite]: #5949
-#5953 := [monotonicity #5950]: #5952
-#5956 := [monotonicity #5913 #5953]: #5955
-#8337 := [monotonicity #5956]: #8336
-#8371 := [trans #8337 #8379]: #8377
-#8279 := [quant-inst #1975]: #8334
-#8381 := [mp #8279 #8371]: #8310
-#9747 := [unit-resolution #8381 #7967]: #5954
-#9748 := [unit-resolution #9747 #9746 #9727 #9742]: false
-#9750 := [lemma #9748]: #9749
-#10617 := [unit-resolution #9750 #10616]: #5951
-#10618 := (or #9732 #7908)
-#10619 := [th-lemma arith triangle-eq]: #10618
-#10620 := [unit-resolution #10619 #10617]: #7908
-#10624 := (not #4081)
-#10623 := (not #7908)
-#10625 := (or #9302 #10623 #1988 #9774 #10624)
-#10626 := [th-lemma arith assign-bounds 1 1 1 1]: #10625
-#10621 := [unit-resolution #10626 #10620 #9459 #10606 #9711]: #9302
-#9349 := (not #9302)
-#10411 := (or #7600 #9301 #9349)
-#9385 := (or #9301 #9349)
-#10539 := (or #7600 #9385)
-#10540 := (iff #10539 #10411)
-#10542 := [rewrite]: #10540
-#10381 := [quant-inst #181 #1975]: #10539
-#10543 := [mp #10381 #10542]: #10411
-#10622 := [unit-resolution #10543 #3612 #10621]: #9301
-#10627 := [symm #10622]: #5727
-#7819 := (not #5727)
-#9916 := (or #5735 #7819)
-#9935 := [def-axiom]: #9916
-#10577 := [unit-resolution #9935 #10601]: #7819
-#10628 := [unit-resolution #10577 #10627]: false
-#10630 := [lemma #10628]: #10629
-#9889 := [unit-resolution #10630 #5973 #5974 #5972]: false
-#9891 := [lemma #9889]: #2865
-#8784 := (or #1962 #3860)
-#4183 := (>= #185 0::Int)
-#3442 := (or #3980 #3717)
-#3422 := [def-axiom]: #3442
-#7614 := [unit-resolution #3422 #7155]: #3717
-#7345 := (or #3722 #4183)
-#7351 := [quant-inst #181]: #7345
-#7620 := [unit-resolution #7351 #7614]: #4183
-#6537 := (f5 #195 ?v0!15)
-#6538 := (f4 #6537)
-#4488 := (* -1::Int #1961)
-#6594 := (+ #4488 #6538)
-#6595 := (+ #185 #6594)
-#8465 := (<= #6595 0::Int)
-#6598 := (= #6595 0::Int)
-#6496 := (f19 f20 ?v0!15)
-#6521 := (* -1::Int #6496)
-#6555 := (+ #6521 #6538)
-#6556 := (+ #185 #6555)
-#6557 := (>= #6556 0::Int)
-#6539 := (* -1::Int #6538)
-#6540 := (+ f3 #6539)
-#6541 := (<= #6540 0::Int)
-#6562 := (or #6541 #6557)
-#6565 := (not #6562)
-#6497 := (= #1961 #6496)
-#9674 := (not #6497)
-#6522 := (+ #1961 #6521)
-#9653 := (>= #6522 0::Int)
-#9661 := (not #9653)
-#6618 := [hypothesis]: #1963
-#9453 := (or #9661 #1962)
-#8190 := (>= #6496 0::Int)
-#8498 := (or #3722 #8190)
-#8474 := [quant-inst #1960]: #8498
-#9421 := [unit-resolution #8474 #7614]: #8190
-#9654 := [hypothesis]: #9653
-#9452 := [th-lemma arith farkas -1 1 1 #6618 #9654 #9421]: false
-#9454 := [lemma #9452]: #9453
-#8733 := [unit-resolution #9454 #6618]: #9661
-#9675 := (or #9674 #9653)
-#8499 := [hypothesis]: #9661
-#8502 := [hypothesis]: #6497
-#9676 := [th-lemma arith triangle-eq]: #9675
-#8503 := [unit-resolution #9676 #8502 #8499]: false
-#8506 := [lemma #8503]: #9675
-#8735 := [unit-resolution #8506 #8733]: #9674
-#6568 := (or #6565 #6497)
-#8443 := (or #3843 #6565 #6497)
-#6542 := (+ #1146 #6539)
-#6543 := (+ #6496 #6542)
-#6544 := (<= #6543 0::Int)
-#6545 := (or #6541 #6544)
-#6546 := (not #6545)
-#6547 := (or #6546 #6497)
-#8444 := (or #3843 #6547)
-#8433 := (iff #8444 #8443)
-#8446 := (or #3843 #6568)
-#8449 := (iff #8446 #8443)
-#8432 := [rewrite]: #8449
-#8447 := (iff #8444 #8446)
-#6569 := (iff #6547 #6568)
-#6566 := (iff #6546 #6565)
-#6563 := (iff #6545 #6562)
-#6560 := (iff #6544 #6557)
-#6548 := (+ #6496 #6539)
-#6549 := (+ #1146 #6548)
-#6552 := (<= #6549 0::Int)
-#6558 := (iff #6552 #6557)
-#6559 := [rewrite]: #6558
-#6553 := (iff #6544 #6552)
-#6550 := (= #6543 #6549)
-#6551 := [rewrite]: #6550
-#6554 := [monotonicity #6551]: #6553
-#6561 := [trans #6554 #6559]: #6560
-#6564 := [monotonicity #6561]: #6563
-#6567 := [monotonicity #6564]: #6566
-#6570 := [monotonicity #6567]: #6569
-#8448 := [monotonicity #6570]: #8447
-#8434 := [trans #8448 #8432]: #8433
-#8445 := [quant-inst #1960]: #8444
-#8435 := [mp #8445 #8434]: #8443
-#9672 := [unit-resolution #8435 #7803]: #6568
-#8736 := [unit-resolution #9672 #8735]: #6565
-#9651 := (or #6562 #6598)
-#9644 := (not #6598)
-#9645 := [hypothesis]: #9644
-#8436 := (not #6541)
-#9646 := [hypothesis]: #6565
-#8431 := (or #6562 #8436)
-#8437 := [def-axiom]: #8431
-#9647 := [unit-resolution #8437 #9646]: #8436
-#8438 := (not #6557)
-#8439 := (or #6562 #8438)
-#8440 := [def-axiom]: #8439
-#9648 := [unit-resolution #8440 #9646]: #8438
-#6601 := (or #6541 #6557 #6598)
-#8452 := (or #3835 #6541 #6557 #6598)
-#6590 := (+ #6538 #4488)
-#6591 := (+ #185 #6590)
-#6592 := (= #6591 0::Int)
-#6593 := (or #6541 #6544 #6592)
-#8453 := (or #3835 #6593)
-#8459 := (iff #8453 #8452)
-#8455 := (or #3835 #6601)
-#8457 := (iff #8455 #8452)
-#8458 := [rewrite]: #8457
-#8450 := (iff #8453 #8455)
-#6602 := (iff #6593 #6601)
-#6599 := (iff #6592 #6598)
-#6596 := (= #6591 #6595)
-#6597 := [rewrite]: #6596
-#6600 := [monotonicity #6597]: #6599
-#6603 := [monotonicity #6561 #6600]: #6602
-#8456 := [monotonicity #6603]: #8450
-#8460 := [trans #8456 #8458]: #8459
-#8454 := [quant-inst #1960]: #8453
-#8464 := [mp #8454 #8460]: #8452
-#9649 := [unit-resolution #8464 #7967]: #6601
-#9650 := [unit-resolution #9649 #9648 #9647 #9645]: false
-#9652 := [lemma #9650]: #9651
-#8729 := [unit-resolution #9652 #8736]: #6598
-#9668 := (or #9644 #8465)
-#9669 := [th-lemma arith triangle-eq]: #9668
-#8737 := [unit-resolution #9669 #8729]: #8465
-#8633 := (>= #6538 0::Int)
-#9565 := (<= #6538 0::Int)
-#9568 := (not #9565)
-#9564 := (= f28 ?v0!15)
-#9601 := (not #9564)
-#9422 := (= ?v0!15 f28)
-#9445 := (not #9422)
-#9602 := (iff #9445 #9601)
-#9625 := (iff #9422 #9564)
-#9571 := (iff #9564 #9422)
-#9624 := [commutativity]: #9571
-#9626 := [symm #9624]: #9625
-#9603 := [monotonicity #9626]: #9602
-#8093 := (f11 f21 ?v0!15)
-#8097 := (= #8093 f1)
-#9428 := (or #9422 #8097)
-#9451 := (not #9428)
-#9414 := (f11 #193 ?v0!15)
-#9415 := (= #9414 f1)
-#9433 := (iff #9415 #9428)
-#8508 := (or #7518 #9433)
-#9423 := (if #9422 #4057 #8097)
-#9424 := (iff #9415 #9423)
-#8509 := (or #7518 #9424)
-#8511 := (iff #8509 #8508)
-#8510 := (iff #8508 #8508)
-#8514 := [rewrite]: #8510
-#9434 := (iff #9424 #9433)
-#9431 := (iff #9423 #9428)
-#9425 := (if #9422 true #8097)
-#9429 := (iff #9425 #9428)
-#9430 := [rewrite]: #9429
-#9426 := (iff #9423 #9425)
-#9427 := [monotonicity #4060]: #9426
-#9432 := [trans #9427 #9430]: #9431
-#9435 := [monotonicity #9432]: #9434
-#8512 := [monotonicity #9435]: #8511
-#8515 := [trans #8512 #8514]: #8511
-#8507 := [quant-inst #115 #181 #3 #1960]: #8509
-#8513 := [mp #8507 #8515]: #8508
-#8738 := [unit-resolution #8513 #3646]: #9433
-#8520 := (not #9415)
-#4485 := (f11 f29 ?v0!15)
-#4486 := (= #4485 f1)
-#4487 := (not #4486)
-#8754 := (iff #4487 #8520)
-#8755 := (iff #4486 #9415)
-#8702 := (iff #9415 #4486)
-#8748 := (= #9414 #4485)
-#8749 := [monotonicity #8146]: #8748
-#8703 := [monotonicity #8749]: #8702
-#8756 := [symm #8703]: #8755
-#8757 := [monotonicity #8756]: #8754
-#6506 := (or #4487 #6497)
-#8739 := [hypothesis]: #3855
-#8421 := (or #3860 #4487 #6497)
-#8422 := (or #3860 #6506)
-#8406 := (iff #8422 #8421)
-#8424 := [rewrite]: #8406
-#8423 := [quant-inst #1960]: #8422
-#8405 := [mp #8423 #8424]: #8421
-#8744 := [unit-resolution #8405 #8739]: #6506
-#8745 := [unit-resolution #8744 #8735]: #4487
-#8758 := [mp #8745 #8757]: #8520
-#8518 := (not #9433)
-#8519 := (or #8518 #9415 #9451)
-#8517 := [def-axiom]: #8519
-#8767 := [unit-resolution #8517 #8758 #8738]: #9451
-#9446 := (or #9428 #9445)
-#9447 := [def-axiom]: #9446
-#8768 := [unit-resolution #9447 #8767]: #9445
-#8774 := [mp #8768 #9603]: #9601
-#9572 := (or #9564 #9568)
-#9574 := (or #7600 #9564 #9568)
-#9575 := (or #7600 #9572)
-#9579 := (iff #9575 #9574)
-#9580 := [rewrite]: #9579
-#9576 := [quant-inst #181 #1960]: #9575
-#9578 := [mp #9576 #9580]: #9574
-#9605 := [unit-resolution #9578 #3612]: #9572
-#8775 := [unit-resolution #9605 #8774]: #9568
-#8773 := (or #8633 #9565)
-#8776 := [th-lemma arith farkas 1 1]: #8773
-#8769 := [unit-resolution #8776 #8775]: #8633
-#8777 := [th-lemma arith farkas 1 1 -1 1 #6618 #8769 #8737 #7620]: false
-#8782 := [lemma #8777]: #8784
-#10054 := [unit-resolution #8782 #10053]: #1962
-#4342 := (f5 #195 f16)
-#4343 := (f4 #4342)
-#7596 := (<= #4343 0::Int)
-#7588 := (not #7596)
-#4163 := (= f28 f16)
-#7583 := (not #4163)
-#7834 := [hypothesis]: #661
-#7847 := (or #7583 #217)
-#4358 := (= #216 #110)
-#7840 := (= #185 #110)
-#7835 := [hypothesis]: #4163
-#7841 := [monotonicity #7835]: #7840
-#7842 := (= #216 #185)
-#7838 := (= #216 #4078)
-#7836 := (= f16 f28)
-#7837 := [symm #7835]: #7836
-#7839 := [monotonicity #7837]: #7838
-#7843 := [trans #7839 #7804]: #7842
-#7844 := [trans #7843 #7841]: #4358
-#7845 := [trans #7844 #6894]: #217
-#7846 := [unit-resolution #7834 #7845]: false
-#7848 := [lemma #7846]: #7847
-#7599 := [unit-resolution #7848 #7834]: #7583
-#7589 := (or #4163 #7588)
-#7601 := (or #7600 #4163 #7588)
-#7602 := (or #7600 #7589)
-#7604 := (iff #7602 #7601)
-#7605 := [rewrite]: #7604
-#7603 := [quant-inst #181 #65]: #7602
-#7606 := [mp #7603 #7605]: #7601
-#7607 := [unit-resolution #7606 #3612]: #7589
-#7598 := [unit-resolution #7607 #7599]: #7588
-#4350 := (* -1::Int #4343)
-#4353 := (+ #1146 #4350)
-#4354 := (+ #110 #4353)
-#4355 := (<= #4354 0::Int)
-#7500 := (not #4355)
-#4351 := (+ f3 #4350)
-#4352 := (<= #4351 0::Int)
-#4356 := (or #4352 #4355)
-#4357 := (not #4356)
-#7612 := (not #4358)
-#7613 := (iff #661 #7612)
-#7610 := (iff #217 #4358)
-#7608 := (iff #4358 #217)
-#7609 := [monotonicity #6894]: #7608
-#7611 := [symm #7609]: #7610
-#7615 := [monotonicity #7611]: #7613
-#7616 := [mp #7834 #7615]: #7612
-#4359 := (or #4357 #4358)
-#7491 := (or #3843 #4357 #4358)
-#7492 := (or #3843 #4359)
-#7494 := (iff #7492 #7491)
-#7495 := [rewrite]: #7494
-#7493 := [quant-inst #65]: #7492
-#7496 := [mp #7493 #7495]: #7491
-#7617 := [unit-resolution #7496 #7803]: #4359
-#7618 := [unit-resolution #7617 #7616]: #4357
-#7501 := (or #4356 #7500)
-#7502 := [def-axiom]: #7501
-#7619 := [unit-resolution #7502 #7618]: #7500
-#4504 := (<= #110 0::Int)
-#7621 := (or #804 #4504)
-#7622 := [th-lemma arith triangle-eq]: #7621
-#7623 := [unit-resolution #7622 #6894]: #4504
-#7624 := [th-lemma arith farkas 1 1 1 1 #7623 #7620 #7619 #7598]: false
-#7626 := [lemma #7624]: #217
-#3476 := (or #3944 #3938)
-#3478 := [def-axiom]: #3476
-#10055 := [unit-resolution #3478 #10052]: #3938
-#3496 := (or #3941 #661 #3935)
-#3486 := [def-axiom]: #3496
-#10056 := [unit-resolution #3486 #10055 #7626]: #3935
-#3488 := (or #3932 #3926)
-#3489 := [def-axiom]: #3488
-#10057 := [unit-resolution #3489 #10056]: #3926
-#3504 := (or #3929 #1963 #3923)
-#3484 := [def-axiom]: #3504
-#10058 := [unit-resolution #3484 #10057 #10054]: #3923
-#3507 := (or #3920 #3914)
-#3512 := [def-axiom]: #3507
-#18172 := [unit-resolution #3512 #10058]: #3914
-#3508 := (or #3917 #2870 #3911)
-#3509 := [def-axiom]: #3508
-#18183 := [unit-resolution #3509 #18172]: #3914
-#18184 := [unit-resolution #18183 #9891]: #3911
-#3517 := (or #3908 #3902)
-#3518 := [def-axiom]: #3517
-#18185 := [unit-resolution #3518 #18184]: #3902
-#18186 := (or #3905 #3899)
-#15845 := [hypothesis]: #2916
-#3530 := (or #2911 #2012)
-#3533 := [def-axiom]: #3530
-#15846 := [unit-resolution #3533 #15845]: #2012
-#6483 := (f5 #195 ?v0!19)
-#6481 := (f4 #6483)
-#6484 := (* -1::Int #6481)
-#15446 := (+ #2008 #6484)
-#15535 := (>= #15446 0::Int)
-#15445 := (= #2008 #6481)
-#15962 := (= #2007 #6483)
-#15956 := (= #2006 #195)
-#5407 := (= ?v1!18 f28)
-#5408 := (f11 f21 ?v1!18)
-#5409 := (= #5408 f1)
-#5415 := (or #5407 #5409)
-#5397 := (f11 #193 ?v1!18)
-#5398 := (= #5397 f1)
-#5420 := (iff #5398 #5415)
-#14616 := (or #7518 #5420)
-#5410 := (if #5407 #4057 #5409)
-#5411 := (iff #5398 #5410)
-#14606 := (or #7518 #5411)
-#14610 := (iff #14606 #14616)
-#14613 := (iff #14616 #14616)
-#14619 := [rewrite]: #14613
-#5421 := (iff #5411 #5420)
-#5418 := (iff #5410 #5415)
-#5412 := (if #5407 true #5409)
-#5416 := (iff #5412 #5415)
-#5417 := [rewrite]: #5416
-#5413 := (iff #5410 #5412)
-#5414 := [monotonicity #4060]: #5413
-#5419 := [trans #5414 #5417]: #5418
-#5422 := [monotonicity #5419]: #5421
-#14620 := [monotonicity #5422]: #14610
-#14630 := [trans #14620 #14619]: #14610
-#14618 := [quant-inst #115 #181 #3 #2002]: #14606
-#14662 := [mp #14618 #14630]: #14616
-#15893 := [unit-resolution #14662 #3646]: #5420
-#3434 := (or #2911 #2005)
-#3529 := [def-axiom]: #3434
-#15907 := [unit-resolution #3529 #15845]: #2005
-#15926 := (= #5397 #2004)
-#15928 := [monotonicity #8146]: #15926
-#15908 := [trans #15928 #15907]: #5398
-#11373 := (not #5398)
-#15102 := (not #5420)
-#15199 := (or #15102 #11373 #5415)
-#15197 := [def-axiom]: #15199
-#15929 := [unit-resolution #15197 #15908 #15893]: #5415
-#6623 := (not #5409)
-#6374 := (f19 f20 ?v0!19)
-#6445 := (* -1::Int #6374)
-#5362 := (f19 f20 ?v1!18)
-#6639 := (+ #5362 #6445)
-#6640 := (+ #2008 #6639)
-#6641 := (>= #6640 0::Int)
-#15933 := (not #6641)
-#3418 := (not #2020)
-#3433 := (or #2911 #3418)
-#3435 := [def-axiom]: #3433
-#15930 := [unit-resolution #3435 #15845]: #3418
-#6446 := (+ #2016 #6445)
-#6447 := (<= #6446 0::Int)
-#12992 := (not #6447)
-#12993 := [hypothesis]: #12992
-#12524 := (or #3851 #6447)
-#6389 := (+ #6374 #2017)
-#6399 := (>= #6389 0::Int)
-#12525 := (or #3851 #6399)
-#12527 := (iff #12525 #12524)
-#12530 := (iff #12524 #12524)
-#12531 := [rewrite]: #12530
-#6468 := (iff #6399 #6447)
-#6440 := (+ #2017 #6374)
-#6443 := (>= #6440 0::Int)
-#6448 := (iff #6443 #6447)
-#6467 := [rewrite]: #6448
-#6438 := (iff #6399 #6443)
-#6441 := (= #6389 #6440)
-#6442 := [rewrite]: #6441
-#6444 := [monotonicity #6442]: #6438
-#6469 := [trans #6444 #6467]: #6468
-#12529 := [monotonicity #6469]: #12527
-#12532 := [trans #12529 #12531]: #12527
-#12526 := [quant-inst #2003]: #12525
-#12533 := [mp #12526 #12532]: #12524
-#12610 := [unit-resolution #12533 #7970 #12993]: false
-#12611 := [lemma #12610]: #6447
-#5459 := (* -1::Int #5362)
-#5460 := (+ #2015 #5459)
-#12226 := (>= #5460 0::Int)
-#5380 := (= #2015 #5362)
-#12307 := (or #3860 #2896 #5380)
-#5383 := (or #2896 #5380)
-#12303 := (or #3860 #5383)
-#12279 := (iff #12303 #12307)
-#12410 := [rewrite]: #12279
-#12304 := [quant-inst #2002]: #12303
-#12398 := [mp #12304 #12410]: #12307
-#15925 := [unit-resolution #12398 #10053 #15907]: #5380
-#15889 := (not #5380)
-#15931 := (or #15889 #12226)
-#15927 := [th-lemma arith triangle-eq]: #15931
-#15932 := [unit-resolution #15927 #15925]: #12226
-#15883 := (not #12226)
-#15934 := (or #15933 #15883 #12992 #2020)
-#15935 := [th-lemma arith assign-bounds 1 -1 -1]: #15934
-#15955 := [unit-resolution #15935 #15932 #12611 #15930]: #15933
-#15957 := (or #6623 #6641)
-#15290 := (or #3739 #6623 #2011 #6641)
-#6642 := (or #6623 #2011 #6641)
-#15291 := (or #3739 #6642)
-#15288 := (iff #15291 #15290)
-#15308 := [rewrite]: #15288
-#15270 := [quant-inst #2003 #2002]: #15291
-#15282 := [mp #15270 #15308]: #15290
-#15958 := [unit-resolution #15282 #7156 #15846]: #15957
-#15954 := [unit-resolution #15958 #15955]: #6623
-#15090 := (not #5415)
-#15098 := (or #15090 #5407 #5409)
-#15091 := [def-axiom]: #15098
-#15960 := [unit-resolution #15091 #15954 #15929]: #5407
-#15961 := [monotonicity #15960]: #15956
-#15959 := [monotonicity #15961]: #15962
-#15963 := [monotonicity #15959]: #15445
-#15964 := (not #15445)
-#13461 := (or #15964 #15535)
-#11575 := [th-lemma arith triangle-eq]: #13461
-#14042 := [unit-resolution #11575 #15963]: #15535
-#6485 := (+ f3 #6484)
-#6486 := (<= #6485 0::Int)
-#6610 := (+ #2017 #6481)
-#6611 := (+ #185 #6610)
-#6614 := (= #6611 0::Int)
-#15542 := (not #6614)
-#14614 := (>= #6611 0::Int)
-#13597 := (not #14614)
-#6363 := (+ #2015 #4079)
-#6364 := (>= #6363 0::Int)
-#7961 := (= #2015 #4078)
-#14062 := [monotonicity #15960]: #7961
-#15213 := (not #7961)
-#15268 := (or #15213 #6364)
-#15479 := [th-lemma arith triangle-eq]: #15268
-#15484 := [unit-resolution #15479 #14062]: #6364
-#12694 := (not #15535)
-#14493 := (not #4439)
-#15269 := (not #6364)
-#12544 := (or #13597 #2020 #15269 #14493 #12694)
-#14974 := [th-lemma arith assign-bounds -1 1 -1 1]: #12544
-#15581 := [unit-resolution #14974 #15930 #10752 #15484 #14042]: #13597
-#15702 := (or #15542 #14614)
-#15714 := [th-lemma arith triangle-eq]: #15702
-#15715 := [unit-resolution #15714 #15581]: #15542
-#6511 := (+ #6445 #6481)
-#6509 := (+ #185 #6511)
-#6512 := (>= #6509 0::Int)
-#13034 := (not #6512)
-#15736 := (or #13034 #12992 #2020 #15269 #14493 #12694)
-#15582 := [th-lemma arith assign-bounds -1 -1 1 -1 1]: #15736
-#15841 := [unit-resolution #15582 #15930 #12611 #10752 #15484 #14042]: #13034
-#5073 := (or #6486 #6512 #6614)
-#14397 := (or #3835 #6486 #6512 #6614)
-#6607 := (+ #6481 #2017)
-#6608 := (+ #185 #6607)
-#6609 := (= #6608 0::Int)
-#6491 := (+ #1146 #6484)
-#6492 := (+ #6374 #6491)
-#6490 := (<= #6492 0::Int)
-#6604 := (or #6486 #6490 #6609)
-#12671 := (or #3835 #6604)
-#14485 := (iff #12671 #14397)
-#14423 := (or #3835 #5073)
-#14478 := (iff #14423 #14397)
-#14483 := [rewrite]: #14478
-#14389 := (iff #12671 #14423)
-#5263 := (iff #6604 #5073)
-#5033 := (iff #6609 #6614)
-#6612 := (= #6608 #6611)
-#6613 := [rewrite]: #6612
-#5034 := [monotonicity #6613]: #5033
-#6529 := (iff #6490 #6512)
-#6499 := (+ #6374 #6484)
-#6500 := (+ #1146 #6499)
-#6502 := (<= #6500 0::Int)
-#6513 := (iff #6502 #6512)
-#6514 := [rewrite]: #6513
-#6503 := (iff #6490 #6502)
-#6498 := (= #6492 #6500)
-#6501 := [rewrite]: #6498
-#6510 := [monotonicity #6501]: #6503
-#6530 := [trans #6510 #6514]: #6529
-#5264 := [monotonicity #6530 #5034]: #5263
-#14396 := [monotonicity #5264]: #14389
-#14472 := [trans #14396 #14483]: #14485
-#14422 := [quant-inst #2003]: #12671
-#14486 := [mp #14422 #14472]: #14397
-#15735 := [unit-resolution #14486 #7967]: #5073
-#15853 := [unit-resolution #15735 #15841 #15715]: #6486
-#15854 := [th-lemma arith farkas -1 1 1 #15853 #14042 #15846]: false
-#15894 := [lemma #15854]: #2911
-#3526 := (or #3905 #2916 #3899)
-#3527 := [def-axiom]: #3526
-#18187 := [unit-resolution #3527 #15894]: #18186
-#18188 := [unit-resolution #18187 #18185]: #3899
-#3441 := (or #3896 #2041)
-#3534 := [def-axiom]: #3441
-#18189 := [unit-resolution #3534 #18188]: #2041
-#15970 := (or #15968 #15969 #2040)
-#13634 := [hypothesis]: #2041
-#15965 := [hypothesis]: #12022
-#15966 := [hypothesis]: #5341
-#15967 := [th-lemma arith farkas -1 1 1 #15966 #15965 #13634]: false
-#15971 := [lemma #15967]: #15970
-#19099 := [unit-resolution #15971 #18189]: #19098
-#19100 := [unit-resolution #19099 #19097]: #15968
-#15703 := (or #5341 #5359)
-#3439 := (or #3896 #2036)
-#3440 := [def-axiom]: #3439
-#19101 := [unit-resolution #3440 #18188]: #2036
-#12207 := (or #3747 #2035 #5341 #5359)
-#5360 := (or #2035 #5341 #5359)
-#12245 := (or #3747 #5360)
-#12251 := (iff #12245 #12207)
-#12352 := [rewrite]: #12251
-#12240 := [quant-inst #2034]: #12245
-#12292 := [mp #12240 #12352]: #12207
-#19102 := [unit-resolution #12292 #7427 #19101]: #15703
-#19103 := [unit-resolution #19102 #19100]: #5359
-#15553 := (or #5358 #5356)
-#15554 := [def-axiom]: #15553
-#18351 := [unit-resolution #15554 #19103]: #5356
-#18355 := (or #5357 #15530)
-#18361 := [th-lemma arith triangle-eq]: #18355
-#18362 := [unit-resolution #18361 #18351]: #15530
-#19106 := (not #13366)
-#18364 := (not #15530)
-#18350 := (or #18243 #18364 #15969 #19106)
-#18365 := [th-lemma arith assign-bounds 1 1 1]: #18350
-#18367 := [unit-resolution #18365 #18362 #19087 #19097]: #18243
-#14134 := (<= #14210 0::Int)
-#14035 := (+ f3 #5353)
-#14131 := (<= #14035 0::Int)
-#18369 := (not #14131)
-#13911 := (>= #5343 0::Int)
-#18177 := (or #3722 #13911)
-#18180 := [quant-inst #5342]: #18177
-#18368 := [unit-resolution #18180 #7614]: #13911
-#18363 := (not #13911)
-#18372 := (or #18369 #18363 #18364 #15969)
-#18371 := (or #18369 #18363 #18364 #2040 #15969)
-#18366 := [th-lemma arith assign-bounds -1 -1 -1 -1]: #18371
-#18373 := [unit-resolution #18366 #18189]: #18372
-#18378 := [unit-resolution #18373 #18368 #18362 #19097]: #18369
-#13322 := (f11 f29 #5342)
-#13321 := (= #13322 f1)
-#15207 := (+ #2037 #13327)
-#15217 := (<= #15207 0::Int)
-#19105 := (not #15217)
-#12311 := (not #5346)
-#12351 := (or #5358 #12311)
-#12297 := [def-axiom]: #12351
-#19104 := [unit-resolution #12297 #19103]: #12311
-#19107 := (or #19105 #5346 #15969 #19106)
-#19108 := [th-lemma arith assign-bounds 1 1 1]: #19107
-#19109 := [unit-resolution #19108 #19104 #19097 #19087]: #19105
-#15548 := (or #13321 #4541 #15217)
-#3523 := (or #3908 #3872)
-#3528 := [def-axiom]: #3523
-#18201 := [unit-resolution #3528 #18184]: #3872
-#10669 := (or #3877 #13321 #4541 #15217)
-#14133 := (+ #13325 #2038)
-#15115 := (>= #14133 0::Int)
-#15556 := (or #13321 #4541 #15115)
-#10675 := (or #3877 #15556)
-#10689 := (iff #10675 #10669)
-#10671 := (or #3877 #15548)
-#10662 := (iff #10671 #10669)
-#10672 := [rewrite]: #10662
-#10679 := (iff #10675 #10671)
-#15557 := (iff #15556 #15548)
-#15222 := (iff #15115 #15217)
-#15259 := (+ #2038 #13325)
-#15255 := (>= #15259 0::Int)
-#15218 := (iff #15255 #15217)
-#15261 := [rewrite]: #15218
-#15208 := (iff #15115 #15255)
-#15253 := (= #14133 #15259)
-#15258 := [rewrite]: #15253
-#15216 := [monotonicity #15258]: #15208
-#15202 := [trans #15216 #15261]: #15222
-#15555 := [monotonicity #15202]: #15557
-#10681 := [monotonicity #15555]: #10679
-#10677 := [trans #10681 #10672]: #10689
-#10676 := [quant-inst #2034 #5342]: #10675
-#10678 := [mp #10676 #10677]: #10669
-#19114 := [unit-resolution #10678 #18201]: #15548
-#18370 := [unit-resolution #19114 #19109 #19086]: #13321
-#13323 := (not #13321)
-#11985 := (or #13323 #14131 #14134)
-#3437 := (or #3896 #3880)
-#3438 := [def-axiom]: #3437
-#18542 := [unit-resolution #3438 #18188]: #3880
-#18156 := (or #3885 #13323 #14131 #14134)
-#14205 := (+ #5352 #14133)
-#12308 := (>= #14205 0::Int)
-#11983 := (or #13323 #14131 #12308)
-#18152 := (or #3885 #11983)
-#18167 := (iff #18152 #18156)
-#18181 := (or #3885 #11985)
-#17096 := (iff #18181 #18156)
-#18166 := [rewrite]: #17096
-#18182 := (iff #18152 #18181)
-#12007 := (iff #11983 #11985)
-#14226 := (iff #12308 #14134)
-#11853 := (+ #5352 #13325)
-#14129 := (+ #2038 #11853)
-#14155 := (>= #14129 0::Int)
-#8205 := (iff #14155 #14134)
-#14211 := [rewrite]: #8205
-#14206 := (iff #12308 #14155)
-#12309 := (= #14205 #14129)
-#12299 := [rewrite]: #12309
-#14118 := [monotonicity #12299]: #14206
-#14132 := [trans #14118 #14211]: #14226
-#11847 := [monotonicity #14132]: #12007
-#17095 := [monotonicity #11847]: #18182
-#18163 := [trans #17095 #18166]: #18167
-#18175 := [quant-inst #2034 #5342]: #18152
-#18168 := [mp #18175 #18163]: #18156
-#18379 := [unit-resolution #18168 #18542]: #11985
-#18376 := [unit-resolution #18379 #18370 #18378]: #14134
-#15298 := (= #14210 0::Int)
-#15302 := (not #15298)
-#15303 := (or #15217 #13323 #15302)
-#3531 := (or #3896 #3888)
-#3535 := [def-axiom]: #3531
-#18380 := [unit-resolution #3535 #18188]: #3888
-#18219 := (or #3893 #15217 #13323 #15302)
-#15198 := (+ #2038 #5352)
-#15203 := (+ #13325 #15198)
-#15204 := (= #15203 0::Int)
-#15256 := (not #15204)
-#15257 := (or #15115 #13323 #15256)
-#18220 := (or #3893 #15257)
-#18239 := (iff #18220 #18219)
-#18247 := (or #3893 #15303)
-#18237 := (iff #18247 #18219)
-#18238 := [rewrite]: #18237
-#18253 := (iff #18220 #18247)
-#15292 := (iff #15257 #15303)
-#15304 := (iff #15256 #15302)
-#15275 := (iff #15204 #15298)
-#15285 := (= #14129 0::Int)
-#15299 := (iff #15285 #15298)
-#15297 := [rewrite]: #15299
-#15279 := (iff #15204 #15285)
-#15264 := (= #15203 #14129)
-#15254 := [rewrite]: #15264
-#15284 := [monotonicity #15254]: #15279
-#15300 := [trans #15284 #15297]: #15275
-#15289 := [monotonicity #15300]: #15304
-#15301 := [monotonicity #15202 #15289]: #15292
-#18254 := [monotonicity #15301]: #18253
-#18236 := [trans #18254 #18238]: #18239
-#18252 := [quant-inst #5342]: #18220
-#18242 := [mp #18252 #18236]: #18219
-#18381 := [unit-resolution #18242 #18380]: #15303
-#18382 := [unit-resolution #18381 #18370 #19109]: #15302
-#18393 := (not #18243)
-#18392 := (not #14134)
-#18388 := (or #15298 #18392 #18393)
-#18394 := [th-lemma arith triangle-eq]: #18388
-#18396 := [unit-resolution #18394 #18382 #18376 #18367]: false
-#18397 := [lemma #18396]: #4541
-#16386 := [mp #18397 #16391]: #9410
-#9307 := (not #9956)
-#9346 := (or #9307 #9944 #9685)
-#9347 := [def-axiom]: #9346
-#16392 := [unit-resolution #9347 #16386 #16374]: #9685
-#9688 := (or #9951 #9687)
-#9684 := [def-axiom]: #9688
-#16383 := [unit-resolution #9684 #16392]: #9687
-#16395 := [unit-resolution #16383 #16394]: false
-#16405 := [lemma #16395]: #15368
-#16379 := [hypothesis]: #15340
-#15365 := (or #7600 #15327 #15358)
-#15362 := (or #15327 #15358)
-#15360 := (or #7600 #15362)
-#15351 := (iff #15360 #15365)
-#15352 := [rewrite]: #15351
-#15364 := [quant-inst #181 #2034]: #15360
-#15353 := [mp #15364 #15352]: #15365
-#16275 := [unit-resolution #15353 #3612 #16379 #16405]: false
-#16404 := [lemma #16275]: #15358
-#16322 := [hypothesis]: #12121
-#16307 := (or #15340 #15379 #15402)
-#15400 := (or #15340 #15379 #15402 #10624)
-#15403 := [th-lemma arith assign-bounds -1 -1 1]: #15400
-#16283 := [unit-resolution #15403 #9459]: #16307
-#16319 := [unit-resolution #16283 #16322 #16404]: #15402
-#5241 := (+ #4079 #5123)
-#5242 := (+ #2037 #5241)
-#5243 := (= #5242 0::Int)
-#12167 := (>= #5242 0::Int)
-#16360 := (or #12167 #15379)
-#15382 := (or #12167 #10624 #15379)
-#14315 := [th-lemma arith assign-bounds -1 1]: #15382
-#16376 := [unit-resolution #14315 #9459]: #16360
-#16377 := [unit-resolution #16376 #16322]: #12167
-#15390 := (not #12167)
-#16258 := (or #5243 #15390)
-#12074 := (<= #5242 0::Int)
-#4339 := (+ #110 #1146)
-#7776 := (<= #4339 0::Int)
-#4340 := (>= #4339 0::Int)
-#9693 := (not #4340)
-#4164 := (?v1!7 f28)
-#4165 := (f19 f20 #4164)
-#4166 := (* -1::Int #4165)
-#4167 := (+ #185 #4166)
-#4168 := (<= #4167 0::Int)
-#7320 := (not #4168)
-#4172 := (f6 f7 #4164)
-#4173 := (f5 #4172 f28)
-#4174 := (f4 #4173)
-#4175 := (* -1::Int #4174)
-#4176 := (+ #4166 #4175)
-#4177 := (+ #185 #4176)
-#4178 := (= #4177 0::Int)
-#4179 := (not #4178)
-#4169 := (f11 f21 #4164)
-#4170 := (= #4169 f1)
-#4171 := (not #4170)
-#4180 := (or #4168 #4171 #4179)
-#4181 := (not #4180)
-#4335 := (f11 f21 f16)
-#4336 := (= #4335 f1)
-#13038 := (or #3896 #2040 #4336)
-#12906 := [hypothesis]: #3899
-#12800 := [unit-resolution #3535 #12906]: #3888
-#7302 := (not #4336)
-#15705 := [hypothesis]: #7302
-#4341 := (or #4336 #4340)
-#7485 := (or #3825 #4336 #4340)
-#7486 := (or #3825 #4341)
-#7488 := (iff #7486 #7485)
-#7489 := [rewrite]: #7488
-#7487 := [quant-inst #65]: #7486
-#7490 := [mp #7487 #7489]: #7485
-#15712 := [unit-resolution #7490 #9789]: #4341
-#15713 := [unit-resolution #15712 #15705]: #4340
-#15716 := (or #9693 #4168)
-#7756 := (>= #4165 0::Int)
-#7707 := (not #7756)
-#7708 := [hypothesis]: #7707
-#7822 := (or #3722 #7756)
-#7824 := [quant-inst #4164]: #7822
-#7859 := [unit-resolution #7824 #7614 #7708]: false
-#7860 := [lemma #7859]: #7756
-#15719 := (or #9693 #7707 #4168)
-#7972 := (not #4504)
-#15717 := (or #9693 #7972 #7707 #4168)
-#15718 := [th-lemma arith assign-bounds -1 1 1]: #15717
-#15721 := [unit-resolution #15718 #7623]: #15719
-#15722 := [unit-resolution #15721 #7860]: #15716
-#15723 := [unit-resolution #15722 #15713]: #4168
-#7326 := (or #4180 #7320)
-#7330 := [def-axiom]: #7326
-#15728 := [unit-resolution #7330 #15723]: #4180
-#7978 := (or #4163 #4181)
-#3458 := (or #3968 #1323)
-#3461 := [def-axiom]: #3458
-#7977 := [unit-resolution #3461 #7802]: #1323
-#7294 := (or #3747 #4163 #1322 #4181)
-#4182 := (or #4163 #1322 #4181)
-#7297 := (or #3747 #4182)
-#7321 := (iff #7297 #7294)
-#7323 := [rewrite]: #7321
-#7298 := [quant-inst #181]: #7297
-#7324 := [mp #7298 #7323]: #7294
-#7979 := [unit-resolution #7324 #7427 #7977]: #7978
-#15748 := [unit-resolution #7979 #15728]: #4163
-#11585 := [unit-resolution #3440 #12906]: #2036
-#15405 := (or #15374 #2035 #7583 #3893)
-#15369 := (iff #2036 #15368)
-#15350 := (iff #2035 #15327)
-#15367 := (iff #15327 #2035)
-#12076 := (= f16 ?v0!20)
-#15366 := (iff #12076 #2035)
-#15354 := [commutativity]: #15366
-#15357 := (iff #15327 #12076)
-#15317 := [monotonicity #7835]: #15357
-#15318 := [trans #15317 #15354]: #15367
-#15336 := [symm #15318]: #15350
-#15370 := [monotonicity #15336]: #15369
-#15363 := [hypothesis]: #2036
-#15373 := [mp #15363 #15370]: #15368
-#15335 := [hypothesis]: #5182
-#15375 := (or #15374 #12121)
-#15376 := [th-lemma arith triangle-eq]: #15375
-#15378 := [unit-resolution #15376 #15335]: #12121
-#15381 := [unit-resolution #14315 #15378 #9459]: #12167
-#12123 := (>= #5179 0::Int)
-#14322 := (or #15374 #12123)
-#11027 := [th-lemma arith triangle-eq]: #14322
-#15349 := [unit-resolution #11027 #15335]: #12123
-#15265 := (not #12123)
-#14563 := (or #12074 #14493 #15265)
-#15266 := [th-lemma arith assign-bounds -1 1]: #14563
-#15286 := [unit-resolution #15266 #15349 #10752]: #12074
-#14562 := (not #12074)
-#15295 := (or #5243 #14562 #15390)
-#15391 := [th-lemma arith triangle-eq]: #15295
-#15392 := [unit-resolution #15391 #15286 #15381]: #5243
-#5248 := (not #5243)
-#15394 := (or #5229 #5248)
-#4429 := (f11 f29 f28)
-#4430 := (= #4429 f1)
-#4055 := (f11 #193 f28)
-#4056 := (= #4055 f1)
-#43 := (:var 0 S1)
-#40 := (:var 2 S7)
-#41 := (f14 f15 #40)
-#42 := (f13 #41 #10)
-#44 := (f12 #42 #43)
-#3633 := (pattern #44)
-#47 := (= #43 f1)
-#45 := (f11 #44 #10)
-#46 := (= #45 f1)
-#48 := (iff #46 #47)
-#3634 := (forall (vars (?v0 S7) (?v1 S2) (?v2 S1)) (:pat #3633) #48)
-#49 := (forall (vars (?v0 S7) (?v1 S2) (?v2 S1)) #48)
-#3637 := (iff #49 #3634)
-#3635 := (iff #48 #48)
-#3636 := [refl]: #3635
-#3638 := [quant-intro #3636]: #3637
-#1494 := (~ #49 #49)
-#1521 := (~ #48 #48)
-#1522 := [refl]: #1521
-#1495 := [nnf-pos #1522]: #1494
-#341 := [asserted]: #49
-#1523 := [mp~ #341 #1495]: #49
-#3639 := [mp #1523 #3638]: #3634
-#7192 := (not #3634)
-#7205 := (or #7192 #4056)
-#4058 := (iff #4056 #4057)
-#7206 := (or #7192 #4058)
-#7227 := (iff #7206 #7205)
-#7229 := (iff #7205 #7205)
-#7200 := [rewrite]: #7229
-#4066 := (iff #4058 #4056)
-#4061 := (iff #4056 true)
-#4064 := (iff #4061 #4056)
-#4065 := [rewrite]: #4064
-#4062 := (iff #4058 #4061)
-#4063 := [monotonicity #4060]: #4062
-#4067 := [trans #4063 #4065]: #4066
-#7228 := [monotonicity #4067]: #7227
-#7201 := [trans #7228 #7200]: #7227
-#7226 := [quant-inst #115 #181 #3]: #7206
-#7199 := [mp #7226 #7201]: #7205
-#15296 := [unit-resolution #7199 #3639]: #4056
-#14014 := (= #4429 #4055)
-#15393 := [monotonicity #8144]: #14014
-#15395 := [trans #15393 #15296]: #4430
-#15396 := [hypothesis]: #3888
-#4431 := (not #4430)
-#12122 := (or #3893 #5229 #4431 #5248)
-#5219 := (+ #4078 #5178)
-#5220 := (= #5219 0::Int)
-#5221 := (not #5220)
-#5209 := (+ #4078 #2038)
-#5210 := (>= #5209 0::Int)
-#5222 := (or #5210 #4431 #5221)
-#12125 := (or #3893 #5222)
-#12149 := (iff #12125 #12122)
-#5251 := (or #5229 #4431 #5248)
-#12146 := (or #3893 #5251)
-#12083 := (iff #12146 #12122)
-#12143 := [rewrite]: #12083
-#12131 := (iff #12125 #12146)
-#5252 := (iff #5222 #5251)
-#5249 := (iff #5221 #5248)
-#5246 := (iff #5220 #5243)
-#5234 := (+ #4078 #5122)
-#5235 := (+ #2038 #5234)
-#5238 := (= #5235 0::Int)
-#5244 := (iff #5238 #5243)
-#5245 := [rewrite]: #5244
-#5239 := (iff #5220 #5238)
-#5236 := (= #5219 #5235)
-#5237 := [rewrite]: #5236
-#5240 := [monotonicity #5237]: #5239
-#5247 := [trans #5240 #5245]: #5246
-#5250 := [monotonicity #5247]: #5249
-#5232 := (iff #5210 #5229)
-#5223 := (+ #2038 #4078)
-#5226 := (>= #5223 0::Int)
-#5230 := (iff #5226 #5229)
-#5231 := [rewrite]: #5230
-#5227 := (iff #5210 #5226)
-#5224 := (= #5209 #5223)
-#5225 := [rewrite]: #5224
-#5228 := [monotonicity #5225]: #5227
-#5233 := [trans #5228 #5231]: #5232
-#5253 := [monotonicity #5233 #5250]: #5252
-#12082 := [monotonicity #5253]: #12131
-#12147 := [trans #12082 #12143]: #12149
-#12075 := [quant-inst #181]: #12125
-#12124 := [mp #12075 #12147]: #12122
-#15399 := [unit-resolution #12124 #15396 #15395]: #15394
-#15401 := [unit-resolution #15399 #15392]: #5229
-#15397 := [unit-resolution #15403 #15378 #15401 #9459]: #15340
-#15398 := [unit-resolution #15353 #3612]: #15362
-#15404 := [unit-resolution #15398 #15397 #15373]: false
-#15406 := [lemma #15404]: #15405
-#11721 := [unit-resolution #15406 #11585 #15748 #12800]: #15374
-#12105 := (not #5125)
-#5149 := (not #5146)
-#15784 := (or #15969 #4336 #2040 #2035)
-#15708 := [unit-resolution #15971 #15965 #13634]: #15968
-#15709 := [unit-resolution #12292 #7427 #15363]: #15703
-#15710 := [unit-resolution #15709 #15708]: #5359
-#12264 := (or #5358 #5348)
-#12306 := [def-axiom]: #12264
-#15453 := [unit-resolution #12306 #15710]: #5348
-#15775 := (= #4335 #5347)
-#15780 := (= #182 #5347)
-#15777 := (= f28 #5342)
-#15752 := (= f16 #5342)
-#13856 := (= #5342 f16)
-#13857 := (?v1!7 #5342)
-#13952 := (f6 f7 #13857)
-#13906 := (f5 #13952 #5342)
-#13881 := (f4 #13906)
-#13913 := (* -1::Int #13881)
-#13835 := (f19 f20 #13857)
-#13871 := (* -1::Int #13835)
-#13806 := (+ #13871 #13913)
-#11518 := (+ #5343 #13806)
-#11879 := (= #11518 0::Int)
-#13861 := (not #11879)
-#13937 := (f11 f21 #13857)
-#13943 := (= #13937 f1)
-#13938 := (not #13943)
-#13870 := (+ #5343 #13871)
-#13936 := (<= #13870 0::Int)
-#11576 := (or #13936 #13938 #13861)
-#12649 := (+ #185 #5344)
-#12802 := (>= #12649 0::Int)
-#11633 := (or #3731 #183 #5349 #12802)
-#11028 := (or #183 #5349 #12802)
-#12259 := (or #3731 #11028)
-#11754 := (iff #12259 #11633)
-#12879 := [rewrite]: #11754
-#12287 := [quant-inst #5342 #181]: #12259
-#12883 := [mp #12287 #12879]: #11633
-#15747 := [unit-resolution #12883 #8156 #8155 #15453]: #12802
-#17243 := (not #12802)
-#17244 := (or #13936 #9693 #17243)
-#17239 := [hypothesis]: #12802
-#14038 := (not #13936)
-#17240 := [hypothesis]: #14038
-#17043 := (>= #13835 0::Int)
-#17063 := (or #3722 #17043)
-#17064 := [quant-inst #13857]: #17063
-#17241 := [unit-resolution #17064 #7614]: #17043
-#9714 := [hypothesis]: #4340
-#17242 := [th-lemma arith farkas -1 1 1 1 1 #7623 #9714 #17241 #17240 #17239]: false
-#17245 := [lemma #17242]: #17244
-#15749 := [unit-resolution #17245 #15747 #15713]: #13936
-#14010 := (or #11576 #14038)
-#14095 := [def-axiom]: #14010
-#15750 := [unit-resolution #14095 #15749]: #11576
-#13847 := (+ f3 #5344)
-#13841 := (<= #13847 0::Int)
-#16579 := (not #13841)
-#15746 := [unit-resolution #12297 #15710]: #12311
-#16524 := (or #16579 #2040 #15969 #5346)
-#16576 := [hypothesis]: #13841
-#16577 := [hypothesis]: #12311
-#16578 := [th-lemma arith farkas 1 1 1 1 #13634 #15965 #16577 #16576]: false
-#16580 := [lemma #16578]: #16524
-#15711 := [unit-resolution #16580 #15965 #13634 #15746]: #16579
-#13912 := (not #11576)
-#13915 := (or #13856 #13841 #13912)
-#14033 := (or #3747 #13856 #13841 #13912)
-#14036 := (or #3747 #13915)
-#14009 := (iff #14036 #14033)
-#14008 := [rewrite]: #14009
-#14011 := [quant-inst #5342]: #14036
-#14037 := [mp #14011 #14008]: #14033
-#15720 := [unit-resolution #14037 #7427]: #13915
-#15751 := [unit-resolution #15720 #15711 #15750]: #13856
-#15753 := [symm #15751]: #15752
-#15778 := [trans #15748 #15753]: #15777
-#15754 := [monotonicity #15778]: #15780
-#15776 := (= #4335 #182)
-#15727 := [symm #15748]: #7836
-#15779 := [monotonicity #15727]: #15776
-#15782 := [trans #15779 #15754]: #15775
-#15783 := [trans #15782 #15453]: #4336
-#15781 := [unit-resolution #15705 #15783]: false
-#15786 := [lemma #15781]: #15784
-#12801 := [unit-resolution #15786 #11585 #13634 #15705]: #15969
-#12564 := [unit-resolution #12636 #12801]: #11867
-#5152 := (or #5149 #5087)
-#11970 := (or #3843 #5149 #5087)
-#5126 := (+ #1146 #5123)
-#5127 := (+ #5080 #5126)
-#5128 := (<= #5127 0::Int)
-#5129 := (or #5125 #5128)
-#5130 := (not #5129)
-#5131 := (or #5130 #5087)
-#12023 := (or #3843 #5131)
-#12067 := (iff #12023 #11970)
-#12068 := (or #3843 #5152)
-#12042 := (iff #12068 #11970)
-#12107 := [rewrite]: #12042
-#12058 := (iff #12023 #12068)
-#5153 := (iff #5131 #5152)
-#5150 := (iff #5130 #5149)
-#5147 := (iff #5129 #5146)
-#5144 := (iff #5128 #5141)
-#5132 := (+ #5080 #5123)
-#5133 := (+ #1146 #5132)
-#5136 := (<= #5133 0::Int)
-#5142 := (iff #5136 #5141)
-#5143 := [rewrite]: #5142
-#5137 := (iff #5128 #5136)
-#5134 := (= #5127 #5133)
-#5135 := [rewrite]: #5134
-#5138 := [monotonicity #5135]: #5137
-#5145 := [trans #5138 #5143]: #5144
-#5148 := [monotonicity #5145]: #5147
-#5151 := [monotonicity #5148]: #5150
-#5154 := [monotonicity #5151]: #5153
-#12062 := [monotonicity #5154]: #12058
-#12070 := [trans #12062 #12107]: #12067
-#12017 := [quant-inst #2034]: #12023
-#12084 := [mp #12017 #12070]: #11970
-#9728 := [unit-resolution #12084 #7803]: #5152
-#12860 := [unit-resolution #9728 #12564]: #5149
-#12072 := (or #5146 #12105)
-#12071 := [def-axiom]: #12072
-#13044 := [unit-resolution #12071 #12860]: #12105
-#12081 := (not #5141)
-#12073 := (or #5146 #12081)
-#12118 := [def-axiom]: #12073
-#13025 := [unit-resolution #12118 #12860]: #12081
-#5185 := (or #5125 #5141 #5182)
-#12077 := (or #3835 #5125 #5141 #5182)
-#5174 := (+ #5122 #2038)
-#5175 := (+ #185 #5174)
-#5176 := (= #5175 0::Int)
-#5177 := (or #5125 #5128 #5176)
-#12079 := (or #3835 #5177)
-#12130 := (iff #12079 #12077)
-#12113 := (or #3835 #5185)
-#12114 := (iff #12113 #12077)
-#12129 := [rewrite]: #12114
-#12117 := (iff #12079 #12113)
-#5186 := (iff #5177 #5185)
-#5183 := (iff #5176 #5182)
-#5180 := (= #5175 #5179)
-#5181 := [rewrite]: #5180
-#5184 := [monotonicity #5181]: #5183
-#5187 := [monotonicity #5145 #5184]: #5186
-#12115 := [monotonicity #5187]: #12117
-#12109 := [trans #12115 #12129]: #12130
-#12120 := [quant-inst #2034]: #12079
-#12128 := [mp #12120 #12109]: #12077
-#11095 := [unit-resolution #12128 #7967]: #5185
-#13026 := [unit-resolution #11095 #13025 #13044 #11721]: false
-#11841 := [lemma #13026]: #13038
-#18190 := [unit-resolution #11841 #18188 #18189]: #4336
-#9665 := (or #7583 #7302)
-#8260 := [hypothesis]: #4336
-#9629 := (= #182 #4335)
-#9567 := [monotonicity #7835]: #9629
-#7820 := [trans #9567 #8260]: #183
-#7888 := [unit-resolution #8155 #7820]: false
-#7765 := [lemma #7888]: #9665
-#18191 := [unit-resolution #7765 #18190]: #7583
-#18192 := [unit-resolution #7979 #18191]: #4181
-#18193 := [unit-resolution #7330 #18192]: #7320
-#16378 := [unit-resolution #15722 #18193]: #9693
-#16277 := (or #7776 #4340)
-#16270 := [th-lemma arith farkas 1 1]: #16277
-#16406 := [unit-resolution #16270 #16378]: #7776
-#9345 := (not #7776)
-#12779 := (or #12074 #9345 #2040 #3885)
-#12732 := [hypothesis]: #14562
-#12904 := (or #5125 #12074)
-#12870 := [hypothesis]: #3880
-#12757 := (or #3885 #4431 #5125 #12074)
-#15309 := (+ #5122 #5209)
-#15305 := (>= #15309 0::Int)
-#15310 := (or #4431 #5125 #15305)
-#12733 := (or #3885 #15310)
-#12657 := (iff #12733 #12757)
-#15344 := (or #4431 #5125 #12074)
-#12731 := (or #3885 #15344)
-#12778 := (iff #12731 #12757)
-#12664 := [rewrite]: #12778
-#12760 := (iff #12733 #12731)
-#15345 := (iff #15310 #15344)
-#15319 := (iff #15305 #12074)
-#15287 := (>= #5235 0::Int)
-#15315 := (iff #15287 #12074)
-#15321 := [rewrite]: #15315
-#15312 := (iff #15305 #15287)
-#15307 := (= #15309 #5235)
-#15311 := [rewrite]: #15307
-#15313 := [monotonicity #15311]: #15312
-#15342 := [trans #15313 #15321]: #15319
-#15343 := [monotonicity #15342]: #15345
-#12770 := [monotonicity #15343]: #12760
-#12730 := [trans #12770 #12664]: #12657
-#12432 := [quant-inst #2034 #181]: #12733
-#12772 := [mp #12432 #12730]: #12757
-#12909 := [unit-resolution #12772 #12870 #15395]: #12904
-#12907 := [unit-resolution #12909 #12732]: #5125
-#12775 := (or #12167 #12074)
-#12881 := [th-lemma arith farkas 1 1]: #12775
-#12865 := [unit-resolution #12881 #12732]: #12167
-#9736 := [hypothesis]: #7776
-#3233 := (>= #110 0::Int)
-#4506 := (or #3722 #3233)
-#4549 := [quant-inst #65]: #4506
-#9790 := [unit-resolution #4549 #7614]: #3233
-#12861 := [th-lemma arith farkas 1 1 1 1 -1 1 #10752 #9790 #9736 #12865 #12907 #13634]: false
-#12858 := [lemma #12861]: #12779
-#16384 := [unit-resolution #12858 #16406 #18189 #18542]: #12074
-#16235 := [unit-resolution #15391 #16384]: #16258
-#16285 := [unit-resolution #16235 #16377]: #5243
-#16416 := [unit-resolution #12124 #18380]: #5251
-#16417 := [unit-resolution #16416 #15395]: #15394
-#16408 := [unit-resolution #16417 #16285 #16319]: false
-#16407 := [lemma #16408]: #15379
-#16321 := [unit-resolution #15376 #16407]: #15374
-#16309 := [hypothesis]: #5149
-#16410 := [unit-resolution #12071 #16309]: #12105
-#16431 := [unit-resolution #12118 #16309]: #12081
-#16434 := [unit-resolution #11095 #16431 #16410 #16321]: false
-#16432 := [lemma #16434]: #5146
-#16614 := [unit-resolution #9728 #16432]: #5087
-#16615 := [unit-resolution #12636 #16614]: #12022
-#16603 := [unit-resolution #19099 #16615]: #15968
-#16597 := [unit-resolution #19102 #16603]: #5359
-#16549 := [unit-resolution #15554 #16597]: #5356
-#16573 := [unit-resolution #18361 #16549]: #15530
-#16590 := [unit-resolution #18365 #19087 #16615 #16573]: #18243
-#16554 := [unit-resolution #18373 #18368 #16573 #16615]: #18369
-#16552 := (f11 #193 #5342)
-#16553 := (= #16552 f1)
-#16559 := (= #5342 f28)
-#16561 := (or #16559 #5348)
-#16545 := (iff #16553 #16561)
-#16568 := (or #7518 #16545)
-#16522 := (if #16559 #4057 #5348)
-#16560 := (iff #16553 #16522)
-#16569 := (or #7518 #16560)
-#16571 := (iff #16569 #16568)
-#16563 := (iff #16568 #16568)
-#16588 := [rewrite]: #16563
-#16567 := (iff #16560 #16545)
-#16564 := (iff #16522 #16561)
-#16557 := (if #16559 true #5348)
-#16565 := (iff #16557 #16561)
-#16547 := [rewrite]: #16565
-#16562 := (iff #16522 #16557)
-#16548 := [monotonicity #4060]: #16562
-#16566 := [trans #16548 #16547]: #16564
-#16546 := [monotonicity #16566]: #16567
-#16574 := [monotonicity #16546]: #16571
-#16586 := [trans #16574 #16588]: #16571
-#16572 := [quant-inst #115 #181 #3 #5342]: #16569
-#16589 := [mp #16572 #16586]: #16568
-#16611 := [unit-resolution #16589 #3646]: #16545
-#16373 := (not #16545)
-#16635 := (or #16373 #16553)
-#16598 := [unit-resolution #12306 #16597]: #5348
-#15874 := (or #16561 #5349)
-#16585 := [def-axiom]: #15874
-#16596 := [unit-resolution #16585 #16598]: #16561
-#15886 := (not #16561)
-#16515 := (or #16373 #16553 #15886)
-#16581 := [def-axiom]: #16515
-#16599 := [unit-resolution #16581 #16596]: #16635
-#16620 := [unit-resolution #16599 #16611]: #16553
-#16621 := (= #13322 #16552)
-#16619 := [monotonicity #8144]: #16621
-#16595 := [trans #16619 #16620]: #13321
-#16610 := [hypothesis]: #13323
-#16622 := [unit-resolution #16610 #16595]: false
-#16600 := [lemma #16622]: #13321
-#16372 := [unit-resolution #18379 #16600 #16554]: #14134
-#16550 := [unit-resolution #12297 #16597]: #12311
-#16043 := [unit-resolution #19108 #19087 #16615 #16550]: #19105
-#16044 := [unit-resolution #18381 #16600 #16043]: #15302
-[unit-resolution #18394 #16044 #16372 #16590]: false
-unsat
--- a/src/HOL/Boogie/Examples/Boogie_Dijkstra.thy	Tue Jul 23 13:14:14 2013 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,93 +0,0 @@
-(*  Title:      HOL/Boogie/Examples/Boogie_Dijkstra.thy
-    Author:     Sascha Boehme, TU Muenchen
-*)
-
-header {* Boogie example: Dijkstra's algorithm *}
-
-theory Boogie_Dijkstra
-imports Boogie
-begin
-
-text {*
-We prove correct the verification condition generated from the following
-Boogie code:
-
-\begin{verbatim}
-type Vertex;
-const G: [Vertex, Vertex] int;
-axiom (forall x: Vertex, y: Vertex ::  x != y ==> 0 < G[x,y]);
-axiom (forall x: Vertex, y: Vertex ::  x == y ==> G[x,y] == 0);
-
-const Infinity: int;
-axiom 0 < Infinity;
-
-const Source: Vertex;
-var SP: [Vertex] int;
-
-procedure Dijkstra();
-  modifies SP;
-  ensures (SP[Source] == 0);
-  ensures (forall z: Vertex, y: Vertex ::
-    SP[y] < Infinity && G[y,z] < Infinity ==> SP[z] <= SP[y] + G[y,z]);
-  ensures (forall z: Vertex :: z != Source && SP[z] < Infinity ==>
-    (exists y: Vertex :: SP[y] < SP[z] && SP[z] == SP[y] + G[y,z]));
-
-implementation Dijkstra()
-{
-  var v: Vertex;
-  var Visited: [Vertex] bool;
-  var oldSP: [Vertex] int;
-
-  havoc SP;
-  assume (forall x: Vertex :: x == Source ==> SP[x] == 0);
-  assume (forall x: Vertex :: x != Source ==> SP[x] == Infinity);
-
-  havoc Visited;
-  assume (forall x: Vertex :: !Visited[x]);
-
-  while ((exists x: Vertex :: !Visited[x] && SP[x] < Infinity))
-    invariant (SP[Source] == 0);
-    invariant (forall x: Vertex :: SP[x] >= 0);
-    invariant (forall y: Vertex, z: Vertex :: 
-      !Visited[z] && Visited[y] ==> SP[y] <= SP[z]);
-    invariant (forall z: Vertex, y: Vertex ::
-      Visited[y] && G[y,z] < Infinity ==> SP[z] <= SP[y] + G[y,z]);
-    invariant (forall z: Vertex :: z != Source && SP[z] < Infinity ==>
-      (exists y: Vertex :: SP[y] < SP[z] && Visited[y] && 
-        SP[z] == SP[y] + G[y,z]));
-  {
-    havoc v;
-    assume (!Visited[v]);
-    assume (SP[v] < Infinity); 
-    assume (forall x: Vertex :: !Visited[x] ==> SP[v] <= SP[x]);
-
-    Visited[v] := true;
-
-    oldSP := SP;
-    havoc SP;
-    assume (forall u: Vertex :: 
-      G[v,u] < Infinity && oldSP[v] + G[v,u] < oldSP[u] ==> 
-        SP[u] == oldSP[v] + G[v,u]);
-    assume (forall u: Vertex :: 
-      !(G[v,u] < Infinity && oldSP[v] + G[v,u] < oldSP[u]) ==> 
-        SP[u] == oldSP[u]);
-    assert (forall z: Vertex:: SP[z] <= oldSP[z]);
-    assert (forall y: Vertex:: Visited[y] ==> SP[y] == oldSP[y]);
-  }
-}
-\end{verbatim}
-*}
-
-
-boogie_open "Boogie_Dijkstra.b2i"
-
-declare [[smt_oracle = false]]
-declare [[smt_certificates = "Boogie_Dijkstra.certs"]]
-declare [[smt_read_only_certificates = true]]
-
-boogie_vc Dijkstra
-  by boogie
-
-boogie_end
-
-end
--- a/src/HOL/Boogie/Examples/Boogie_Max.b2i	Tue Jul 23 13:14:14 2013 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,748 +0,0 @@
-vc max 1
-    implies
-    label pos 10 7
-    true
-    implies
-    <
-    int-num 0
-    var length
-      int
-    implies
-    true
-    implies
-    =
-    var max@0
-      int
-    select 2
-    var array
-      array 2
-        int
-        int
-    int-num 0
-    implies
-    and 4
-    <=
-    int-num 0
-    int-num 0
-    <=
-    int-num 0
-    int-num 0
-    <=
-    int-num 1
-    int-num 1
-    <=
-    int-num 1
-    int-num 1
-    and 2
-    label neg 14 5
-    forall 1 0 3
-      var i
-        int
-      attribute qid 1
-        string-attr BoogieMa.14:23
-      attribute uniqueId 1
-        string-attr 2
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 2
-    <=
-    int-num 0
-    var i
-      int
-    <
-    var i
-      int
-    int-num 1
-    <=
-    select 2
-    var array
-      array 2
-        int
-        int
-    var i
-      int
-    var max@0
-      int
-    implies
-    forall 1 0 3
-      var i
-        int
-      attribute qid 1
-        string-attr BoogieMa.14:23
-      attribute uniqueId 1
-        string-attr 2
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 2
-    <=
-    int-num 0
-    var i
-      int
-    <
-    var i
-      int
-    int-num 1
-    <=
-    select 2
-    var array
-      array 2
-        int
-        int
-    var i
-      int
-    var max@0
-      int
-    and 2
-    label neg 15 5
-    =
-    select 2
-    var array
-      array 2
-        int
-        int
-    int-num 0
-    var max@0
-      int
-    implies
-    =
-    select 2
-    var array
-      array 2
-        int
-        int
-    int-num 0
-    var max@0
-      int
-    implies
-    label pos 13 3
-    true
-    implies
-    and 2
-    <=
-    int-num 0
-    var k@0
-      int
-    <=
-    int-num 1
-    var p@0
-      int
-    implies
-    forall 1 0 3
-      var i
-        int
-      attribute qid 1
-        string-attr BoogieMa.14:23
-      attribute uniqueId 1
-        string-attr 2
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 2
-    <=
-    int-num 0
-    var i
-      int
-    <
-    var i
-      int
-    var p@0
-      int
-    <=
-    select 2
-    var array
-      array 2
-        int
-        int
-    var i
-      int
-    var max@1
-      int
-    implies
-    =
-    select 2
-    var array
-      array 2
-        int
-        int
-    var k@0
-      int
-    var max@1
-      int
-    implies
-    and 2
-    <=
-    int-num 0
-    var k@0
-      int
-    <=
-    int-num 1
-    var p@0
-      int
-    and 2
-    implies
-    label pos 13 3
-    true
-    implies
-    and 2
-    <=
-    int-num 0
-    var k@0
-      int
-    <=
-    int-num 1
-    var p@0
-      int
-    implies
-    >=
-    var p@0
-      int
-    var length
-      int
-    implies
-    and 2
-    <=
-    int-num 0
-    var k@0
-      int
-    <=
-    int-num 1
-    var p@0
-      int
-    implies
-    label pos 0 0
-    true
-    implies
-    =
-    var k@2
-      int
-    var k@0
-      int
-    implies
-    =
-    var max@4
-      int
-    var max@1
-      int
-    implies
-    =
-    var p@2
-      int
-    var p@0
-      int
-    implies
-    label pos 0 0
-    true
-    and 2
-    label neg 5 3
-    exists 1 0 3
-      var i
-        int
-      attribute qid 1
-        string-attr BoogieMa.5:19
-      attribute uniqueId 1
-        string-attr 1
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 2
-    <=
-    int-num 0
-    var i
-      int
-    <
-    var i
-      int
-    var length
-      int
-    =
-    select 2
-    var array
-      array 2
-        int
-        int
-    var i
-      int
-    var max@4
-      int
-    implies
-    exists 1 0 3
-      var i
-        int
-      attribute qid 1
-        string-attr BoogieMa.5:19
-      attribute uniqueId 1
-        string-attr 1
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 2
-    <=
-    int-num 0
-    var i
-      int
-    <
-    var i
-      int
-    var length
-      int
-    =
-    select 2
-    var array
-      array 2
-        int
-        int
-    var i
-      int
-    var max@4
-      int
-    and 2
-    label neg 4 3
-    forall 1 0 3
-      var i
-        int
-      attribute qid 1
-        string-attr BoogieMa.4:19
-      attribute uniqueId 1
-        string-attr 0
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 2
-    <=
-    int-num 0
-    var i
-      int
-    <
-    var i
-      int
-    var length
-      int
-    <=
-    select 2
-    var array
-      array 2
-        int
-        int
-    var i
-      int
-    var max@4
-      int
-    implies
-    forall 1 0 3
-      var i
-        int
-      attribute qid 1
-        string-attr BoogieMa.4:19
-      attribute uniqueId 1
-        string-attr 0
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 2
-    <=
-    int-num 0
-    var i
-      int
-    <
-    var i
-      int
-    var length
-      int
-    <=
-    select 2
-    var array
-      array 2
-        int
-        int
-    var i
-      int
-    var max@4
-      int
-    true
-    implies
-    label pos 17 5
-    true
-    implies
-    and 2
-    <=
-    int-num 0
-    var k@0
-      int
-    <=
-    int-num 1
-    var p@0
-      int
-    implies
-    <
-    var p@0
-      int
-    var length
-      int
-    implies
-    and 2
-    <=
-    int-num 0
-    var k@0
-      int
-    <=
-    int-num 1
-    var p@0
-      int
-    and 2
-    implies
-    label pos 17 31
-    true
-    implies
-    and 2
-    <=
-    int-num 0
-    var k@0
-      int
-    <=
-    int-num 1
-    var p@0
-      int
-    implies
-    >
-    select 2
-    var array
-      array 2
-        int
-        int
-    var p@0
-      int
-    var max@1
-      int
-    implies
-    =
-    var max@2
-      int
-    select 2
-    var array
-      array 2
-        int
-        int
-    var p@0
-      int
-    implies
-    and 2
-    <=
-    int-num 1
-    var p@0
-      int
-    <=
-    int-num 1
-    var p@0
-      int
-    implies
-    label pos 0 0
-    true
-    implies
-    =
-    var k@1
-      int
-    var p@0
-      int
-    implies
-    =
-    var max@3
-      int
-    var max@2
-      int
-    implies
-    label pos 18 7
-    true
-    implies
-    and 2
-    <=
-    int-num 0
-    var k@1
-      int
-    <=
-    int-num 1
-    var p@0
-      int
-    implies
-    =
-    var p@1
-      int
-    +
-    var p@0
-      int
-    int-num 1
-    implies
-    and 2
-    <=
-    int-num 0
-    var k@1
-      int
-    <=
-    int-num 2
-    var p@1
-      int
-    implies
-    label pos 0 0
-    true
-    and 2
-    label neg 14 5
-    forall 1 0 3
-      var i
-        int
-      attribute qid 1
-        string-attr BoogieMa.14:23
-      attribute uniqueId 1
-        string-attr 2
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 2
-    <=
-    int-num 0
-    var i
-      int
-    <
-    var i
-      int
-    var p@1
-      int
-    <=
-    select 2
-    var array
-      array 2
-        int
-        int
-    var i
-      int
-    var max@3
-      int
-    implies
-    forall 1 0 3
-      var i
-        int
-      attribute qid 1
-        string-attr BoogieMa.14:23
-      attribute uniqueId 1
-        string-attr 2
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 2
-    <=
-    int-num 0
-    var i
-      int
-    <
-    var i
-      int
-    var p@1
-      int
-    <=
-    select 2
-    var array
-      array 2
-        int
-        int
-    var i
-      int
-    var max@3
-      int
-    and 2
-    label neg 15 5
-    =
-    select 2
-    var array
-      array 2
-        int
-        int
-    var k@1
-      int
-    var max@3
-      int
-    implies
-    =
-    select 2
-    var array
-      array 2
-        int
-        int
-    var k@1
-      int
-    var max@3
-      int
-    implies
-    false
-    true
-    implies
-    label pos 17 5
-    true
-    implies
-    and 2
-    <=
-    int-num 0
-    var k@0
-      int
-    <=
-    int-num 1
-    var p@0
-      int
-    implies
-    <=
-    select 2
-    var array
-      array 2
-        int
-        int
-    var p@0
-      int
-    var max@1
-      int
-    implies
-    and 2
-    <=
-    int-num 0
-    var k@0
-      int
-    <=
-    int-num 1
-    var p@0
-      int
-    implies
-    label pos 0 0
-    true
-    implies
-    =
-    var k@1
-      int
-    var k@0
-      int
-    implies
-    =
-    var max@3
-      int
-    var max@1
-      int
-    implies
-    label pos 18 7
-    true
-    implies
-    and 2
-    <=
-    int-num 0
-    var k@1
-      int
-    <=
-    int-num 1
-    var p@0
-      int
-    implies
-    =
-    var p@1
-      int
-    +
-    var p@0
-      int
-    int-num 1
-    implies
-    and 2
-    <=
-    int-num 0
-    var k@1
-      int
-    <=
-    int-num 2
-    var p@1
-      int
-    implies
-    label pos 0 0
-    true
-    and 2
-    label neg 14 5
-    forall 1 0 3
-      var i
-        int
-      attribute qid 1
-        string-attr BoogieMa.14:23
-      attribute uniqueId 1
-        string-attr 2
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 2
-    <=
-    int-num 0
-    var i
-      int
-    <
-    var i
-      int
-    var p@1
-      int
-    <=
-    select 2
-    var array
-      array 2
-        int
-        int
-    var i
-      int
-    var max@3
-      int
-    implies
-    forall 1 0 3
-      var i
-        int
-      attribute qid 1
-        string-attr BoogieMa.14:23
-      attribute uniqueId 1
-        string-attr 2
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 2
-    <=
-    int-num 0
-    var i
-      int
-    <
-    var i
-      int
-    var p@1
-      int
-    <=
-    select 2
-    var array
-      array 2
-        int
-        int
-    var i
-      int
-    var max@3
-      int
-    and 2
-    label neg 15 5
-    =
-    select 2
-    var array
-      array 2
-        int
-        int
-    var k@1
-      int
-    var max@3
-      int
-    implies
-    =
-    select 2
-    var array
-      array 2
-        int
-        int
-    var k@1
-      int
-    var max@3
-      int
-    implies
-    false
-    true
--- a/src/HOL/Boogie/Examples/Boogie_Max.certs	Tue Jul 23 13:14:14 2013 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,2079 +0,0 @@
-43f8affa0232e904defc8acf7ee4c00d186b27a1 2078 0
-#2 := false
-#7 := 0::Int
-decl ?v0!3 :: Int
-#1093 := ?v0!3
-#616 := -1::Int
-#1238 := (* -1::Int ?v0!3)
-decl f8 :: Int
-#31 := f8
-#2314 := (+ f8 #1238)
-#2315 := (<= #2314 0::Int)
-#2420 := (not #2315)
-#2347 := (>= #2314 0::Int)
-decl f16 :: Int
-#79 := f16
-#725 := (* -1::Int f16)
-#757 := (+ f8 #725)
-#1758 := (>= #757 -1::Int)
-#756 := (= #757 -1::Int)
-decl f15 :: Int
-#75 := f15
-decl f5 :: (-> S2 Int Int)
-decl f14 :: Int
-#73 := f14
-decl f6 :: S2
-#11 := f6
-#90 := (f5 f6 f14)
-#91 := (= #90 f15)
-#20 := (:var 0 Int)
-#24 := (f5 f6 #20)
-#2076 := (pattern #24)
-#736 := (* -1::Int f15)
-#737 := (+ #24 #736)
-#738 := (<= #737 0::Int)
-#726 := (+ #20 #725)
-#724 := (>= #726 0::Int)
-#605 := (>= #20 0::Int)
-#1341 := (not #605)
-#1515 := (or #1341 #724 #738)
-#2110 := (forall (vars (?v0 Int)) (:pat #2076) #1515)
-#2115 := (not #2110)
-#2118 := (or #2115 #91)
-#2121 := (not #2118)
-#1100 := (f5 f6 ?v0!3)
-#1260 := (* -1::Int #1100)
-#1261 := (+ f15 #1260)
-#1262 := (>= #1261 0::Int)
-#1239 := (+ f16 #1238)
-#1240 := (<= #1239 0::Int)
-#1094 := (>= ?v0!3 0::Int)
-#1478 := (not #1094)
-#1493 := (or #1478 #1240 #1262)
-#1498 := (not #1493)
-#2124 := (or #1498 #2121)
-#2127 := (not #2124)
-#82 := 2::Int
-#716 := (>= f16 2::Int)
-#1540 := (not #716)
-#713 := (>= f14 0::Int)
-#1539 := (not #713)
-#760 := (not #756)
-#15 := 1::Int
-#635 := (>= f8 1::Int)
-#769 := (not #635)
-decl f9 :: Int
-#36 := f9
-#108 := (= f15 f9)
-#395 := (not #108)
-decl f7 :: Int
-#29 := f7
-#107 := (= f14 f7)
-#404 := (not #107)
-#68 := (f5 f6 f8)
-#774 := (* -1::Int #68)
-#775 := (+ f9 #774)
-#773 := (>= #775 0::Int)
-#772 := (not #773)
-#632 := (>= f7 0::Int)
-#1470 := (not #632)
-#2136 := (or #1470 #772 #404 #395 #769 #760 #1539 #1540 #2127)
-#2139 := (not #2136)
-decl f13 :: Int
-#70 := f13
-#76 := (= f15 f13)
-#335 := (not #76)
-#74 := (= f14 f8)
-#344 := (not #74)
-#71 := (= f13 #68)
-#362 := (not #71)
-#2130 := (or #1470 #773 #362 #344 #335 #769 #760 #1539 #1540 #2127)
-#2133 := (not #2130)
-#2142 := (or #2133 #2139)
-#2145 := (not #2142)
-#681 := (* -1::Int f8)
-decl f3 :: Int
-#8 := f3
-#682 := (+ f3 #681)
-#683 := (<= #682 0::Int)
-#2148 := (or #1470 #769 #683 #2145)
-#2151 := (not #2148)
-decl ?v0!2 :: Int
-#1038 := ?v0!2
-#1046 := (f5 f6 ?v0!2)
-#1191 := (* -1::Int #1046)
-decl f11 :: Int
-#45 := f11
-#1192 := (+ f11 #1191)
-#1193 := (>= #1192 0::Int)
-#1040 := (* -1::Int ?v0!2)
-#1041 := (+ f3 #1040)
-#1042 := (<= #1041 0::Int)
-#1039 := (>= ?v0!2 0::Int)
-#1431 := (not #1039)
-decl ?v0!1 :: Int
-#1020 := ?v0!1
-#1028 := (f5 f6 ?v0!1)
-#1029 := (= #1028 f11)
-#1022 := (* -1::Int ?v0!1)
-#1023 := (+ f3 #1022)
-#1024 := (<= #1023 0::Int)
-#1021 := (>= ?v0!1 0::Int)
-#1411 := (not #1021)
-#1426 := (or #1411 #1024 #1029)
-#1457 := (not #1426)
-#1458 := (or #1457 #1431 #1042 #1193)
-#1459 := (not #1458)
-#51 := (= #24 f11)
-#643 := (* -1::Int #20)
-#644 := (+ f3 #643)
-#645 := (<= #644 0::Int)
-#1400 := (or #1341 #645 #51)
-#1405 := (not #1400)
-#2093 := (forall (vars (?v0 Int)) (:pat #2076) #1405)
-#2098 := (or #2093 #1459)
-#2101 := (not #2098)
-decl f12 :: Int
-#47 := f12
-#48 := (= f12 f8)
-#212 := (not #48)
-#46 := (= f11 f9)
-#221 := (not #46)
-decl f10 :: Int
-#43 := f10
-#44 := (= f10 f7)
-#230 := (not #44)
-#686 := (not #683)
-#2104 := (or #1470 #769 #686 #230 #221 #212 #2101)
-#2107 := (not #2104)
-#2154 := (or #2107 #2151)
-#2157 := (not #2154)
-#40 := (f5 f6 f7)
-#41 := (= #40 f9)
-#491 := (not #41)
-#881 := (* -1::Int f9)
-#882 := (+ #24 #881)
-#883 := (<= #882 0::Int)
-#871 := (+ #20 #681)
-#870 := (>= #871 0::Int)
-#1378 := (or #1341 #870 #883)
-#2085 := (forall (vars (?v0 Int)) (:pat #2076) #1378)
-#2090 := (not #2085)
-decl f4 :: Int
-#10 := f4
-#12 := (f5 f6 0::Int)
-#28 := (= #12 f4)
-#524 := (not #28)
-#2160 := (or #524 #1470 #769 #2090 #491 #2157)
-#2163 := (not #2160)
-#2166 := (or #524 #2163)
-#2169 := (not #2166)
-#619 := (* -1::Int #24)
-#620 := (+ f4 #619)
-#618 := (>= #620 0::Int)
-#608 := (>= #20 1::Int)
-#1356 := (or #1341 #608 #618)
-#2077 := (forall (vars (?v0 Int)) (:pat #2076) #1356)
-#2082 := (not #2077)
-#2172 := (or #2082 #2169)
-#2175 := (not #2172)
-decl ?v0!0 :: Int
-#968 := ?v0!0
-#962 := (f5 f6 ?v0!0)
-#963 := (* -1::Int #962)
-#960 := (+ f4 #963)
-#961 := (>= #960 0::Int)
-#970 := (>= ?v0!0 1::Int)
-#969 := (>= ?v0!0 0::Int)
-#1033 := (not #969)
-#1333 := (or #1033 #970 #961)
-#1927 := (= f4 #962)
-#1853 := (= #12 #962)
-#1893 := (= #962 #12)
-#1903 := (= ?v0!0 0::Int)
-#971 := (not #970)
-#1338 := (not #1333)
-#1926 := [hypothesis]: #1338
-#1663 := (or #1333 #971)
-#1748 := [def-axiom]: #1663
-#1928 := [unit-resolution #1748 #1926]: #971
-#1662 := (or #1333 #969)
-#1747 := [def-axiom]: #1662
-#1896 := [unit-resolution #1747 #1926]: #969
-#1862 := [th-lemma arith eq-propagate 0 0 #1896 #1928]: #1903
-#1895 := [monotonicity #1862]: #1893
-#1854 := [symm #1895]: #1853
-#13 := (= f4 #12)
-#727 := (not #724)
-#730 := (and #605 #727)
-#733 := (not #730)
-#741 := (or #733 #738)
-#744 := (forall (vars (?v0 Int)) #741)
-#747 := (not #744)
-#750 := (or #747 #91)
-#753 := (and #744 #750)
-#718 := (and #713 #716)
-#721 := (not #718)
-#763 := (and #713 #635)
-#766 := (not #763)
-#637 := (and #632 #635)
-#640 := (not #637)
-#836 := (or #640 #772 #404 #395 #766 #760 #721 #753)
-#812 := (or #640 #773 #362 #769 #344 #335 #766 #760 #721 #753)
-#841 := (and #812 #836)
-#862 := (or #640 #683 #841)
-#664 := (* -1::Int f11)
-#665 := (+ #24 #664)
-#666 := (<= #665 0::Int)
-#646 := (not #645)
-#649 := (and #605 #646)
-#652 := (not #649)
-#669 := (or #652 #666)
-#672 := (forall (vars (?v0 Int)) #669)
-#655 := (or #652 #51)
-#658 := (exists (vars (?v0 Int)) #655)
-#661 := (not #658)
-#675 := (or #661 #672)
-#678 := (and #658 #675)
-#707 := (or #640 #686 #230 #221 #212 #678)
-#867 := (and #707 #862)
-#872 := (not #870)
-#875 := (and #605 #872)
-#878 := (not #875)
-#886 := (or #878 #883)
-#889 := (forall (vars (?v0 Int)) #886)
-#892 := (not #889)
-#910 := (or #524 #640 #892 #491 #867)
-#915 := (and #28 #910)
-#606 := (not #608)
-#610 := (and #605 #606)
-#613 := (not #610)
-#622 := (or #613 #618)
-#625 := (forall (vars (?v0 Int)) #622)
-#628 := (not #625)
-#918 := (or #628 #915)
-#921 := (and #625 #918)
-#557 := (not #13)
-#924 := (<= f3 0::Int)
-#944 := (or #924 #557 #921)
-#949 := (not #944)
-#1 := true
-#92 := (and #91 true)
-#87 := (<= #24 f15)
-#85 := (< #20 f16)
-#21 := (<= 0::Int #20)
-#86 := (and #21 #85)
-#88 := (implies #86 #87)
-#89 := (forall (vars (?v0 Int)) #88)
-#93 := (implies #89 #92)
-#94 := (and #89 #93)
-#83 := (<= 2::Int f16)
-#77 := (<= 0::Int f14)
-#84 := (and #77 #83)
-#95 := (implies #84 #94)
-#80 := (+ f8 1::Int)
-#81 := (= f16 #80)
-#96 := (implies #81 #95)
-#32 := (<= 1::Int f8)
-#78 := (and #77 #32)
-#97 := (implies #78 #96)
-#98 := (implies true #97)
-#109 := (implies #108 #98)
-#110 := (implies #107 #109)
-#30 := (<= 0::Int f7)
-#33 := (and #30 #32)
-#111 := (implies #33 #110)
-#106 := (<= #68 f9)
-#112 := (implies #106 #111)
-#113 := (implies #33 #112)
-#114 := (implies true #113)
-#99 := (implies #76 #98)
-#100 := (implies #74 #99)
-#72 := (and #32 #32)
-#101 := (implies #72 #100)
-#102 := (implies #71 #101)
-#69 := (< f9 #68)
-#103 := (implies #69 #102)
-#104 := (implies #33 #103)
-#105 := (implies true #104)
-#115 := (and #105 #114)
-#116 := (implies #33 #115)
-#67 := (< f8 f3)
-#117 := (implies #67 #116)
-#118 := (implies #33 #117)
-#119 := (implies true #118)
-#54 := (<= #24 f11)
-#49 := (< #20 f3)
-#50 := (and #21 #49)
-#55 := (implies #50 #54)
-#56 := (forall (vars (?v0 Int)) #55)
-#57 := (and #56 true)
-#52 := (implies #50 #51)
-#53 := (exists (vars (?v0 Int)) #52)
-#58 := (implies #53 #57)
-#59 := (and #53 #58)
-#60 := (implies #48 #59)
-#61 := (implies #46 #60)
-#62 := (implies #44 #61)
-#63 := (implies #33 #62)
-#42 := (<= f3 f8)
-#64 := (implies #42 #63)
-#65 := (implies #33 #64)
-#66 := (implies true #65)
-#120 := (and #66 #119)
-#121 := (implies #33 #120)
-#122 := (implies #41 #121)
-#37 := (<= #24 f9)
-#34 := (< #20 f8)
-#35 := (and #21 #34)
-#38 := (implies #35 #37)
-#39 := (forall (vars (?v0 Int)) #38)
-#123 := (implies #39 #122)
-#124 := (implies #33 #123)
-#125 := (implies true #124)
-#126 := (implies #28 #125)
-#127 := (and #28 #126)
-#25 := (<= #24 f4)
-#22 := (< #20 1::Int)
-#23 := (and #21 #22)
-#26 := (implies #23 #25)
-#27 := (forall (vars (?v0 Int)) #26)
-#128 := (implies #27 #127)
-#129 := (and #27 #128)
-#16 := (<= 1::Int 1::Int)
-#17 := (and #16 #16)
-#14 := (<= 0::Int 0::Int)
-#18 := (and #14 #17)
-#19 := (and #14 #18)
-#130 := (implies #19 #129)
-#131 := (implies #13 #130)
-#9 := (< 0::Int f3)
-#132 := (implies #9 #131)
-#133 := (implies true #132)
-#134 := (not #133)
-#952 := (iff #134 #949)
-#277 := (not #86)
-#278 := (or #277 #87)
-#281 := (forall (vars (?v0 Int)) #278)
-#289 := (not #281)
-#290 := (or #289 #91)
-#295 := (and #281 #290)
-#301 := (not #84)
-#302 := (or #301 #295)
-#271 := (+ 1::Int f8)
-#274 := (= f16 #271)
-#310 := (not #274)
-#311 := (or #310 #302)
-#319 := (not #78)
-#320 := (or #319 #311)
-#396 := (or #395 #320)
-#405 := (or #404 #396)
-#239 := (not #33)
-#413 := (or #239 #405)
-#421 := (not #106)
-#422 := (or #421 #413)
-#430 := (or #239 #422)
-#336 := (or #335 #320)
-#345 := (or #344 #336)
-#353 := (not #32)
-#354 := (or #353 #345)
-#363 := (or #362 #354)
-#371 := (not #69)
-#372 := (or #371 #363)
-#380 := (or #239 #372)
-#442 := (and #380 #430)
-#448 := (or #239 #442)
-#456 := (not #67)
-#457 := (or #456 #448)
-#465 := (or #239 #457)
-#177 := (not #50)
-#184 := (or #177 #54)
-#187 := (forall (vars (?v0 Int)) #184)
-#178 := (or #177 #51)
-#181 := (exists (vars (?v0 Int)) #178)
-#200 := (not #181)
-#201 := (or #200 #187)
-#206 := (and #181 #201)
-#213 := (or #212 #206)
-#222 := (or #221 #213)
-#231 := (or #230 #222)
-#240 := (or #239 #231)
-#248 := (not #42)
-#249 := (or #248 #240)
-#257 := (or #239 #249)
-#477 := (and #257 #465)
-#483 := (or #239 #477)
-#492 := (or #491 #483)
-#170 := (not #35)
-#171 := (or #170 #37)
-#174 := (forall (vars (?v0 Int)) #171)
-#500 := (not #174)
-#501 := (or #500 #492)
-#509 := (or #239 #501)
-#525 := (or #524 #509)
-#530 := (and #28 #525)
-#163 := (not #23)
-#164 := (or #163 #25)
-#167 := (forall (vars (?v0 Int)) #164)
-#536 := (not #167)
-#537 := (or #536 #530)
-#542 := (and #167 #537)
-#157 := (and #14 #16)
-#160 := (and #14 #157)
-#548 := (not #160)
-#549 := (or #548 #542)
-#558 := (or #557 #549)
-#566 := (not #9)
-#567 := (or #566 #558)
-#579 := (not #567)
-#950 := (iff #579 #949)
-#947 := (iff #567 #944)
-#935 := (or false #921)
-#938 := (or #557 #935)
-#941 := (or #924 #938)
-#945 := (iff #941 #944)
-#946 := [rewrite]: #945
-#942 := (iff #567 #941)
-#939 := (iff #558 #938)
-#936 := (iff #549 #935)
-#922 := (iff #542 #921)
-#919 := (iff #537 #918)
-#916 := (iff #530 #915)
-#913 := (iff #525 #910)
-#895 := (or #640 #867)
-#898 := (or #491 #895)
-#901 := (or #892 #898)
-#904 := (or #640 #901)
-#907 := (or #524 #904)
-#911 := (iff #907 #910)
-#912 := [rewrite]: #911
-#908 := (iff #525 #907)
-#905 := (iff #509 #904)
-#902 := (iff #501 #901)
-#899 := (iff #492 #898)
-#896 := (iff #483 #895)
-#868 := (iff #477 #867)
-#865 := (iff #465 #862)
-#853 := (or #640 #841)
-#856 := (or #683 #853)
-#859 := (or #640 #856)
-#863 := (iff #859 #862)
-#864 := [rewrite]: #863
-#860 := (iff #465 #859)
-#857 := (iff #457 #856)
-#854 := (iff #448 #853)
-#842 := (iff #442 #841)
-#839 := (iff #430 #836)
-#785 := (or #721 #753)
-#788 := (or #760 #785)
-#791 := (or #766 #788)
-#821 := (or #395 #791)
-#824 := (or #404 #821)
-#827 := (or #640 #824)
-#830 := (or #772 #827)
-#833 := (or #640 #830)
-#837 := (iff #833 #836)
-#838 := [rewrite]: #837
-#834 := (iff #430 #833)
-#831 := (iff #422 #830)
-#828 := (iff #413 #827)
-#825 := (iff #405 #824)
-#822 := (iff #396 #821)
-#792 := (iff #320 #791)
-#789 := (iff #311 #788)
-#786 := (iff #302 #785)
-#754 := (iff #295 #753)
-#751 := (iff #290 #750)
-#748 := (iff #289 #747)
-#745 := (iff #281 #744)
-#742 := (iff #278 #741)
-#739 := (iff #87 #738)
-#740 := [rewrite]: #739
-#734 := (iff #277 #733)
-#731 := (iff #86 #730)
-#728 := (iff #85 #727)
-#729 := [rewrite]: #728
-#603 := (iff #21 #605)
-#604 := [rewrite]: #603
-#732 := [monotonicity #604 #729]: #731
-#735 := [monotonicity #732]: #734
-#743 := [monotonicity #735 #740]: #742
-#746 := [quant-intro #743]: #745
-#749 := [monotonicity #746]: #748
-#752 := [monotonicity #749]: #751
-#755 := [monotonicity #746 #752]: #754
-#722 := (iff #301 #721)
-#719 := (iff #84 #718)
-#715 := (iff #83 #716)
-#717 := [rewrite]: #715
-#712 := (iff #77 #713)
-#714 := [rewrite]: #712
-#720 := [monotonicity #714 #717]: #719
-#723 := [monotonicity #720]: #722
-#787 := [monotonicity #723 #755]: #786
-#761 := (iff #310 #760)
-#758 := (iff #274 #756)
-#759 := [rewrite]: #758
-#762 := [monotonicity #759]: #761
-#790 := [monotonicity #762 #787]: #789
-#767 := (iff #319 #766)
-#764 := (iff #78 #763)
-#634 := (iff #32 #635)
-#636 := [rewrite]: #634
-#765 := [monotonicity #714 #636]: #764
-#768 := [monotonicity #765]: #767
-#793 := [monotonicity #768 #790]: #792
-#823 := [monotonicity #793]: #822
-#826 := [monotonicity #823]: #825
-#641 := (iff #239 #640)
-#638 := (iff #33 #637)
-#631 := (iff #30 #632)
-#633 := [rewrite]: #631
-#639 := [monotonicity #633 #636]: #638
-#642 := [monotonicity #639]: #641
-#829 := [monotonicity #642 #826]: #828
-#819 := (iff #421 #772)
-#817 := (iff #106 #773)
-#818 := [rewrite]: #817
-#820 := [monotonicity #818]: #819
-#832 := [monotonicity #820 #829]: #831
-#835 := [monotonicity #642 #832]: #834
-#840 := [trans #835 #838]: #839
-#815 := (iff #380 #812)
-#794 := (or #335 #791)
-#797 := (or #344 #794)
-#800 := (or #769 #797)
-#803 := (or #362 #800)
-#806 := (or #773 #803)
-#809 := (or #640 #806)
-#813 := (iff #809 #812)
-#814 := [rewrite]: #813
-#810 := (iff #380 #809)
-#807 := (iff #372 #806)
-#804 := (iff #363 #803)
-#801 := (iff #354 #800)
-#798 := (iff #345 #797)
-#795 := (iff #336 #794)
-#796 := [monotonicity #793]: #795
-#799 := [monotonicity #796]: #798
-#770 := (iff #353 #769)
-#771 := [monotonicity #636]: #770
-#802 := [monotonicity #771 #799]: #801
-#805 := [monotonicity #802]: #804
-#783 := (iff #371 #773)
-#778 := (not #772)
-#781 := (iff #778 #773)
-#782 := [rewrite]: #781
-#779 := (iff #371 #778)
-#776 := (iff #69 #772)
-#777 := [rewrite]: #776
-#780 := [monotonicity #777]: #779
-#784 := [trans #780 #782]: #783
-#808 := [monotonicity #784 #805]: #807
-#811 := [monotonicity #642 #808]: #810
-#816 := [trans #811 #814]: #815
-#843 := [monotonicity #816 #840]: #842
-#855 := [monotonicity #642 #843]: #854
-#851 := (iff #456 #683)
-#846 := (not #686)
-#849 := (iff #846 #683)
-#850 := [rewrite]: #849
-#847 := (iff #456 #846)
-#844 := (iff #67 #686)
-#845 := [rewrite]: #844
-#848 := [monotonicity #845]: #847
-#852 := [trans #848 #850]: #851
-#858 := [monotonicity #852 #855]: #857
-#861 := [monotonicity #642 #858]: #860
-#866 := [trans #861 #864]: #865
-#710 := (iff #257 #707)
-#689 := (or #212 #678)
-#692 := (or #221 #689)
-#695 := (or #230 #692)
-#698 := (or #640 #695)
-#701 := (or #686 #698)
-#704 := (or #640 #701)
-#708 := (iff #704 #707)
-#709 := [rewrite]: #708
-#705 := (iff #257 #704)
-#702 := (iff #249 #701)
-#699 := (iff #240 #698)
-#696 := (iff #231 #695)
-#693 := (iff #222 #692)
-#690 := (iff #213 #689)
-#679 := (iff #206 #678)
-#676 := (iff #201 #675)
-#673 := (iff #187 #672)
-#670 := (iff #184 #669)
-#667 := (iff #54 #666)
-#668 := [rewrite]: #667
-#653 := (iff #177 #652)
-#650 := (iff #50 #649)
-#647 := (iff #49 #646)
-#648 := [rewrite]: #647
-#651 := [monotonicity #604 #648]: #650
-#654 := [monotonicity #651]: #653
-#671 := [monotonicity #654 #668]: #670
-#674 := [quant-intro #671]: #673
-#662 := (iff #200 #661)
-#659 := (iff #181 #658)
-#656 := (iff #178 #655)
-#657 := [monotonicity #654]: #656
-#660 := [quant-intro #657]: #659
-#663 := [monotonicity #660]: #662
-#677 := [monotonicity #663 #674]: #676
-#680 := [monotonicity #660 #677]: #679
-#691 := [monotonicity #680]: #690
-#694 := [monotonicity #691]: #693
-#697 := [monotonicity #694]: #696
-#700 := [monotonicity #642 #697]: #699
-#687 := (iff #248 #686)
-#684 := (iff #42 #683)
-#685 := [rewrite]: #684
-#688 := [monotonicity #685]: #687
-#703 := [monotonicity #688 #700]: #702
-#706 := [monotonicity #642 #703]: #705
-#711 := [trans #706 #709]: #710
-#869 := [monotonicity #711 #866]: #868
-#897 := [monotonicity #642 #869]: #896
-#900 := [monotonicity #897]: #899
-#893 := (iff #500 #892)
-#890 := (iff #174 #889)
-#887 := (iff #171 #886)
-#884 := (iff #37 #883)
-#885 := [rewrite]: #884
-#879 := (iff #170 #878)
-#876 := (iff #35 #875)
-#873 := (iff #34 #872)
-#874 := [rewrite]: #873
-#877 := [monotonicity #604 #874]: #876
-#880 := [monotonicity #877]: #879
-#888 := [monotonicity #880 #885]: #887
-#891 := [quant-intro #888]: #890
-#894 := [monotonicity #891]: #893
-#903 := [monotonicity #894 #900]: #902
-#906 := [monotonicity #642 #903]: #905
-#909 := [monotonicity #906]: #908
-#914 := [trans #909 #912]: #913
-#917 := [monotonicity #914]: #916
-#629 := (iff #536 #628)
-#626 := (iff #167 #625)
-#623 := (iff #164 #622)
-#617 := (iff #25 #618)
-#621 := [rewrite]: #617
-#614 := (iff #163 #613)
-#611 := (iff #23 #610)
-#607 := (iff #22 #606)
-#609 := [rewrite]: #607
-#612 := [monotonicity #604 #609]: #611
-#615 := [monotonicity #612]: #614
-#624 := [monotonicity #615 #621]: #623
-#627 := [quant-intro #624]: #626
-#630 := [monotonicity #627]: #629
-#920 := [monotonicity #630 #917]: #919
-#923 := [monotonicity #627 #920]: #922
-#601 := (iff #548 false)
-#596 := (not true)
-#599 := (iff #596 false)
-#600 := [rewrite]: #599
-#597 := (iff #548 #596)
-#594 := (iff #160 true)
-#586 := (and true true)
-#589 := (and true #586)
-#592 := (iff #589 true)
-#593 := [rewrite]: #592
-#590 := (iff #160 #589)
-#587 := (iff #157 #586)
-#584 := (iff #16 true)
-#585 := [rewrite]: #584
-#582 := (iff #14 true)
-#583 := [rewrite]: #582
-#588 := [monotonicity #583 #585]: #587
-#591 := [monotonicity #583 #588]: #590
-#595 := [trans #591 #593]: #594
-#598 := [monotonicity #595]: #597
-#602 := [trans #598 #600]: #601
-#937 := [monotonicity #602 #923]: #936
-#940 := [monotonicity #937]: #939
-#933 := (iff #566 #924)
-#925 := (not #924)
-#928 := (not #925)
-#931 := (iff #928 #924)
-#932 := [rewrite]: #931
-#929 := (iff #566 #928)
-#926 := (iff #9 #925)
-#927 := [rewrite]: #926
-#930 := [monotonicity #927]: #929
-#934 := [trans #930 #932]: #933
-#943 := [monotonicity #934 #940]: #942
-#948 := [trans #943 #946]: #947
-#951 := [monotonicity #948]: #950
-#580 := (iff #134 #579)
-#577 := (iff #133 #567)
-#572 := (implies true #567)
-#575 := (iff #572 #567)
-#576 := [rewrite]: #575
-#573 := (iff #133 #572)
-#570 := (iff #132 #567)
-#563 := (implies #9 #558)
-#568 := (iff #563 #567)
-#569 := [rewrite]: #568
-#564 := (iff #132 #563)
-#561 := (iff #131 #558)
-#554 := (implies #13 #549)
-#559 := (iff #554 #558)
-#560 := [rewrite]: #559
-#555 := (iff #131 #554)
-#552 := (iff #130 #549)
-#545 := (implies #160 #542)
-#550 := (iff #545 #549)
-#551 := [rewrite]: #550
-#546 := (iff #130 #545)
-#543 := (iff #129 #542)
-#540 := (iff #128 #537)
-#533 := (implies #167 #530)
-#538 := (iff #533 #537)
-#539 := [rewrite]: #538
-#534 := (iff #128 #533)
-#531 := (iff #127 #530)
-#528 := (iff #126 #525)
-#521 := (implies #28 #509)
-#526 := (iff #521 #525)
-#527 := [rewrite]: #526
-#522 := (iff #126 #521)
-#519 := (iff #125 #509)
-#514 := (implies true #509)
-#517 := (iff #514 #509)
-#518 := [rewrite]: #517
-#515 := (iff #125 #514)
-#512 := (iff #124 #509)
-#506 := (implies #33 #501)
-#510 := (iff #506 #509)
-#511 := [rewrite]: #510
-#507 := (iff #124 #506)
-#504 := (iff #123 #501)
-#497 := (implies #174 #492)
-#502 := (iff #497 #501)
-#503 := [rewrite]: #502
-#498 := (iff #123 #497)
-#495 := (iff #122 #492)
-#488 := (implies #41 #483)
-#493 := (iff #488 #492)
-#494 := [rewrite]: #493
-#489 := (iff #122 #488)
-#486 := (iff #121 #483)
-#480 := (implies #33 #477)
-#484 := (iff #480 #483)
-#485 := [rewrite]: #484
-#481 := (iff #121 #480)
-#478 := (iff #120 #477)
-#475 := (iff #119 #465)
-#470 := (implies true #465)
-#473 := (iff #470 #465)
-#474 := [rewrite]: #473
-#471 := (iff #119 #470)
-#468 := (iff #118 #465)
-#462 := (implies #33 #457)
-#466 := (iff #462 #465)
-#467 := [rewrite]: #466
-#463 := (iff #118 #462)
-#460 := (iff #117 #457)
-#453 := (implies #67 #448)
-#458 := (iff #453 #457)
-#459 := [rewrite]: #458
-#454 := (iff #117 #453)
-#451 := (iff #116 #448)
-#445 := (implies #33 #442)
-#449 := (iff #445 #448)
-#450 := [rewrite]: #449
-#446 := (iff #116 #445)
-#443 := (iff #115 #442)
-#440 := (iff #114 #430)
-#435 := (implies true #430)
-#438 := (iff #435 #430)
-#439 := [rewrite]: #438
-#436 := (iff #114 #435)
-#433 := (iff #113 #430)
-#427 := (implies #33 #422)
-#431 := (iff #427 #430)
-#432 := [rewrite]: #431
-#428 := (iff #113 #427)
-#425 := (iff #112 #422)
-#418 := (implies #106 #413)
-#423 := (iff #418 #422)
-#424 := [rewrite]: #423
-#419 := (iff #112 #418)
-#416 := (iff #111 #413)
-#410 := (implies #33 #405)
-#414 := (iff #410 #413)
-#415 := [rewrite]: #414
-#411 := (iff #111 #410)
-#408 := (iff #110 #405)
-#401 := (implies #107 #396)
-#406 := (iff #401 #405)
-#407 := [rewrite]: #406
-#402 := (iff #110 #401)
-#399 := (iff #109 #396)
-#392 := (implies #108 #320)
-#397 := (iff #392 #396)
-#398 := [rewrite]: #397
-#393 := (iff #109 #392)
-#330 := (iff #98 #320)
-#325 := (implies true #320)
-#328 := (iff #325 #320)
-#329 := [rewrite]: #328
-#326 := (iff #98 #325)
-#323 := (iff #97 #320)
-#316 := (implies #78 #311)
-#321 := (iff #316 #320)
-#322 := [rewrite]: #321
-#317 := (iff #97 #316)
-#314 := (iff #96 #311)
-#307 := (implies #274 #302)
-#312 := (iff #307 #311)
-#313 := [rewrite]: #312
-#308 := (iff #96 #307)
-#305 := (iff #95 #302)
-#298 := (implies #84 #295)
-#303 := (iff #298 #302)
-#304 := [rewrite]: #303
-#299 := (iff #95 #298)
-#296 := (iff #94 #295)
-#293 := (iff #93 #290)
-#286 := (implies #281 #91)
-#291 := (iff #286 #290)
-#292 := [rewrite]: #291
-#287 := (iff #93 #286)
-#284 := (iff #92 #91)
-#285 := [rewrite]: #284
-#282 := (iff #89 #281)
-#279 := (iff #88 #278)
-#280 := [rewrite]: #279
-#283 := [quant-intro #280]: #282
-#288 := [monotonicity #283 #285]: #287
-#294 := [trans #288 #292]: #293
-#297 := [monotonicity #283 #294]: #296
-#300 := [monotonicity #297]: #299
-#306 := [trans #300 #304]: #305
-#275 := (iff #81 #274)
-#272 := (= #80 #271)
-#273 := [rewrite]: #272
-#276 := [monotonicity #273]: #275
-#309 := [monotonicity #276 #306]: #308
-#315 := [trans #309 #313]: #314
-#318 := [monotonicity #315]: #317
-#324 := [trans #318 #322]: #323
-#327 := [monotonicity #324]: #326
-#331 := [trans #327 #329]: #330
-#394 := [monotonicity #331]: #393
-#400 := [trans #394 #398]: #399
-#403 := [monotonicity #400]: #402
-#409 := [trans #403 #407]: #408
-#412 := [monotonicity #409]: #411
-#417 := [trans #412 #415]: #416
-#420 := [monotonicity #417]: #419
-#426 := [trans #420 #424]: #425
-#429 := [monotonicity #426]: #428
-#434 := [trans #429 #432]: #433
-#437 := [monotonicity #434]: #436
-#441 := [trans #437 #439]: #440
-#390 := (iff #105 #380)
-#385 := (implies true #380)
-#388 := (iff #385 #380)
-#389 := [rewrite]: #388
-#386 := (iff #105 #385)
-#383 := (iff #104 #380)
-#377 := (implies #33 #372)
-#381 := (iff #377 #380)
-#382 := [rewrite]: #381
-#378 := (iff #104 #377)
-#375 := (iff #103 #372)
-#368 := (implies #69 #363)
-#373 := (iff #368 #372)
-#374 := [rewrite]: #373
-#369 := (iff #103 #368)
-#366 := (iff #102 #363)
-#359 := (implies #71 #354)
-#364 := (iff #359 #363)
-#365 := [rewrite]: #364
-#360 := (iff #102 #359)
-#357 := (iff #101 #354)
-#350 := (implies #32 #345)
-#355 := (iff #350 #354)
-#356 := [rewrite]: #355
-#351 := (iff #101 #350)
-#348 := (iff #100 #345)
-#341 := (implies #74 #336)
-#346 := (iff #341 #345)
-#347 := [rewrite]: #346
-#342 := (iff #100 #341)
-#339 := (iff #99 #336)
-#332 := (implies #76 #320)
-#337 := (iff #332 #336)
-#338 := [rewrite]: #337
-#333 := (iff #99 #332)
-#334 := [monotonicity #331]: #333
-#340 := [trans #334 #338]: #339
-#343 := [monotonicity #340]: #342
-#349 := [trans #343 #347]: #348
-#269 := (iff #72 #32)
-#270 := [rewrite]: #269
-#352 := [monotonicity #270 #349]: #351
-#358 := [trans #352 #356]: #357
-#361 := [monotonicity #358]: #360
-#367 := [trans #361 #365]: #366
-#370 := [monotonicity #367]: #369
-#376 := [trans #370 #374]: #375
-#379 := [monotonicity #376]: #378
-#384 := [trans #379 #382]: #383
-#387 := [monotonicity #384]: #386
-#391 := [trans #387 #389]: #390
-#444 := [monotonicity #391 #441]: #443
-#447 := [monotonicity #444]: #446
-#452 := [trans #447 #450]: #451
-#455 := [monotonicity #452]: #454
-#461 := [trans #455 #459]: #460
-#464 := [monotonicity #461]: #463
-#469 := [trans #464 #467]: #468
-#472 := [monotonicity #469]: #471
-#476 := [trans #472 #474]: #475
-#267 := (iff #66 #257)
-#262 := (implies true #257)
-#265 := (iff #262 #257)
-#266 := [rewrite]: #265
-#263 := (iff #66 #262)
-#260 := (iff #65 #257)
-#254 := (implies #33 #249)
-#258 := (iff #254 #257)
-#259 := [rewrite]: #258
-#255 := (iff #65 #254)
-#252 := (iff #64 #249)
-#245 := (implies #42 #240)
-#250 := (iff #245 #249)
-#251 := [rewrite]: #250
-#246 := (iff #64 #245)
-#243 := (iff #63 #240)
-#236 := (implies #33 #231)
-#241 := (iff #236 #240)
-#242 := [rewrite]: #241
-#237 := (iff #63 #236)
-#234 := (iff #62 #231)
-#227 := (implies #44 #222)
-#232 := (iff #227 #231)
-#233 := [rewrite]: #232
-#228 := (iff #62 #227)
-#225 := (iff #61 #222)
-#218 := (implies #46 #213)
-#223 := (iff #218 #222)
-#224 := [rewrite]: #223
-#219 := (iff #61 #218)
-#216 := (iff #60 #213)
-#209 := (implies #48 #206)
-#214 := (iff #209 #213)
-#215 := [rewrite]: #214
-#210 := (iff #60 #209)
-#207 := (iff #59 #206)
-#204 := (iff #58 #201)
-#197 := (implies #181 #187)
-#202 := (iff #197 #201)
-#203 := [rewrite]: #202
-#198 := (iff #58 #197)
-#195 := (iff #57 #187)
-#190 := (and #187 true)
-#193 := (iff #190 #187)
-#194 := [rewrite]: #193
-#191 := (iff #57 #190)
-#188 := (iff #56 #187)
-#185 := (iff #55 #184)
-#186 := [rewrite]: #185
-#189 := [quant-intro #186]: #188
-#192 := [monotonicity #189]: #191
-#196 := [trans #192 #194]: #195
-#182 := (iff #53 #181)
-#179 := (iff #52 #178)
-#180 := [rewrite]: #179
-#183 := [quant-intro #180]: #182
-#199 := [monotonicity #183 #196]: #198
-#205 := [trans #199 #203]: #204
-#208 := [monotonicity #183 #205]: #207
-#211 := [monotonicity #208]: #210
-#217 := [trans #211 #215]: #216
-#220 := [monotonicity #217]: #219
-#226 := [trans #220 #224]: #225
-#229 := [monotonicity #226]: #228
-#235 := [trans #229 #233]: #234
-#238 := [monotonicity #235]: #237
-#244 := [trans #238 #242]: #243
-#247 := [monotonicity #244]: #246
-#253 := [trans #247 #251]: #252
-#256 := [monotonicity #253]: #255
-#261 := [trans #256 #259]: #260
-#264 := [monotonicity #261]: #263
-#268 := [trans #264 #266]: #267
-#479 := [monotonicity #268 #476]: #478
-#482 := [monotonicity #479]: #481
-#487 := [trans #482 #485]: #486
-#490 := [monotonicity #487]: #489
-#496 := [trans #490 #494]: #495
-#175 := (iff #39 #174)
-#172 := (iff #38 #171)
-#173 := [rewrite]: #172
-#176 := [quant-intro #173]: #175
-#499 := [monotonicity #176 #496]: #498
-#505 := [trans #499 #503]: #504
-#508 := [monotonicity #505]: #507
-#513 := [trans #508 #511]: #512
-#516 := [monotonicity #513]: #515
-#520 := [trans #516 #518]: #519
-#523 := [monotonicity #520]: #522
-#529 := [trans #523 #527]: #528
-#532 := [monotonicity #529]: #531
-#168 := (iff #27 #167)
-#165 := (iff #26 #164)
-#166 := [rewrite]: #165
-#169 := [quant-intro #166]: #168
-#535 := [monotonicity #169 #532]: #534
-#541 := [trans #535 #539]: #540
-#544 := [monotonicity #169 #541]: #543
-#161 := (iff #19 #160)
-#158 := (iff #18 #157)
-#155 := (iff #17 #16)
-#156 := [rewrite]: #155
-#159 := [monotonicity #156]: #158
-#162 := [monotonicity #159]: #161
-#547 := [monotonicity #162 #544]: #546
-#553 := [trans #547 #551]: #552
-#556 := [monotonicity #553]: #555
-#562 := [trans #556 #560]: #561
-#565 := [monotonicity #562]: #564
-#571 := [trans #565 #569]: #570
-#574 := [monotonicity #571]: #573
-#578 := [trans #574 #576]: #577
-#581 := [monotonicity #578]: #580
-#953 := [trans #581 #951]: #952
-#154 := [asserted]: #134
-#954 := [mp #154 #953]: #949
-#956 := [not-or-elim #954]: #13
-#1861 := [trans #956 #1854]: #1927
-#1749 := (not #961)
-#1740 := (or #1333 #1749)
-#1751 := [def-axiom]: #1740
-#1863 := [unit-resolution #1751 #1926]: #1749
-#1864 := (not #1927)
-#1892 := (or #1864 #961)
-#1865 := [th-lemma arith triangle-eq]: #1892
-#1867 := [unit-resolution #1865 #1863 #1861]: false
-#1868 := [lemma #1867]: #1333
-#2178 := (or #1338 #2175)
-#1520 := (forall (vars (?v0 Int)) #1515)
-#1526 := (not #1520)
-#1527 := (or #1526 #91)
-#1528 := (not #1527)
-#1533 := (or #1498 #1528)
-#1541 := (not #1533)
-#1551 := (or #1470 #772 #404 #395 #769 #760 #1539 #1540 #1541)
-#1552 := (not #1551)
-#1542 := (or #1470 #773 #362 #344 #335 #769 #760 #1539 #1540 #1541)
-#1543 := (not #1542)
-#1557 := (or #1543 #1552)
-#1563 := (not #1557)
-#1564 := (or #1470 #769 #683 #1563)
-#1565 := (not #1564)
-#1408 := (forall (vars (?v0 Int)) #1405)
-#1464 := (or #1408 #1459)
-#1471 := (not #1464)
-#1472 := (or #1470 #769 #686 #230 #221 #212 #1471)
-#1473 := (not #1472)
-#1570 := (or #1473 #1565)
-#1577 := (not #1570)
-#1383 := (forall (vars (?v0 Int)) #1378)
-#1576 := (not #1383)
-#1578 := (or #524 #1470 #769 #1576 #491 #1577)
-#1579 := (not #1578)
-#1584 := (or #524 #1579)
-#1591 := (not #1584)
-#1361 := (forall (vars (?v0 Int)) #1356)
-#1590 := (not #1361)
-#1592 := (or #1590 #1591)
-#1593 := (not #1592)
-#1598 := (or #1338 #1593)
-#2179 := (iff #1598 #2178)
-#2176 := (iff #1593 #2175)
-#2173 := (iff #1592 #2172)
-#2170 := (iff #1591 #2169)
-#2167 := (iff #1584 #2166)
-#2164 := (iff #1579 #2163)
-#2161 := (iff #1578 #2160)
-#2158 := (iff #1577 #2157)
-#2155 := (iff #1570 #2154)
-#2152 := (iff #1565 #2151)
-#2149 := (iff #1564 #2148)
-#2146 := (iff #1563 #2145)
-#2143 := (iff #1557 #2142)
-#2140 := (iff #1552 #2139)
-#2137 := (iff #1551 #2136)
-#2128 := (iff #1541 #2127)
-#2125 := (iff #1533 #2124)
-#2122 := (iff #1528 #2121)
-#2119 := (iff #1527 #2118)
-#2116 := (iff #1526 #2115)
-#2113 := (iff #1520 #2110)
-#2111 := (iff #1515 #1515)
-#2112 := [refl]: #2111
-#2114 := [quant-intro #2112]: #2113
-#2117 := [monotonicity #2114]: #2116
-#2120 := [monotonicity #2117]: #2119
-#2123 := [monotonicity #2120]: #2122
-#2126 := [monotonicity #2123]: #2125
-#2129 := [monotonicity #2126]: #2128
-#2138 := [monotonicity #2129]: #2137
-#2141 := [monotonicity #2138]: #2140
-#2134 := (iff #1543 #2133)
-#2131 := (iff #1542 #2130)
-#2132 := [monotonicity #2129]: #2131
-#2135 := [monotonicity #2132]: #2134
-#2144 := [monotonicity #2135 #2141]: #2143
-#2147 := [monotonicity #2144]: #2146
-#2150 := [monotonicity #2147]: #2149
-#2153 := [monotonicity #2150]: #2152
-#2108 := (iff #1473 #2107)
-#2105 := (iff #1472 #2104)
-#2102 := (iff #1471 #2101)
-#2099 := (iff #1464 #2098)
-#2096 := (iff #1408 #2093)
-#2094 := (iff #1405 #1405)
-#2095 := [refl]: #2094
-#2097 := [quant-intro #2095]: #2096
-#2100 := [monotonicity #2097]: #2099
-#2103 := [monotonicity #2100]: #2102
-#2106 := [monotonicity #2103]: #2105
-#2109 := [monotonicity #2106]: #2108
-#2156 := [monotonicity #2109 #2153]: #2155
-#2159 := [monotonicity #2156]: #2158
-#2091 := (iff #1576 #2090)
-#2088 := (iff #1383 #2085)
-#2086 := (iff #1378 #1378)
-#2087 := [refl]: #2086
-#2089 := [quant-intro #2087]: #2088
-#2092 := [monotonicity #2089]: #2091
-#2162 := [monotonicity #2092 #2159]: #2161
-#2165 := [monotonicity #2162]: #2164
-#2168 := [monotonicity #2165]: #2167
-#2171 := [monotonicity #2168]: #2170
-#2083 := (iff #1590 #2082)
-#2080 := (iff #1361 #2077)
-#2078 := (iff #1356 #1356)
-#2079 := [refl]: #2078
-#2081 := [quant-intro #2079]: #2080
-#2084 := [monotonicity #2081]: #2083
-#2174 := [monotonicity #2084 #2171]: #2173
-#2177 := [monotonicity #2174]: #2176
-#2180 := [monotonicity #2177]: #2179
-#1116 := (not #91)
-#1119 := (and #744 #1116)
-#1245 := (not #1240)
-#1248 := (and #1094 #1245)
-#1251 := (not #1248)
-#1267 := (or #1251 #1262)
-#1270 := (not #1267)
-#1273 := (or #1270 #1119)
-#1291 := (and #632 #773 #107 #108 #635 #756 #713 #716 #1273)
-#1279 := (and #632 #772 #71 #74 #76 #635 #756 #713 #716 #1273)
-#1296 := (or #1279 #1291)
-#1302 := (and #632 #635 #686 #1296)
-#1043 := (not #1042)
-#1044 := (and #1039 #1043)
-#1045 := (not #1044)
-#1198 := (or #1045 #1193)
-#1201 := (not #1198)
-#1025 := (not #1024)
-#1026 := (and #1021 #1025)
-#1027 := (not #1026)
-#1030 := (or #1027 #1029)
-#1204 := (and #1030 #1201)
-#1014 := (not #655)
-#1017 := (forall (vars (?v0 Int)) #1014)
-#1207 := (or #1017 #1204)
-#1213 := (and #632 #635 #683 #44 #46 #48 #1207)
-#1307 := (or #1213 #1302)
-#1313 := (and #28 #632 #635 #889 #41 #1307)
-#1318 := (or #524 #1313)
-#1321 := (and #625 #1318)
-#964 := (and #969 #971)
-#965 := (not #964)
-#972 := (or #965 #961)
-#973 := (not #972)
-#1324 := (or #973 #1321)
-#1599 := (iff #1324 #1598)
-#1596 := (iff #1321 #1593)
-#1587 := (and #1361 #1584)
-#1594 := (iff #1587 #1593)
-#1595 := [rewrite]: #1594
-#1588 := (iff #1321 #1587)
-#1585 := (iff #1318 #1584)
-#1582 := (iff #1313 #1579)
-#1573 := (and #28 #632 #635 #1383 #41 #1570)
-#1580 := (iff #1573 #1579)
-#1581 := [rewrite]: #1580
-#1574 := (iff #1313 #1573)
-#1571 := (iff #1307 #1570)
-#1568 := (iff #1302 #1565)
-#1560 := (and #632 #635 #686 #1557)
-#1566 := (iff #1560 #1565)
-#1567 := [rewrite]: #1566
-#1561 := (iff #1302 #1560)
-#1558 := (iff #1296 #1557)
-#1555 := (iff #1291 #1552)
-#1548 := (and #632 #773 #107 #108 #635 #756 #713 #716 #1533)
-#1553 := (iff #1548 #1552)
-#1554 := [rewrite]: #1553
-#1549 := (iff #1291 #1548)
-#1534 := (iff #1273 #1533)
-#1531 := (iff #1119 #1528)
-#1523 := (and #1520 #1116)
-#1529 := (iff #1523 #1528)
-#1530 := [rewrite]: #1529
-#1524 := (iff #1119 #1523)
-#1521 := (iff #744 #1520)
-#1518 := (iff #741 #1515)
-#1501 := (or #1341 #724)
-#1512 := (or #1501 #738)
-#1516 := (iff #1512 #1515)
-#1517 := [rewrite]: #1516
-#1513 := (iff #741 #1512)
-#1510 := (iff #733 #1501)
-#1502 := (not #1501)
-#1505 := (not #1502)
-#1508 := (iff #1505 #1501)
-#1509 := [rewrite]: #1508
-#1506 := (iff #733 #1505)
-#1503 := (iff #730 #1502)
-#1504 := [rewrite]: #1503
-#1507 := [monotonicity #1504]: #1506
-#1511 := [trans #1507 #1509]: #1510
-#1514 := [monotonicity #1511]: #1513
-#1519 := [trans #1514 #1517]: #1518
-#1522 := [quant-intro #1519]: #1521
-#1525 := [monotonicity #1522]: #1524
-#1532 := [trans #1525 #1530]: #1531
-#1499 := (iff #1270 #1498)
-#1496 := (iff #1267 #1493)
-#1479 := (or #1478 #1240)
-#1490 := (or #1479 #1262)
-#1494 := (iff #1490 #1493)
-#1495 := [rewrite]: #1494
-#1491 := (iff #1267 #1490)
-#1488 := (iff #1251 #1479)
-#1480 := (not #1479)
-#1483 := (not #1480)
-#1486 := (iff #1483 #1479)
-#1487 := [rewrite]: #1486
-#1484 := (iff #1251 #1483)
-#1481 := (iff #1248 #1480)
-#1482 := [rewrite]: #1481
-#1485 := [monotonicity #1482]: #1484
-#1489 := [trans #1485 #1487]: #1488
-#1492 := [monotonicity #1489]: #1491
-#1497 := [trans #1492 #1495]: #1496
-#1500 := [monotonicity #1497]: #1499
-#1535 := [monotonicity #1500 #1532]: #1534
-#1550 := [monotonicity #1535]: #1549
-#1556 := [trans #1550 #1554]: #1555
-#1546 := (iff #1279 #1543)
-#1536 := (and #632 #772 #71 #74 #76 #635 #756 #713 #716 #1533)
-#1544 := (iff #1536 #1543)
-#1545 := [rewrite]: #1544
-#1537 := (iff #1279 #1536)
-#1538 := [monotonicity #1535]: #1537
-#1547 := [trans #1538 #1545]: #1546
-#1559 := [monotonicity #1547 #1556]: #1558
-#1562 := [monotonicity #1559]: #1561
-#1569 := [trans #1562 #1567]: #1568
-#1476 := (iff #1213 #1473)
-#1467 := (and #632 #635 #683 #44 #46 #48 #1464)
-#1474 := (iff #1467 #1473)
-#1475 := [rewrite]: #1474
-#1468 := (iff #1213 #1467)
-#1465 := (iff #1207 #1464)
-#1462 := (iff #1204 #1459)
-#1446 := (or #1431 #1042 #1193)
-#1451 := (not #1446)
-#1454 := (and #1426 #1451)
-#1460 := (iff #1454 #1459)
-#1461 := [rewrite]: #1460
-#1455 := (iff #1204 #1454)
-#1452 := (iff #1201 #1451)
-#1449 := (iff #1198 #1446)
-#1432 := (or #1431 #1042)
-#1443 := (or #1432 #1193)
-#1447 := (iff #1443 #1446)
-#1448 := [rewrite]: #1447
-#1444 := (iff #1198 #1443)
-#1441 := (iff #1045 #1432)
-#1433 := (not #1432)
-#1436 := (not #1433)
-#1439 := (iff #1436 #1432)
-#1440 := [rewrite]: #1439
-#1437 := (iff #1045 #1436)
-#1434 := (iff #1044 #1433)
-#1435 := [rewrite]: #1434
-#1438 := [monotonicity #1435]: #1437
-#1442 := [trans #1438 #1440]: #1441
-#1445 := [monotonicity #1442]: #1444
-#1450 := [trans #1445 #1448]: #1449
-#1453 := [monotonicity #1450]: #1452
-#1429 := (iff #1030 #1426)
-#1412 := (or #1411 #1024)
-#1423 := (or #1412 #1029)
-#1427 := (iff #1423 #1426)
-#1428 := [rewrite]: #1427
-#1424 := (iff #1030 #1423)
-#1421 := (iff #1027 #1412)
-#1413 := (not #1412)
-#1416 := (not #1413)
-#1419 := (iff #1416 #1412)
-#1420 := [rewrite]: #1419
-#1417 := (iff #1027 #1416)
-#1414 := (iff #1026 #1413)
-#1415 := [rewrite]: #1414
-#1418 := [monotonicity #1415]: #1417
-#1422 := [trans #1418 #1420]: #1421
-#1425 := [monotonicity #1422]: #1424
-#1430 := [trans #1425 #1428]: #1429
-#1456 := [monotonicity #1430 #1453]: #1455
-#1463 := [trans #1456 #1461]: #1462
-#1409 := (iff #1017 #1408)
-#1406 := (iff #1014 #1405)
-#1403 := (iff #655 #1400)
-#1386 := (or #1341 #645)
-#1397 := (or #1386 #51)
-#1401 := (iff #1397 #1400)
-#1402 := [rewrite]: #1401
-#1398 := (iff #655 #1397)
-#1395 := (iff #652 #1386)
-#1387 := (not #1386)
-#1390 := (not #1387)
-#1393 := (iff #1390 #1386)
-#1394 := [rewrite]: #1393
-#1391 := (iff #652 #1390)
-#1388 := (iff #649 #1387)
-#1389 := [rewrite]: #1388
-#1392 := [monotonicity #1389]: #1391
-#1396 := [trans #1392 #1394]: #1395
-#1399 := [monotonicity #1396]: #1398
-#1404 := [trans #1399 #1402]: #1403
-#1407 := [monotonicity #1404]: #1406
-#1410 := [quant-intro #1407]: #1409
-#1466 := [monotonicity #1410 #1463]: #1465
-#1469 := [monotonicity #1466]: #1468
-#1477 := [trans #1469 #1475]: #1476
-#1572 := [monotonicity #1477 #1569]: #1571
-#1384 := (iff #889 #1383)
-#1381 := (iff #886 #1378)
-#1364 := (or #1341 #870)
-#1375 := (or #1364 #883)
-#1379 := (iff #1375 #1378)
-#1380 := [rewrite]: #1379
-#1376 := (iff #886 #1375)
-#1373 := (iff #878 #1364)
-#1365 := (not #1364)
-#1368 := (not #1365)
-#1371 := (iff #1368 #1364)
-#1372 := [rewrite]: #1371
-#1369 := (iff #878 #1368)
-#1366 := (iff #875 #1365)
-#1367 := [rewrite]: #1366
-#1370 := [monotonicity #1367]: #1369
-#1374 := [trans #1370 #1372]: #1373
-#1377 := [monotonicity #1374]: #1376
-#1382 := [trans #1377 #1380]: #1381
-#1385 := [quant-intro #1382]: #1384
-#1575 := [monotonicity #1385 #1572]: #1574
-#1583 := [trans #1575 #1581]: #1582
-#1586 := [monotonicity #1583]: #1585
-#1362 := (iff #625 #1361)
-#1359 := (iff #622 #1356)
-#1342 := (or #1341 #608)
-#1353 := (or #1342 #618)
-#1357 := (iff #1353 #1356)
-#1358 := [rewrite]: #1357
-#1354 := (iff #622 #1353)
-#1351 := (iff #613 #1342)
-#1343 := (not #1342)
-#1346 := (not #1343)
-#1349 := (iff #1346 #1342)
-#1350 := [rewrite]: #1349
-#1347 := (iff #613 #1346)
-#1344 := (iff #610 #1343)
-#1345 := [rewrite]: #1344
-#1348 := [monotonicity #1345]: #1347
-#1352 := [trans #1348 #1350]: #1351
-#1355 := [monotonicity #1352]: #1354
-#1360 := [trans #1355 #1358]: #1359
-#1363 := [quant-intro #1360]: #1362
-#1589 := [monotonicity #1363 #1586]: #1588
-#1597 := [trans #1589 #1595]: #1596
-#1339 := (iff #973 #1338)
-#1336 := (iff #972 #1333)
-#1034 := (or #1033 #970)
-#1330 := (or #1034 #961)
-#1334 := (iff #1330 #1333)
-#1335 := [rewrite]: #1334
-#1331 := (iff #972 #1330)
-#1328 := (iff #965 #1034)
-#1054 := (not #1034)
-#1108 := (not #1054)
-#1172 := (iff #1108 #1034)
-#1327 := [rewrite]: #1172
-#976 := (iff #965 #1108)
-#1055 := (iff #964 #1054)
-#1107 := [rewrite]: #1055
-#977 := [monotonicity #1107]: #976
-#1329 := [trans #977 #1327]: #1328
-#1332 := [monotonicity #1329]: #1331
-#1337 := [trans #1332 #1335]: #1336
-#1340 := [monotonicity #1337]: #1339
-#1600 := [monotonicity #1340 #1597]: #1599
-#1101 := (+ #1100 #736)
-#1102 := (<= #1101 0::Int)
-#1095 := (+ ?v0!3 #725)
-#1096 := (>= #1095 0::Int)
-#1097 := (not #1096)
-#1098 := (and #1094 #1097)
-#1099 := (not #1098)
-#1103 := (or #1099 #1102)
-#1104 := (not #1103)
-#1123 := (or #1104 #1119)
-#1090 := (not #721)
-#1087 := (not #760)
-#1084 := (not #766)
-#1136 := (not #395)
-#1133 := (not #404)
-#990 := (not #640)
-#1139 := (and #990 #778 #1133 #1136 #1084 #1087 #1090 #1123)
-#1081 := (not #335)
-#1078 := (not #344)
-#1075 := (not #769)
-#1072 := (not #362)
-#1127 := (and #990 #772 #1072 #1075 #1078 #1081 #1084 #1087 #1090 #1123)
-#1143 := (or #1127 #1139)
-#1147 := (and #990 #686 #1143)
-#1047 := (+ #1046 #664)
-#1048 := (<= #1047 0::Int)
-#1049 := (or #1045 #1048)
-#1050 := (not #1049)
-#1056 := (and #1030 #1050)
-#1060 := (or #1017 #1056)
-#1011 := (not #212)
-#1008 := (not #221)
-#1005 := (not #230)
-#1064 := (and #990 #846 #1005 #1008 #1011 #1060)
-#1151 := (or #1064 #1147)
-#1000 := (not #491)
-#987 := (not #524)
-#1155 := (and #987 #990 #889 #1000 #1151)
-#1159 := (or #524 #1155)
-#1163 := (and #625 #1159)
-#1167 := (or #973 #1163)
-#1325 := (iff #1167 #1324)
-#1322 := (iff #1163 #1321)
-#1319 := (iff #1159 #1318)
-#1316 := (iff #1155 #1313)
-#1310 := (and #28 #637 #889 #41 #1307)
-#1314 := (iff #1310 #1313)
-#1315 := [rewrite]: #1314
-#1311 := (iff #1155 #1310)
-#1308 := (iff #1151 #1307)
-#1305 := (iff #1147 #1302)
-#1299 := (and #637 #686 #1296)
-#1303 := (iff #1299 #1302)
-#1304 := [rewrite]: #1303
-#1300 := (iff #1147 #1299)
-#1297 := (iff #1143 #1296)
-#1294 := (iff #1139 #1291)
-#1288 := (and #637 #773 #107 #108 #763 #756 #718 #1273)
-#1292 := (iff #1288 #1291)
-#1293 := [rewrite]: #1292
-#1289 := (iff #1139 #1288)
-#1274 := (iff #1123 #1273)
-#1271 := (iff #1104 #1270)
-#1268 := (iff #1103 #1267)
-#1265 := (iff #1102 #1262)
-#1254 := (+ #736 #1100)
-#1257 := (<= #1254 0::Int)
-#1263 := (iff #1257 #1262)
-#1264 := [rewrite]: #1263
-#1258 := (iff #1102 #1257)
-#1255 := (= #1101 #1254)
-#1256 := [rewrite]: #1255
-#1259 := [monotonicity #1256]: #1258
-#1266 := [trans #1259 #1264]: #1265
-#1252 := (iff #1099 #1251)
-#1249 := (iff #1098 #1248)
-#1246 := (iff #1097 #1245)
-#1243 := (iff #1096 #1240)
-#1232 := (+ #725 ?v0!3)
-#1235 := (>= #1232 0::Int)
-#1241 := (iff #1235 #1240)
-#1242 := [rewrite]: #1241
-#1236 := (iff #1096 #1235)
-#1233 := (= #1095 #1232)
-#1234 := [rewrite]: #1233
-#1237 := [monotonicity #1234]: #1236
-#1244 := [trans #1237 #1242]: #1243
-#1247 := [monotonicity #1244]: #1246
-#1250 := [monotonicity #1247]: #1249
-#1253 := [monotonicity #1250]: #1252
-#1269 := [monotonicity #1253 #1266]: #1268
-#1272 := [monotonicity #1269]: #1271
-#1275 := [monotonicity #1272]: #1274
-#1230 := (iff #1090 #718)
-#1231 := [rewrite]: #1230
-#1228 := (iff #1087 #756)
-#1229 := [rewrite]: #1228
-#1226 := (iff #1084 #763)
-#1227 := [rewrite]: #1226
-#1286 := (iff #1136 #108)
-#1287 := [rewrite]: #1286
-#1284 := (iff #1133 #107)
-#1285 := [rewrite]: #1284
-#1175 := (iff #990 #637)
-#1176 := [rewrite]: #1175
-#1290 := [monotonicity #1176 #782 #1285 #1287 #1227 #1229 #1231 #1275]: #1289
-#1295 := [trans #1290 #1293]: #1294
-#1282 := (iff #1127 #1279)
-#1276 := (and #637 #772 #71 #635 #74 #76 #763 #756 #718 #1273)
-#1280 := (iff #1276 #1279)
-#1281 := [rewrite]: #1280
-#1277 := (iff #1127 #1276)
-#1224 := (iff #1081 #76)
-#1225 := [rewrite]: #1224
-#1222 := (iff #1078 #74)
-#1223 := [rewrite]: #1222
-#1220 := (iff #1075 #635)
-#1221 := [rewrite]: #1220
-#1218 := (iff #1072 #71)
-#1219 := [rewrite]: #1218
-#1278 := [monotonicity #1176 #1219 #1221 #1223 #1225 #1227 #1229 #1231 #1275]: #1277
-#1283 := [trans #1278 #1281]: #1282
-#1298 := [monotonicity #1283 #1295]: #1297
-#1301 := [monotonicity #1176 #1298]: #1300
-#1306 := [trans #1301 #1304]: #1305
-#1216 := (iff #1064 #1213)
-#1210 := (and #637 #683 #44 #46 #48 #1207)
-#1214 := (iff #1210 #1213)
-#1215 := [rewrite]: #1214
-#1211 := (iff #1064 #1210)
-#1208 := (iff #1060 #1207)
-#1205 := (iff #1056 #1204)
-#1202 := (iff #1050 #1201)
-#1199 := (iff #1049 #1198)
-#1196 := (iff #1048 #1193)
-#1185 := (+ #664 #1046)
-#1188 := (<= #1185 0::Int)
-#1194 := (iff #1188 #1193)
-#1195 := [rewrite]: #1194
-#1189 := (iff #1048 #1188)
-#1186 := (= #1047 #1185)
-#1187 := [rewrite]: #1186
-#1190 := [monotonicity #1187]: #1189
-#1197 := [trans #1190 #1195]: #1196
-#1200 := [monotonicity #1197]: #1199
-#1203 := [monotonicity #1200]: #1202
-#1206 := [monotonicity #1203]: #1205
-#1209 := [monotonicity #1206]: #1208
-#1183 := (iff #1011 #48)
-#1184 := [rewrite]: #1183
-#1181 := (iff #1008 #46)
-#1182 := [rewrite]: #1181
-#1179 := (iff #1005 #44)
-#1180 := [rewrite]: #1179
-#1212 := [monotonicity #1176 #850 #1180 #1182 #1184 #1209]: #1211
-#1217 := [trans #1212 #1215]: #1216
-#1309 := [monotonicity #1217 #1306]: #1308
-#1177 := (iff #1000 #41)
-#1178 := [rewrite]: #1177
-#1173 := (iff #987 #28)
-#1174 := [rewrite]: #1173
-#1312 := [monotonicity #1174 #1176 #1178 #1309]: #1311
-#1317 := [trans #1312 #1315]: #1316
-#1320 := [monotonicity #1317]: #1319
-#1323 := [monotonicity #1320]: #1322
-#1326 := [monotonicity #1323]: #1325
-#957 := (not #921)
-#1168 := (~ #957 #1167)
-#1164 := (not #918)
-#1165 := (~ #1164 #1163)
-#1160 := (not #915)
-#1161 := (~ #1160 #1159)
-#1156 := (not #910)
-#1157 := (~ #1156 #1155)
-#1152 := (not #867)
-#1153 := (~ #1152 #1151)
-#1148 := (not #862)
-#1149 := (~ #1148 #1147)
-#1144 := (not #841)
-#1145 := (~ #1144 #1143)
-#1140 := (not #836)
-#1141 := (~ #1140 #1139)
-#1124 := (not #753)
-#1125 := (~ #1124 #1123)
-#1120 := (not #750)
-#1121 := (~ #1120 #1119)
-#1117 := (~ #1116 #1116)
-#1118 := [refl]: #1117
-#1113 := (not #747)
-#1114 := (~ #1113 #744)
-#1111 := (~ #744 #744)
-#1109 := (~ #741 #741)
-#1110 := [refl]: #1109
-#1112 := [nnf-pos #1110]: #1111
-#1115 := [nnf-neg #1112]: #1114
-#1122 := [nnf-neg #1115 #1118]: #1121
-#1105 := (~ #747 #1104)
-#1106 := [sk]: #1105
-#1126 := [nnf-neg #1106 #1122]: #1125
-#1091 := (~ #1090 #1090)
-#1092 := [refl]: #1091
-#1088 := (~ #1087 #1087)
-#1089 := [refl]: #1088
-#1085 := (~ #1084 #1084)
-#1086 := [refl]: #1085
-#1137 := (~ #1136 #1136)
-#1138 := [refl]: #1137
-#1134 := (~ #1133 #1133)
-#1135 := [refl]: #1134
-#1131 := (~ #778 #778)
-#1132 := [refl]: #1131
-#991 := (~ #990 #990)
-#992 := [refl]: #991
-#1142 := [nnf-neg #992 #1132 #1135 #1138 #1086 #1089 #1092 #1126]: #1141
-#1128 := (not #812)
-#1129 := (~ #1128 #1127)
-#1082 := (~ #1081 #1081)
-#1083 := [refl]: #1082
-#1079 := (~ #1078 #1078)
-#1080 := [refl]: #1079
-#1076 := (~ #1075 #1075)
-#1077 := [refl]: #1076
-#1073 := (~ #1072 #1072)
-#1074 := [refl]: #1073
-#1070 := (~ #772 #772)
-#1071 := [refl]: #1070
-#1130 := [nnf-neg #992 #1071 #1074 #1077 #1080 #1083 #1086 #1089 #1092 #1126]: #1129
-#1146 := [nnf-neg #1130 #1142]: #1145
-#1068 := (~ #686 #686)
-#1069 := [refl]: #1068
-#1150 := [nnf-neg #992 #1069 #1146]: #1149
-#1065 := (not #707)
-#1066 := (~ #1065 #1064)
-#1061 := (not #678)
-#1062 := (~ #1061 #1060)
-#1057 := (not #675)
-#1058 := (~ #1057 #1056)
-#1051 := (not #672)
-#1052 := (~ #1051 #1050)
-#1053 := [sk]: #1052
-#1035 := (not #661)
-#1036 := (~ #1035 #1030)
-#1031 := (~ #658 #1030)
-#1032 := [sk]: #1031
-#1037 := [nnf-neg #1032]: #1036
-#1059 := [nnf-neg #1037 #1053]: #1058
-#1018 := (~ #661 #1017)
-#1015 := (~ #1014 #1014)
-#1016 := [refl]: #1015
-#1019 := [nnf-neg #1016]: #1018
-#1063 := [nnf-neg #1019 #1059]: #1062
-#1012 := (~ #1011 #1011)
-#1013 := [refl]: #1012
-#1009 := (~ #1008 #1008)
-#1010 := [refl]: #1009
-#1006 := (~ #1005 #1005)
-#1007 := [refl]: #1006
-#1003 := (~ #846 #846)
-#1004 := [refl]: #1003
-#1067 := [nnf-neg #992 #1004 #1007 #1010 #1013 #1063]: #1066
-#1154 := [nnf-neg #1067 #1150]: #1153
-#1001 := (~ #1000 #1000)
-#1002 := [refl]: #1001
-#997 := (not #892)
-#998 := (~ #997 #889)
-#995 := (~ #889 #889)
-#993 := (~ #886 #886)
-#994 := [refl]: #993
-#996 := [nnf-pos #994]: #995
-#999 := [nnf-neg #996]: #998
-#988 := (~ #987 #987)
-#989 := [refl]: #988
-#1158 := [nnf-neg #989 #992 #999 #1002 #1154]: #1157
-#985 := (~ #524 #524)
-#986 := [refl]: #985
-#1162 := [nnf-neg #986 #1158]: #1161
-#982 := (not #628)
-#983 := (~ #982 #625)
-#980 := (~ #625 #625)
-#978 := (~ #622 #622)
-#979 := [refl]: #978
-#981 := [nnf-pos #979]: #980
-#984 := [nnf-neg #981]: #983
-#1166 := [nnf-neg #984 #1162]: #1165
-#974 := (~ #628 #973)
-#975 := [sk]: #974
-#1169 := [nnf-neg #975 #1166]: #1168
-#958 := [not-or-elim #954]: #957
-#1170 := [mp~ #958 #1169]: #1167
-#1171 := [mp #1170 #1326]: #1324
-#1601 := [mp #1171 #1600]: #1598
-#2181 := [mp #1601 #2180]: #2178
-#1676 := [unit-resolution #2181 #1868]: #2175
-#1946 := (or #2172 #2166)
-#1947 := [def-axiom]: #1946
-#1672 := [unit-resolution #1947 #1676]: #2166
-#1668 := (or #2169 #2163)
-#1666 := (iff #13 #28)
-#1677 := (iff #28 #13)
-#1664 := [commutativity]: #1677
-#1667 := [symm #1664]: #1666
-#1665 := [mp #956 #1667]: #28
-#1945 := (or #2169 #524 #2163)
-#1941 := [def-axiom]: #1945
-#2182 := [unit-resolution #1941 #1665]: #1668
-#2183 := [unit-resolution #2182 #1672]: #2163
-#1952 := (or #2160 #2154)
-#1954 := [def-axiom]: #1952
-#2263 := [unit-resolution #1954 #2183]: #2154
-#2221 := [hypothesis]: #2107
-#1769 := (or #2104 #2098)
-#2043 := [def-axiom]: #1769
-#2222 := [unit-resolution #2043 #2221]: #2098
-#2058 := (not #2093)
-#1669 := (or #2104 #46)
-#2045 := [def-axiom]: #1669
-#2223 := [unit-resolution #2045 #2221]: #46
-#2256 := (or #2058 #221)
-#2231 := (= #40 f11)
-#2228 := (* -1::Int f7)
-#2229 := (+ f3 #2228)
-#2230 := (<= #2229 0::Int)
-#2232 := (or #1470 #2230 #2231)
-#1785 := (= f9 f11)
-#2234 := [hypothesis]: #46
-#2251 := [symm #2234]: #1785
-#1973 := (or #2160 #41)
-#1951 := [def-axiom]: #1973
-#2235 := [unit-resolution #1951 #2183]: #41
-#2252 := [trans #2235 #2251]: #2231
-#2246 := (not #2231)
-#2247 := (or #2232 #2246)
-#2248 := [def-axiom]: #2247
-#2253 := [unit-resolution #2248 #2252]: #2232
-#2254 := [hypothesis]: #2093
-#2233 := (not #2232)
-#2236 := (or #2058 #2233)
-#2237 := [quant-inst #29]: #2236
-#2255 := [unit-resolution #2237 #2254 #2253]: false
-#2257 := [lemma #2255]: #2256
-#2224 := [unit-resolution #2257 #2223]: #2058
-#1702 := (or #2101 #2093 #1459)
-#2062 := [def-axiom]: #1702
-#2225 := [unit-resolution #2062 #2224 #2222]: #1459
-#1715 := (or #1458 #1039)
-#1716 := [def-axiom]: #1715
-#2226 := [unit-resolution #1716 #2225]: #1039
-#1744 := (+ f8 #1040)
-#1745 := (<= #1744 0::Int)
-#1695 := (not #1745)
-#1772 := (or #2104 #683)
-#1773 := [def-axiom]: #1772
-#2227 := [unit-resolution #1773 #2221]: #683
-#1717 := (or #1458 #1043)
-#2053 := [def-axiom]: #1717
-#2210 := [unit-resolution #2053 #2225]: #1043
-#1678 := (or #1695 #686 #1042)
-#1691 := [hypothesis]: #1043
-#1692 := [hypothesis]: #1745
-#1694 := [hypothesis]: #683
-#1690 := [th-lemma arith farkas -1 -1 1 #1694 #1692 #1691]: false
-#1681 := [lemma #1690]: #1678
-#2211 := [unit-resolution #1681 #2210 #2227]: #1695
-#1721 := (+ f9 #1191)
-#1722 := (>= #1721 0::Int)
-#1680 := (not #1722)
-#1787 := (+ f9 #664)
-#1776 := (<= #1787 0::Int)
-#2209 := [symm #2223]: #1785
-#2212 := (not #1785)
-#2213 := (or #2212 #1776)
-#2214 := [th-lemma arith triangle-eq]: #2213
-#2215 := [unit-resolution #2214 #2209]: #1776
-#2054 := (not #1193)
-#2055 := (or #1458 #2054)
-#2056 := [def-axiom]: #2055
-#2216 := [unit-resolution #2056 #2225]: #2054
-#1688 := (not #1776)
-#1673 := (or #1680 #1193 #1688)
-#1682 := [hypothesis]: #1776
-#1685 := [hypothesis]: #2054
-#1686 := [hypothesis]: #1722
-#1687 := [th-lemma arith farkas -1 1 1 #1686 #1685 #1682]: false
-#1670 := [lemma #1687]: #1673
-#2217 := [unit-resolution #1670 #2216 #2215]: #1680
-#1707 := (or #1431 #1745 #1722)
-#1671 := [hypothesis]: #1680
-#1674 := [hypothesis]: #1695
-#1675 := [hypothesis]: #1039
-#1972 := (or #2160 #2085)
-#1962 := [def-axiom]: #1972
-#2184 := [unit-resolution #1962 #2183]: #2085
-#1763 := (or #2090 #1431 #1745 #1722)
-#1757 := (+ #1046 #881)
-#1767 := (<= #1757 0::Int)
-#1771 := (+ ?v0!2 #681)
-#1781 := (>= #1771 0::Int)
-#1734 := (or #1431 #1781 #1767)
-#1764 := (or #2090 #1734)
-#1683 := (iff #1764 #1763)
-#1703 := (or #2090 #1707)
-#1704 := (iff #1703 #1763)
-#1679 := [rewrite]: #1704
-#1698 := (iff #1764 #1703)
-#1708 := (iff #1734 #1707)
-#1724 := (iff #1767 #1722)
-#1732 := (+ #881 #1046)
-#1718 := (<= #1732 0::Int)
-#1723 := (iff #1718 #1722)
-#1712 := [rewrite]: #1723
-#1719 := (iff #1767 #1718)
-#1711 := (= #1757 #1732)
-#1713 := [rewrite]: #1711
-#1720 := [monotonicity #1713]: #1719
-#1705 := [trans #1720 #1712]: #1724
-#1729 := (iff #1781 #1745)
-#1736 := (+ #681 ?v0!2)
-#1741 := (>= #1736 0::Int)
-#1735 := (iff #1741 #1745)
-#1746 := [rewrite]: #1735
-#1742 := (iff #1781 #1741)
-#1737 := (= #1771 #1736)
-#1728 := [rewrite]: #1737
-#1743 := [monotonicity #1728]: #1742
-#1731 := [trans #1743 #1746]: #1729
-#1706 := [monotonicity #1731 #1705]: #1708
-#1709 := [monotonicity #1706]: #1698
-#1684 := [trans #1709 #1679]: #1683
-#1700 := [quant-inst #1038]: #1764
-#1689 := [mp #1700 #1684]: #1763
-#2185 := [unit-resolution #1689 #2184 #1675 #1674 #1671]: false
-#2186 := [lemma #2185]: #1707
-#2241 := [unit-resolution #2186 #2217 #2211 #2226]: false
-#2242 := [lemma #2241]: #2104
-#1964 := (or #2157 #2107 #2151)
-#1965 := [def-axiom]: #1964
-#2261 := [unit-resolution #1965 #2242 #2263]: #2151
-#1977 := (or #2148 #2142)
-#1978 := [def-axiom]: #1977
-#2427 := [unit-resolution #1978 #2261]: #2142
-#2348 := (= #68 #1100)
-#2362 := (not #2348)
-#2349 := (+ #68 #1260)
-#2351 := (>= #2349 0::Int)
-#2356 := (not #2351)
-#2250 := (+ #68 #736)
-#2218 := (<= #2250 0::Int)
-#2249 := (= #68 f15)
-#2383 := (= f13 f15)
-#2378 := [hypothesis]: #2133
-#1859 := (or #2130 #76)
-#2012 := [def-axiom]: #1859
-#2379 := [unit-resolution #2012 #2378]: #76
-#2384 := [symm #2379]: #2383
-#2381 := (= #68 f13)
-#2020 := (or #2130 #71)
-#2027 := [def-axiom]: #2020
-#2380 := [unit-resolution #2027 #2378]: #71
-#2382 := [symm #2380]: #2381
-#2385 := [trans #2382 #2384]: #2249
-#2386 := (not #2249)
-#2387 := (or #2386 #2218)
-#2388 := [th-lemma arith triangle-eq]: #2387
-#2389 := [unit-resolution #2388 #2385]: #2218
-#2040 := (not #1262)
-#1860 := (or #2130 #2124)
-#2008 := [def-axiom]: #1860
-#2390 := [unit-resolution #2008 #2378]: #2124
-#2394 := (= #90 f13)
-#2392 := (= #90 #68)
-#1856 := (or #2130 #74)
-#1858 := [def-axiom]: #1856
-#2391 := [unit-resolution #1858 #2378]: #74
-#2393 := [monotonicity #2391]: #2392
-#2395 := [trans #2393 #2382]: #2394
-#2396 := [trans #2395 #2384]: #91
-#2032 := (or #2118 #1116)
-#2028 := [def-axiom]: #2032
-#2397 := [unit-resolution #2028 #2396]: #2118
-#2023 := (or #2127 #1498 #2121)
-#2024 := [def-axiom]: #2023
-#2398 := [unit-resolution #2024 #2397 #2390]: #1498
-#1755 := (or #1493 #2040)
-#2037 := [def-axiom]: #1755
-#2399 := [unit-resolution #2037 #2398]: #2040
-#2357 := (not #2218)
-#2358 := (or #2356 #1262 #2357)
-#2352 := [hypothesis]: #2351
-#2353 := [hypothesis]: #2218
-#2354 := [hypothesis]: #2040
-#2355 := [th-lemma arith farkas -1 -1 1 #2354 #2353 #2352]: false
-#2359 := [lemma #2355]: #2358
-#2400 := [unit-resolution #2359 #2399 #2389]: #2356
-#2363 := (or #2362 #2351)
-#2364 := [th-lemma arith triangle-eq]: #2363
-#2401 := [unit-resolution #2364 #2400]: #2362
-#2350 := (= f8 ?v0!3)
-#2016 := (or #2130 #756)
-#2014 := [def-axiom]: #2016
-#2402 := [unit-resolution #2014 #2378]: #756
-#2403 := (or #760 #1758)
-#2404 := [th-lemma arith triangle-eq]: #2403
-#2405 := [unit-resolution #2404 #2402]: #1758
-#1761 := (or #1493 #1245)
-#2039 := [def-axiom]: #1761
-#2406 := [unit-resolution #2039 #2398]: #1245
-#2407 := (not #1758)
-#2408 := (or #2347 #1240 #2407)
-#2409 := [th-lemma arith assign-bounds 1 1]: #2408
-#2410 := [unit-resolution #2409 #2406 #2405]: #2347
-#2326 := (+ f9 #1260)
-#2327 := (>= #2326 0::Int)
-#2412 := (not #2327)
-#2025 := (or #2130 #772)
-#2026 := [def-axiom]: #2025
-#2411 := [unit-resolution #2026 #2378]: #772
-#2413 := (or #2412 #1262 #2357 #773)
-#2414 := [th-lemma arith assign-bounds 1 1 1]: #2413
-#2415 := [unit-resolution #2414 #2399 #2389 #2411]: #2412
-#2417 := (or #2315 #2327)
-#1759 := (or #1493 #1094)
-#1760 := [def-axiom]: #1759
-#2416 := [unit-resolution #1760 #2398]: #1094
-#2335 := (or #2090 #1478 #2315 #2327)
-#2305 := (+ #1100 #881)
-#2306 := (<= #2305 0::Int)
-#2297 := (+ ?v0!3 #681)
-#2298 := (>= #2297 0::Int)
-#2307 := (or #1478 #2298 #2306)
-#2336 := (or #2090 #2307)
-#2343 := (iff #2336 #2335)
-#2332 := (or #1478 #2315 #2327)
-#2338 := (or #2090 #2332)
-#2341 := (iff #2338 #2335)
-#2342 := [rewrite]: #2341
-#2339 := (iff #2336 #2338)
-#2333 := (iff #2307 #2332)
-#2330 := (iff #2306 #2327)
-#2320 := (+ #881 #1100)
-#2323 := (<= #2320 0::Int)
-#2328 := (iff #2323 #2327)
-#2329 := [rewrite]: #2328
-#2324 := (iff #2306 #2323)
-#2321 := (= #2305 #2320)
-#2322 := [rewrite]: #2321
-#2325 := [monotonicity #2322]: #2324
-#2331 := [trans #2325 #2329]: #2330
-#2318 := (iff #2298 #2315)
-#2308 := (+ #681 ?v0!3)
-#2311 := (>= #2308 0::Int)
-#2316 := (iff #2311 #2315)
-#2317 := [rewrite]: #2316
-#2312 := (iff #2298 #2311)
-#2309 := (= #2297 #2308)
-#2310 := [rewrite]: #2309
-#2313 := [monotonicity #2310]: #2312
-#2319 := [trans #2313 #2317]: #2318
-#2334 := [monotonicity #2319 #2331]: #2333
-#2340 := [monotonicity #2334]: #2339
-#2344 := [trans #2340 #2342]: #2343
-#2337 := [quant-inst #1093]: #2336
-#2345 := [mp #2337 #2344]: #2335
-#2418 := [unit-resolution #2345 #2184 #2416]: #2417
-#2419 := [unit-resolution #2418 #2415]: #2315
-#2421 := (not #2347)
-#2422 := (or #2350 #2420 #2421)
-#2423 := [th-lemma arith triangle-eq]: #2422
-#2424 := [unit-resolution #2423 #2419 #2410]: #2350
-#2375 := (not #2350)
-#2376 := (or #2375 #2348)
-#2371 := (= #1100 #68)
-#2369 := (= ?v0!3 f8)
-#2368 := [hypothesis]: #2350
-#2370 := [symm #2368]: #2369
-#2372 := [monotonicity #2370]: #2371
-#2373 := [symm #2372]: #2348
-#2367 := [hypothesis]: #2362
-#2374 := [unit-resolution #2367 #2373]: false
-#2377 := [lemma #2374]: #2376
-#2425 := [unit-resolution #2377 #2424 #2401]: false
-#2426 := [lemma #2425]: #2130
-#1984 := (or #2145 #2133 #2139)
-#1985 := [def-axiom]: #1984
-#2428 := [unit-resolution #1985 #2426 #2427]: #2139
-#1991 := (or #2136 #756)
-#2001 := [def-axiom]: #1991
-#2429 := [unit-resolution #2001 #2428]: #756
-#2430 := [unit-resolution #2404 #2429]: #1758
-#1993 := (or #2136 #2124)
-#1994 := [def-axiom]: #1993
-#2431 := [unit-resolution #1994 #2428]: #2124
-#1912 := (= f9 f15)
-#1998 := (or #2136 #108)
-#2000 := [def-axiom]: #1998
-#2432 := [unit-resolution #2000 #2428]: #108
-#2436 := [symm #2432]: #1912
-#2437 := (= #90 f9)
-#2434 := (= #90 #40)
-#2007 := (or #2136 #107)
-#2011 := [def-axiom]: #2007
-#2433 := [unit-resolution #2011 #2428]: #107
-#2435 := [monotonicity #2433]: #2434
-#2438 := [trans #2435 #2235]: #2437
-#2439 := [trans #2438 #2436]: #91
-#2440 := [unit-resolution #2028 #2439]: #2118
-#2441 := [unit-resolution #2024 #2440 #2431]: #1498
-#2442 := [unit-resolution #2039 #2441]: #1245
-#2443 := [unit-resolution #2409 #2442 #2430]: #2347
-#1917 := (or #2136 #773)
-#2010 := [def-axiom]: #1917
-#2444 := [unit-resolution #2010 #2428]: #773
-#1905 := (+ f9 #736)
-#1913 := (<= #1905 0::Int)
-#1808 := (or #395 #1912)
-#1804 := (iff #108 #1912)
-#1802 := (iff #1912 #108)
-#1803 := [commutativity]: #1802
-#1805 := [symm #1803]: #1804
-#1801 := [hypothesis]: #108
-#1806 := [mp #1801 #1805]: #1912
-#1798 := (not #1912)
-#1800 := [hypothesis]: #1798
-#1807 := [unit-resolution #1800 #1806]: false
-#1809 := [lemma #1807]: #1808
-#2445 := [unit-resolution #1809 #2432]: #1912
-#2446 := (or #1798 #1913)
-#2447 := [th-lemma arith triangle-eq]: #2446
-#2448 := [unit-resolution #2447 #2445]: #1913
-#2449 := (not #1913)
-#2450 := (or #2218 #772 #2449)
-#2451 := [th-lemma arith assign-bounds 1 -1]: #2450
-#2452 := [unit-resolution #2451 #2448 #2444]: #2218
-#2453 := [unit-resolution #2037 #2441]: #2040
-#2454 := [unit-resolution #2359 #2453 #2452]: #2356
-#2455 := [unit-resolution #2364 #2454]: #2362
-#2456 := [unit-resolution #2377 #2455]: #2375
-#2457 := [unit-resolution #2423 #2456 #2443]: #2420
-#2458 := (or #2412 #1262 #2449)
-#2459 := [th-lemma arith assign-bounds -1 -1]: #2458
-#2460 := [unit-resolution #2459 #2448 #2453]: #2412
-#2461 := [unit-resolution #1760 #2441]: #1094
-[unit-resolution #2345 #2184 #2461 #2460 #2457]: false
-unsat
--- a/src/HOL/Boogie/Examples/Boogie_Max.thy	Tue Jul 23 13:14:14 2013 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,50 +0,0 @@
-(*  Title:      HOL/Boogie/Examples/Boogie_Max.thy
-    Author:     Sascha Boehme, TU Muenchen
-*)
-
-header {* Boogie example: get the greatest element of a list *}
-
-theory Boogie_Max
-imports Boogie
-begin
-
-text {*
-We prove correct the verification condition generated from the following
-Boogie code:
-
-\begin{verbatim}
-procedure max(array : [int]int, length : int)
-  returns (max : int);
-  requires (0 < length);
-  ensures (forall i: int :: 0 <= i && i < length ==> array[i] <= max);
-  ensures (exists i: int :: 0 <= i && i < length ==> array[i] == max);
-
-implementation max(array : [int]int, length : int) returns (max : int)
-{
-  var p : int, k : int;
-  max := array[0];
-  k := 0;
-  p := 1;
-  while (p < length)
-    invariant (forall i: int :: 0 <= i && i < p ==> array[i] <= max);
-    invariant (array[k] == max);
-  {
-    if (array[p] > max) { max := array[p]; k := p; }
-    p := p + 1;
-  }
-}
-\end{verbatim}
-*}
-
-boogie_open "Boogie_Max.b2i"
-
-declare [[smt_oracle = false]]
-declare [[smt_certificates = "Boogie_Max.certs"]]
-declare [[smt_read_only_certificates = true]]
-
-boogie_vc max
-  by boogie
-
-boogie_end
-
-end
--- a/src/HOL/Boogie/Examples/Boogie_Max_Stepwise.thy	Tue Jul 23 13:14:14 2013 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,115 +0,0 @@
-(*  Title:      HOL/Boogie/Examples/Boogie_Max_Stepwise.thy
-    Author:     Sascha Boehme, TU Muenchen
-*)
-
-header {* Boogie example: get the greatest element of a list *}
-
-theory Boogie_Max_Stepwise
-imports Boogie
-begin
-
-text {*
-We prove correct the verification condition generated from the following
-Boogie code:
-
-\begin{verbatim}
-procedure max(array : [int]int, length : int)
-  returns (max : int);
-  requires (0 < length);
-  ensures (forall i: int :: 0 <= i && i < length ==> array[i] <= max);
-  ensures (exists i: int :: 0 <= i && i < length ==> array[i] == max);
-
-implementation max(array : [int]int, length : int) returns (max : int)
-{
-  var p : int, k : int;
-  max := array[0];
-  k := 0;
-  p := 1;
-  while (p < length)
-    invariant (forall i: int :: 0 <= i && i < p ==> array[i] <= max);
-    invariant (array[k] == max);
-  {
-    if (array[p] > max) { max := array[p]; k := p; }
-    p := p + 1;
-  }
-}
-\end{verbatim}
-*}
-
-boogie_open "Boogie_Max.b2i"
-
-boogie_status -- {* gives an overview of all verification conditions *}
-
-boogie_status max -- {* shows the names of all unproved assertions
-  of the verification condition @{text max} *}
-
-boogie_status (full) max -- {* shows the state of all assertions
-  of the verification condition @{text max} *}
-
-boogie_status (paths) max -- {* shows the names of all unproved assertions,
-  grouped by program paths, of the verification condition @{text max} *}
-
-boogie_status (full paths) max -- {* shows the state of all assertions
-  of the verification condition @{text max}, grouped by program paths,
-  with already proved or irrelevant assertions written in parentheses *}
-
-
-text {* Let's find out which assertions of @{text max} are hard to prove: *}
-
-boogie_status (narrow step_timeout: 4 final_timeout: 15) max
-  (simp add: labels, (fast | metis)?)
-  -- {* The argument @{text step_timeout} is optional, restricting the runtime
-        of each intermediate narrowing step by the given number of seconds. *}
-  -- {* The argument @{text final_timeout} is optional, restricting the
-        runtime of the final narrowing steps by the given number of seconds. *}
-  -- {* The last argument should be a method to be applied at each narrowing
-        step. Note that different methods may be composed to one method by
-        a language similar to regular expressions. *}
-
-boogie_status (scan timeout: 4) max
-  (simp add: labels, (fast | metis)?)
-  -- {* finds the first hard assertion wrt. to the given method *}
-
-
-text {* Now, let's prove the three hard assertions of @{text max}: *}
-
-boogie_vc max (examine: L_14_5c L_14_5b L_14_5a) 
-proof boogie_cases
-  case L_14_5c
-  thus ?case by (metis less_le_not_le zle_add1_eq_le less_linear)
-next
-  case L_14_5b
-  thus ?case by (metis less_le_not_le xt1(10) linorder_linear zless_add1_eq)
-next
-  case L_14_5a
-  note max0 = `max0 = array 0`
-  {
-    fix i :: int
-    assume "0 \<le> i \<and> i < 1"
-    hence "i = 0" by simp
-    with max0 have "array i \<le> max0" by simp
-  }
-  thus ?case by simp
-qed
-
-
-boogie_status max -- {* the above proved assertions are not shown anymore *}
-
-boogie_status (full) max -- {* states the above proved assertions as proved
-  and all other assertions as not proved *}
-
-boogie_status (paths) max
-
-boogie_status (full paths) max
-
-
-text {* Now we prove the remaining assertions all at once: *}
-
-boogie_vc max by (auto simp add: labels)
-
-
-boogie_status -- {* @{text max} has been completely proved *}
-
-boogie_end
-
-end
--- a/src/HOL/Boogie/Examples/VCC_Max.b2i	Tue Jul 23 13:14:14 2013 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,19619 +0,0 @@
-type-decl $ctype 0 0
-type-decl $ptr 0 0
-type-decl $field 0 0
-type-decl $kind 0 0
-type-decl $type_state 0 0
-type-decl $status 0 0
-type-decl $primitive 0 0
-type-decl $struct 0 0
-type-decl $token 0 0
-type-decl $state 0 0
-type-decl $pure_function 0 0
-type-decl $label 0 0
-type-decl $memory_t 0 0
-type-decl $typemap_t 0 0
-type-decl $statusmap_t 0 0
-type-decl $record 0 0
-type-decl $version 0 0
-type-decl $vol_version 0 0
-type-decl $ptrset 0 0
-fun-decl $kind_of 2 0
-    type-con $ctype 0
-    type-con $kind 0
-fun-decl $kind_composite 1 1
-    type-con $kind 0
-  attribute unique 0
-fun-decl $kind_primitive 1 1
-    type-con $kind 0
-  attribute unique 0
-fun-decl $kind_array 1 1
-    type-con $kind 0
-  attribute unique 0
-fun-decl $kind_thread 1 1
-    type-con $kind 0
-  attribute unique 0
-fun-decl $sizeof 2 0
-    type-con $ctype 0
-    int
-fun-decl ^^i1 1 1
-    type-con $ctype 0
-  attribute unique 0
-fun-decl ^^i2 1 1
-    type-con $ctype 0
-  attribute unique 0
-fun-decl ^^i4 1 1
-    type-con $ctype 0
-  attribute unique 0
-fun-decl ^^i8 1 1
-    type-con $ctype 0
-  attribute unique 0
-fun-decl ^^u1 1 1
-    type-con $ctype 0
-  attribute unique 0
-fun-decl ^^u2 1 1
-    type-con $ctype 0
-  attribute unique 0
-fun-decl ^^u4 1 1
-    type-con $ctype 0
-  attribute unique 0
-fun-decl ^^u8 1 1
-    type-con $ctype 0
-  attribute unique 0
-fun-decl ^^void 1 1
-    type-con $ctype 0
-  attribute unique 0
-fun-decl ^^bool 1 1
-    type-con $ctype 0
-  attribute unique 0
-fun-decl ^^f4 1 1
-    type-con $ctype 0
-  attribute unique 0
-fun-decl ^^f8 1 1
-    type-con $ctype 0
-  attribute unique 0
-fun-decl ^^claim 1 1
-    type-con $ctype 0
-  attribute unique 0
-fun-decl ^^root_emb 1 1
-    type-con $ctype 0
-  attribute unique 0
-fun-decl ^^mathint 1 1
-    type-con $ctype 0
-  attribute unique 0
-fun-decl ^$#thread_id_t 1 1
-    type-con $ctype 0
-  attribute unique 0
-fun-decl ^$#ptrset 1 1
-    type-con $ctype 0
-  attribute unique 0
-fun-decl ^$#state_t 1 1
-    type-con $ctype 0
-  attribute unique 0
-fun-decl ^$#struct 1 1
-    type-con $ctype 0
-  attribute unique 0
-fun-decl $ptr_to 2 0
-    type-con $ctype 0
-    type-con $ctype 0
-fun-decl $unptr_to 2 0
-    type-con $ctype 0
-    type-con $ctype 0
-fun-decl $ptr_level 2 0
-    type-con $ctype 0
-    int
-fun-decl $map_t 3 0
-    type-con $ctype 0
-    type-con $ctype 0
-    type-con $ctype 0
-fun-decl $map_domain 2 0
-    type-con $ctype 0
-    type-con $ctype 0
-fun-decl $map_range 2 0
-    type-con $ctype 0
-    type-con $ctype 0
-fun-decl $is_primitive 2 1
-    type-con $ctype 0
-    bool
-  attribute weight 1
-    expr-attr
-      int-num 0
-fun-decl $is_primitive_ch 2 1
-    type-con $ctype 0
-    bool
-  attribute inline 1
-    expr-attr
-      true
-fun-decl $is_composite 2 1
-    type-con $ctype 0
-    bool
-  attribute weight 1
-    expr-attr
-      int-num 0
-fun-decl $is_composite_ch 2 1
-    type-con $ctype 0
-    bool
-  attribute inline 1
-    expr-attr
-      true
-fun-decl $is_arraytype 2 1
-    type-con $ctype 0
-    bool
-  attribute weight 1
-    expr-attr
-      int-num 0
-fun-decl $is_arraytype_ch 2 1
-    type-con $ctype 0
-    bool
-  attribute inline 1
-    expr-attr
-      true
-fun-decl $is_threadtype 2 1
-    type-con $ctype 0
-    bool
-  attribute weight 1
-    expr-attr
-      int-num 0
-fun-decl $is_thread 2 1
-    type-con $ptr 0
-    bool
-  attribute inline 1
-    expr-attr
-      true
-fun-decl $is_ptr_to_composite 2 1
-    type-con $ptr 0
-    bool
-  attribute inline 1
-    expr-attr
-      true
-fun-decl $field_offset 2 0
-    type-con $field 0
-    int
-fun-decl $field_parent_type 2 0
-    type-con $field 0
-    type-con $ctype 0
-fun-decl $is_non_primitive 2 0
-    type-con $ctype 0
-    bool
-fun-decl $is_non_primitive_ch 2 1
-    type-con $ctype 0
-    bool
-  attribute inline 1
-    expr-attr
-      true
-fun-decl $is_non_primitive_ptr 2 1
-    type-con $ptr 0
-    bool
-  attribute inline 1
-    expr-attr
-      true
-fun-decl $me_ref 1 0
-    int
-fun-decl $me 1 0
-    type-con $ptr 0
-fun-decl $current_state 2 1
-    type-con $state 0
-    type-con $state 0
-  attribute inline 1
-    expr-attr
-      true
-fun-decl $select.mem 3 0
-    type-con $memory_t 0
-    type-con $ptr 0
-    int
-fun-decl $store.mem 4 0
-    type-con $memory_t 0
-    type-con $ptr 0
-    int
-    type-con $memory_t 0
-fun-decl $select.tm 3 0
-    type-con $typemap_t 0
-    type-con $ptr 0
-    type-con $type_state 0
-fun-decl $store.tm 4 0
-    type-con $typemap_t 0
-    type-con $ptr 0
-    type-con $type_state 0
-    type-con $typemap_t 0
-fun-decl $select.sm 3 0
-    type-con $statusmap_t 0
-    type-con $ptr 0
-    type-con $status 0
-fun-decl $store.sm 4 0
-    type-con $statusmap_t 0
-    type-con $ptr 0
-    type-con $status 0
-    type-con $statusmap_t 0
-fun-decl $memory 2 0
-    type-con $state 0
-    type-con $memory_t 0
-fun-decl $typemap 2 0
-    type-con $state 0
-    type-con $typemap_t 0
-fun-decl $statusmap 2 0
-    type-con $state 0
-    type-con $statusmap_t 0
-fun-decl $mem 3 1
-    type-con $state 0
-    type-con $ptr 0
-    int
-  attribute inline 1
-    expr-attr
-      true
-fun-decl $read_any 3 1
-    type-con $state 0
-    type-con $ptr 0
-    int
-  attribute inline 1
-    expr-attr
-      true
-fun-decl $mem_eq 4 1
-    type-con $state 0
-    type-con $state 0
-    type-con $ptr 0
-    bool
-  attribute inline 1
-    expr-attr
-      true
-fun-decl $st_eq 4 1
-    type-con $state 0
-    type-con $state 0
-    type-con $ptr 0
-    bool
-  attribute inline 1
-    expr-attr
-      true
-fun-decl $ts_eq 4 1
-    type-con $state 0
-    type-con $state 0
-    type-con $ptr 0
-    bool
-  attribute inline 1
-    expr-attr
-      true
-fun-decl $extent_hint 3 0
-    type-con $ptr 0
-    type-con $ptr 0
-    bool
-fun-decl $nesting_level 2 0
-    type-con $ctype 0
-    int
-fun-decl $is_nested 3 0
-    type-con $ctype 0
-    type-con $ctype 0
-    bool
-fun-decl $nesting_min 3 0
-    type-con $ctype 0
-    type-con $ctype 0
-    int
-fun-decl $nesting_max 3 0
-    type-con $ctype 0
-    type-con $ctype 0
-    int
-fun-decl $is_nested_range 5 0
-    type-con $ctype 0
-    type-con $ctype 0
-    int
-    int
-    bool
-fun-decl $typ 2 0
-    type-con $ptr 0
-    type-con $ctype 0
-fun-decl $ref 2 0
-    type-con $ptr 0
-    int
-fun-decl $ptr 3 0
-    type-con $ctype 0
-    int
-    type-con $ptr 0
-fun-decl $ghost_ref 3 0
-    type-con $ptr 0
-    type-con $field 0
-    int
-fun-decl $ghost_emb 2 0
-    int
-    type-con $ptr 0
-fun-decl $ghost_path 2 0
-    int
-    type-con $field 0
-fun-decl $physical_ref 3 0
-    type-con $ptr 0
-    type-con $field 0
-    int
-fun-decl $array_path 3 0
-    type-con $field 0
-    int
-    type-con $field 0
-fun-decl $is_base_field 2 0
-    type-con $field 0
-    bool
-fun-decl $array_path_1 2 0
-    type-con $field 0
-    type-con $field 0
-fun-decl $array_path_2 2 0
-    type-con $field 0
-    int
-fun-decl $null 1 0
-    type-con $ptr 0
-fun-decl $is 3 0
-    type-con $ptr 0
-    type-con $ctype 0
-    bool
-fun-decl $ptr_cast 3 1
-    type-con $ptr 0
-    type-con $ctype 0
-    type-con $ptr 0
-  attribute inline 1
-    expr-attr
-      true
-fun-decl $read_ptr 4 1
-    type-con $state 0
-    type-con $ptr 0
-    type-con $ctype 0
-    type-con $ptr 0
-  attribute inline 1
-    expr-attr
-      true
-fun-decl $dot 3 0
-    type-con $ptr 0
-    type-con $field 0
-    type-con $ptr 0
-fun-decl $emb 3 1
-    type-con $state 0
-    type-con $ptr 0
-    type-con $ptr 0
-  attribute inline 1
-    expr-attr
-      true
-fun-decl $path 3 1
-    type-con $state 0
-    type-con $ptr 0
-    type-con $field 0
-  attribute inline 1
-    expr-attr
-      true
-fun-decl $containing_struct 3 0
-    type-con $ptr 0
-    type-con $field 0
-    type-con $ptr 0
-fun-decl $containing_struct_ref 3 0
-    type-con $ptr 0
-    type-con $field 0
-    int
-fun-decl $is_primitive_non_volatile_field 2 0
-    type-con $field 0
-    bool
-fun-decl $is_primitive_volatile_field 2 0
-    type-con $field 0
-    bool
-fun-decl $is_primitive_embedded_array 3 0
-    type-con $field 0
-    int
-    bool
-fun-decl $is_primitive_embedded_volatile_array 4 0
-    type-con $field 0
-    int
-    type-con $ctype 0
-    bool
-fun-decl $static_field_properties 3 1
-    type-con $field 0
-    type-con $ctype 0
-    bool
-  attribute inline 1
-    expr-attr
-      true
-fun-decl $field_properties 6 1
-    type-con $state 0
-    type-con $ptr 0
-    type-con $field 0
-    type-con $ctype 0
-    bool
-    bool
-  attribute inline 1
-    expr-attr
-      true
-fun-decl $ts_typed 2 0
-    type-con $type_state 0
-    bool
-fun-decl $ts_emb 2 0
-    type-con $type_state 0
-    type-con $ptr 0
-fun-decl $ts_path 2 0
-    type-con $type_state 0
-    type-con $field 0
-fun-decl $ts_is_array_elt 2 0
-    type-con $type_state 0
-    bool
-fun-decl $ts_is_volatile 2 0
-    type-con $type_state 0
-    bool
-fun-decl $is_object_root 3 1
-    type-con $state 0
-    type-con $ptr 0
-    bool
-  attribute inline 1
-    expr-attr
-      true
-fun-decl $is_volatile 3 1
-    type-con $state 0
-    type-con $ptr 0
-    bool
-  attribute inline 1
-    expr-attr
-      true
-fun-decl $is_malloc_root 3 1
-    type-con $state 0
-    type-con $ptr 0
-    bool
-  attribute inline 1
-    expr-attr
-      true
-fun-decl $current_timestamp 2 0
-    type-con $state 0
-    int
-fun-decl $is_fresh 4 1
-    type-con $state 0
-    type-con $state 0
-    type-con $ptr 0
-    bool
-  attribute inline 1
-    expr-attr
-      true
-fun-decl $in_writes_at 3 0
-    int
-    type-con $ptr 0
-    bool
-fun-decl $writable 4 1
-    type-con $state 0
-    int
-    type-con $ptr 0
-    bool
-  attribute inline 1
-    expr-attr
-      true
-fun-decl $top_writable 4 1
-    type-con $state 0
-    int
-    type-con $ptr 0
-    bool
-  attribute inline 1
-    expr-attr
-      true
-fun-decl $struct_zero 1 0
-    type-con $struct 0
-fun-decl $vs_base 3 1
-    type-con $struct 0
-    type-con $ctype 0
-    type-con $ptr 0
-  attribute inline 1
-    expr-attr
-      true
-fun-decl $vs_base_ref 2 0
-    type-con $struct 0
-    int
-fun-decl $vs_state 2 0
-    type-con $struct 0
-    type-con $state 0
-fun-decl $vs_ctor 3 0
-    type-con $state 0
-    type-con $ptr 0
-    type-con $struct 0
-fun-decl $rec_zero 1 0
-    type-con $record 0
-fun-decl $rec_update 4 0
-    type-con $record 0
-    type-con $field 0
-    int
-    type-con $record 0
-fun-decl $rec_fetch 3 0
-    type-con $record 0
-    type-con $field 0
-    int
-fun-decl $rec_update_bv 7 0
-    type-con $record 0
-    type-con $field 0
-    int
-    int
-    int
-    int
-    type-con $record 0
-fun-decl $is_record_type 2 0
-    type-con $ctype 0
-    bool
-fun-decl $is_record_field 4 0
-    type-con $ctype 0
-    type-con $field 0
-    type-con $ctype 0
-    bool
-fun-decl $as_record_record_field 2 0
-    type-con $field 0
-    type-con $field 0
-fun-decl $rec_eq 3 0
-    type-con $record 0
-    type-con $record 0
-    bool
-fun-decl $rec_base_eq 3 0
-    int
-    int
-    bool
-fun-decl $int_to_record 2 0
-    int
-    type-con $record 0
-fun-decl $record_to_int 2 0
-    type-con $record 0
-    int
-fun-decl $good_state 2 0
-    type-con $state 0
-    bool
-fun-decl $invok_state 2 0
-    type-con $state 0
-    bool
-fun-decl $has_volatile_owns_set 2 0
-    type-con $ctype 0
-    bool
-fun-decl $owns_set_field 2 0
-    type-con $ctype 0
-    type-con $field 0
-fun-decl $st_owner 2 0
-    type-con $status 0
-    type-con $ptr 0
-fun-decl $st_closed 2 0
-    type-con $status 0
-    bool
-fun-decl $st_timestamp 2 0
-    type-con $status 0
-    int
-fun-decl $st_ref_cnt 2 0
-    type-con $status 0
-    int
-fun-decl $owner 3 0
-    type-con $state 0
-    type-con $ptr 0
-    type-con $ptr 0
-fun-decl $closed 3 0
-    type-con $state 0
-    type-con $ptr 0
-    bool
-fun-decl $timestamp 3 0
-    type-con $state 0
-    type-con $ptr 0
-    int
-fun-decl $position_marker 1 0
-    bool
-fun-decl $st 3 1
-    type-con $state 0
-    type-con $ptr 0
-    type-con $status 0
-  attribute inline 1
-    expr-attr
-      true
-fun-decl $ts 3 1
-    type-con $state 0
-    type-con $ptr 0
-    type-con $type_state 0
-  attribute inline 1
-    expr-attr
-      true
-fun-decl $owns 3 1
-    type-con $state 0
-    type-con $ptr 0
-    type-con $ptrset 0
-  attribute weight 1
-    expr-attr
-      int-num 0
-fun-decl $nested 3 1
-    type-con $state 0
-    type-con $ptr 0
-    bool
-  attribute inline 1
-    expr-attr
-      true
-fun-decl $nested_in 4 1
-    type-con $state 0
-    type-con $ptr 0
-    type-con $ptr 0
-    bool
-  attribute inline 1
-    expr-attr
-      true
-fun-decl $wrapped 4 1
-    type-con $state 0
-    type-con $ptr 0
-    type-con $ctype 0
-    bool
-  attribute inline 1
-    expr-attr
-      true
-fun-decl $irrelevant 3 1
-    type-con $state 0
-    type-con $ptr 0
-    bool
-  attribute inline 1
-    expr-attr
-      true
-fun-decl $mutable 3 1
-    type-con $state 0
-    type-con $ptr 0
-    bool
-  attribute weight 1
-    expr-attr
-      int-num 0
-fun-decl $thread_owned 3 1
-    type-con $state 0
-    type-con $ptr 0
-    bool
-  attribute inline 1
-    expr-attr
-      true
-fun-decl $thread_owned_or_even_mutable 3 1
-    type-con $state 0
-    type-con $ptr 0
-    bool
-  attribute inline 1
-    expr-attr
-      true
-fun-decl $typed 3 0
-    type-con $state 0
-    type-con $ptr 0
-    bool
-fun-decl $typed2 4 1
-    type-con $state 0
-    type-con $ptr 0
-    type-con $ctype 0
-    bool
-  attribute inline 1
-    expr-attr
-      true
-fun-decl $ptr_eq 3 1
-    type-con $ptr 0
-    type-con $ptr 0
-    bool
-  attribute inline 1
-    expr-attr
-      true
-fun-decl $ptr_neq 3 1
-    type-con $ptr 0
-    type-con $ptr 0
-    bool
-  attribute inline 1
-    expr-attr
-      true
-fun-decl $is_primitive_field_of 4 1
-    type-con $state 0
-    type-con $ptr 0
-    type-con $ptr 0
-    bool
-  attribute inline 1
-    expr-attr
-      true
-fun-decl $instantiate_st 2 0
-    type-con $status 0
-    bool
-fun-decl $is_domain_root 3 0
-    type-con $state 0
-    type-con $ptr 0
-    bool
-fun-decl $in_wrapped_domain 3 0
-    type-con $state 0
-    type-con $ptr 0
-    bool
-fun-decl $thread_local_np 3 1
-    type-con $state 0
-    type-con $ptr 0
-    bool
-  attribute inline 1
-    expr-attr
-      true
-fun-decl $thread_local 3 0
-    type-con $state 0
-    type-con $ptr 0
-    bool
-fun-decl $thread_local2 4 1
-    type-con $state 0
-    type-con $ptr 0
-    type-con $ctype 0
-    bool
-  attribute inline 1
-    expr-attr
-      true
-fun-decl $dont_instantiate 2 0
-    type-con $ptr 0
-    bool
-fun-decl $dont_instantiate_int 2 0
-    int
-    bool
-fun-decl $dont_instantiate_state 2 0
-    type-con $state 0
-    bool
-fun-decl $instantiate_int 2 0
-    int
-    bool
-fun-decl $instantiate_bool 2 0
-    bool
-    bool
-fun-decl $instantiate_ptr 2 0
-    type-con $ptr 0
-    bool
-fun-decl $instantiate_ptrset 2 0
-    type-con $ptrset 0
-    bool
-fun-decl $inv 4 1
-    type-con $state 0
-    type-con $ptr 0
-    type-con $ctype 0
-    bool
-  attribute inline 1
-    expr-attr
-      true
-fun-decl $inv2nt 4 1
-    type-con $state 0
-    type-con $state 0
-    type-con $ptr 0
-    bool
-  attribute inline 1
-    expr-attr
-      true
-fun-decl $imply_inv 4 0
-    type-con $state 0
-    type-con $ptr 0
-    type-con $ctype 0
-    bool
-fun-decl $inv2 5 0
-    type-con $state 0
-    type-con $state 0
-    type-con $ptr 0
-    type-con $ctype 0
-    bool
-fun-decl $inv2_when_closed 5 1
-    type-con $state 0
-    type-con $state 0
-    type-con $ptr 0
-    type-con $ctype 0
-    bool
-  attribute inline 1
-    expr-attr
-      true
-fun-decl $sequential 5 1
-    type-con $state 0
-    type-con $state 0
-    type-con $ptr 0
-    type-con $ctype 0
-    bool
-  attribute weight 1
-    expr-attr
-      int-num 0
-fun-decl $depends 5 1
-    type-con $state 0
-    type-con $state 0
-    type-con $ptr 0
-    type-con $ptr 0
-    bool
-  attribute weight 1
-    expr-attr
-      int-num 0
-fun-decl $spans_the_same 5 1
-    type-con $state 0
-    type-con $state 0
-    type-con $ptr 0
-    type-con $ctype 0
-    bool
-  attribute weight 1
-    expr-attr
-      int-num 0
-fun-decl $state_spans_the_same 5 0
-    type-con $state 0
-    type-con $state 0
-    type-con $ptr 0
-    type-con $ctype 0
-    bool
-fun-decl $nonvolatile_spans_the_same 5 1
-    type-con $state 0
-    type-con $state 0
-    type-con $ptr 0
-    type-con $ctype 0
-    bool
-  attribute weight 1
-    expr-attr
-      int-num 0
-fun-decl $state_nonvolatile_spans_the_same 5 0
-    type-con $state 0
-    type-con $state 0
-    type-con $ptr 0
-    type-con $ctype 0
-    bool
-fun-decl $in_extent_of 4 1
-    type-con $state 0
-    type-con $ptr 0
-    type-con $ptr 0
-    bool
-  attribute inline 1
-    expr-attr
-      true
-fun-decl $in_full_extent_of 3 1
-    type-con $ptr 0
-    type-con $ptr 0
-    bool
-  attribute inline 1
-    expr-attr
-      true
-fun-decl $extent_mutable 3 0
-    type-con $state 0
-    type-con $ptr 0
-    bool
-fun-decl $extent_is_fresh 4 0
-    type-con $state 0
-    type-con $state 0
-    type-con $ptr 0
-    bool
-fun-decl $forall_inv2_when_closed 3 1
-    type-con $state 0
-    type-con $state 0
-    bool
-  attribute inline 1
-    expr-attr
-      true
-fun-decl $function_entry 2 0
-    type-con $state 0
-    bool
-fun-decl $full_stop 2 0
-    type-con $state 0
-    bool
-fun-decl $full_stop_ext 3 1
-    type-con $token 0
-    type-con $state 0
-    bool
-  attribute inline 1
-    expr-attr
-      true
-fun-decl $file_name_is 3 0
-    int
-    type-con $token 0
-    bool
-fun-decl $closed_is_transitive 2 1
-    type-con $state 0
-    bool
-  attribute inline 1
-    expr-attr
-      true
-fun-decl $call_transition 3 0
-    type-con $state 0
-    type-con $state 0
-    bool
-fun-decl $good_state_ext 3 0
-    type-con $token 0
-    type-con $state 0
-    bool
-fun-decl $local_value_is 6 0
-    type-con $state 0
-    type-con $token 0
-    type-con $token 0
-    int
-    type-con $ctype 0
-    bool
-fun-decl $local_value_is_ptr 6 0
-    type-con $state 0
-    type-con $token 0
-    type-con $token 0
-    type-con $ptr 0
-    type-con $ctype 0
-    bool
-fun-decl $read_ptr_m 4 0
-    type-con $state 0
-    type-con $ptr 0
-    type-con $ctype 0
-    type-con $ptr 0
-fun-decl $type_code_is 3 0
-    int
-    type-con $ctype 0
-    bool
-fun-decl $function_arg_type 4 0
-    type-con $pure_function 0
-    int
-    type-con $ctype 0
-    bool
-fun-decl $ver_domain 2 0
-    type-con $version 0
-    type-con $ptrset 0
-fun-decl $read_version 3 1
-    type-con $state 0
-    type-con $ptr 0
-    type-con $version 0
-  attribute weight 1
-    expr-attr
-      int-num 0
-fun-decl $domain 3 1
-    type-con $state 0
-    type-con $ptr 0
-    type-con $ptrset 0
-  attribute weight 1
-    expr-attr
-      int-num 0
-fun-decl $in_domain 4 0
-    type-con $state 0
-    type-con $ptr 0
-    type-con $ptr 0
-    bool
-fun-decl $in_vdomain 4 0
-    type-con $state 0
-    type-con $ptr 0
-    type-con $ptr 0
-    bool
-fun-decl $in_domain_lab 5 0
-    type-con $state 0
-    type-con $ptr 0
-    type-con $ptr 0
-    type-con $label 0
-    bool
-fun-decl $in_vdomain_lab 5 0
-    type-con $state 0
-    type-con $ptr 0
-    type-con $ptr 0
-    type-con $label 0
-    bool
-fun-decl $inv_lab 4 0
-    type-con $state 0
-    type-con $ptr 0
-    type-con $label 0
-    bool
-fun-decl $dom_thread_local 3 1
-    type-con $state 0
-    type-con $ptr 0
-    bool
-  attribute inline 1
-    expr-attr
-      true
-fun-decl $fetch_from_domain 3 0
-    type-con $version 0
-    type-con $ptr 0
-    int
-fun-decl $in_claim_domain 3 0
-    type-con $ptr 0
-    type-con $ptr 0
-    bool
-fun-decl $by_claim 5 1
-    type-con $state 0
-    type-con $ptr 0
-    type-con $ptr 0
-    type-con $ptr 0
-    type-con $ptr 0
-  attribute weight 1
-    expr-attr
-      int-num 0
-fun-decl $claim_version 2 0
-    type-con $ptr 0
-    type-con $version 0
-fun-decl $read_vol_version 3 1
-    type-con $state 0
-    type-con $ptr 0
-    type-con $vol_version 0
-  attribute weight 1
-    expr-attr
-      int-num 0
-fun-decl $fetch_from_vv 3 0
-    type-con $vol_version 0
-    type-con $ptr 0
-    int
-fun-decl $fetch_vol_field 4 1
-    type-con $state 0
-    type-con $ptr 0
-    type-con $field 0
-    int
-  attribute inline 1
-    expr-attr
-      true
-fun-decl $is_approved_by 4 0
-    type-con $ctype 0
-    type-con $field 0
-    type-con $field 0
-    bool
-fun-decl $inv_is_approved_by_ptr 6 1
-    type-con $state 0
-    type-con $state 0
-    type-con $ptr 0
-    type-con $ptr 0
-    type-con $field 0
-    bool
-  attribute inline 1
-    expr-attr
-      true
-fun-decl $inv_is_approved_by 6 1
-    type-con $state 0
-    type-con $state 0
-    type-con $ptr 0
-    type-con $field 0
-    type-con $field 0
-    bool
-  attribute inline 1
-    expr-attr
-      true
-fun-decl $is_owner_approved 3 0
-    type-con $ctype 0
-    type-con $field 0
-    bool
-fun-decl $inv_is_owner_approved 5 1
-    type-con $state 0
-    type-con $state 0
-    type-con $ptr 0
-    type-con $field 0
-    bool
-  attribute inline 1
-    expr-attr
-      true
-fun-decl $good_for_admissibility 2 0
-    type-con $state 0
-    bool
-fun-decl $good_for_post_admissibility 2 0
-    type-con $state 0
-    bool
-fun-decl $stuttering_pre 3 1
-    type-con $state 0
-    type-con $ptr 0
-    bool
-  attribute inline 1
-    expr-attr
-      true
-fun-decl $admissibility_pre 3 1
-    type-con $state 0
-    type-con $ptr 0
-    bool
-  attribute inline 1
-    expr-attr
-      true
-fun-decl $mutable_increases 3 1
-    type-con $state 0
-    type-con $state 0
-    bool
-  attribute inline 1
-    expr-attr
-      true
-fun-decl $meta_eq 3 1
-    type-con $state 0
-    type-con $state 0
-    bool
-  attribute inline 1
-    expr-attr
-      true
-fun-decl $is_stuttering_check 1 0
-    bool
-fun-decl $is_unwrap_check 1 0
-    bool
-fun-decl $is_admissibility_check 1 1
-    bool
-  attribute inline 1
-    expr-attr
-      true
-fun-decl $good_for_pre_can_unwrap 2 0
-    type-con $state 0
-    bool
-fun-decl $good_for_post_can_unwrap 2 0
-    type-con $state 0
-    bool
-fun-decl $unwrap_check_pre 3 1
-    type-con $state 0
-    type-con $ptr 0
-    bool
-  attribute inline 1
-    expr-attr
-      true
-fun-decl $update_int 4 0
-    type-con $state 0
-    type-con $ptr 0
-    int
-    type-con $state 0
-fun-decl $timestamp_is_now 3 1
-    type-con $state 0
-    type-con $ptr 0
-    bool
-  attribute inline 1
-    expr-attr
-      true
-fun-decl $now_writable 3 1
-    type-con $state 0
-    type-con $ptr 0
-    bool
-  attribute inline 1
-    expr-attr
-      true
-fun-decl $timestamp_post 3 1
-    type-con $state 0
-    type-con $state 0
-    bool
-  attribute inline 1
-    expr-attr
-      true
-fun-decl $timestamp_post_strict 3 1
-    type-con $state 0
-    type-con $state 0
-    bool
-  attribute inline 1
-    expr-attr
-      true
-fun-decl $pre_wrap 2 0
-    type-con $state 0
-    bool
-fun-decl $pre_unwrap 2 0
-    type-con $state 0
-    bool
-fun-decl $pre_static_wrap 2 0
-    type-con $state 0
-    bool
-fun-decl $pre_static_unwrap 2 0
-    type-con $state 0
-    bool
-fun-decl $unwrap_post 5 1
-    type-con $state 0
-    type-con $state 0
-    type-con $ptr 0
-    type-con $ptr 0
-    bool
-  attribute inline 1
-    expr-attr
-      true
-fun-decl $unwrap_post_claimable 5 1
-    type-con $state 0
-    type-con $state 0
-    type-con $ptr 0
-    type-con $ptr 0
-    bool
-  attribute inline 1
-    expr-attr
-      true
-fun-decl $keeps 4 2
-    type-con $state 0
-    type-con $ptr 0
-    type-con $ptr 0
-    bool
-  attribute inline 1
-    expr-attr
-      true
-  attribute expand 1
-    expr-attr
-      true
-fun-decl $expect_unreachable 1 0
-    bool
-fun-decl $taken_over 4 0
-    type-con $state 0
-    type-con $ptr 0
-    type-con $ptr 0
-    type-con $status 0
-fun-decl $take_over 4 0
-    type-con $state 0
-    type-con $ptr 0
-    type-con $ptr 0
-    type-con $state 0
-fun-decl $released 4 0
-    type-con $state 0
-    type-con $ptr 0
-    type-con $ptr 0
-    type-con $status 0
-fun-decl $release 5 0
-    type-con $state 0
-    type-con $state 0
-    type-con $ptr 0
-    type-con $ptr 0
-    type-con $state 0
-fun-decl $post_unwrap 3 0
-    type-con $state 0
-    type-con $state 0
-    bool
-fun-decl $new_ownees 4 1
-    type-con $state 0
-    type-con $ptr 0
-    type-con $ptrset 0
-    type-con $ptrset 0
-  attribute inline 1
-    expr-attr
-      true
-fun-decl $get_memory_allocator 1 0
-    type-con $ptr 0
-fun-decl $memory_allocator_type 1 1
-    type-con $ctype 0
-  attribute unique 0
-fun-decl $memory_allocator_ref 1 0
-    int
-fun-decl $program_entry_point 2 0
-    type-con $state 0
-    bool
-fun-decl $program_entry_point_ch 2 0
-    type-con $state 0
-    bool
-fun-decl $is_global 3 1
-    type-con $ptr 0
-    type-con $ctype 0
-    bool
-  attribute inline 1
-    expr-attr
-      true
-fun-decl $is_global_array 4 1
-    type-con $ptr 0
-    type-con $ctype 0
-    int
-    bool
-  attribute inline 1
-    expr-attr
-      true
-fun-decl $active_option 3 1
-    type-con $state 0
-    type-con $ptr 0
-    type-con $field 0
-  attribute inline 1
-    expr-attr
-      true
-fun-decl $ts_active_option 2 0
-    type-con $type_state 0
-    type-con $field 0
-fun-decl $union_active 4 1
-    type-con $state 0
-    type-con $ptr 0
-    type-con $field 0
-    bool
-  attribute inline 1
-    expr-attr
-      true
-fun-decl $is_union_field 3 0
-    type-con $ctype 0
-    type-con $field 0
-    bool
-fun-decl $union_havoced 4 1
-    type-con $state 0
-    type-con $state 0
-    type-con $ptr 0
-    bool
-  attribute inline 1
-    expr-attr
-      true
-fun-decl $full_extent 2 0
-    type-con $ptr 0
-    type-con $ptrset 0
-fun-decl $extent 3 0
-    type-con $state 0
-    type-con $ptr 0
-    type-con $ptrset 0
-fun-decl $span 2 0
-    type-con $ptr 0
-    type-con $ptrset 0
-fun-decl $in_span_of 3 1
-    type-con $ptr 0
-    type-con $ptr 0
-    bool
-  attribute inline 1
-    expr-attr
-      true
-fun-decl $first_option_typed 3 0
-    type-con $state 0
-    type-con $ptr 0
-    bool
-fun-decl $struct_extent 2 1
-    type-con $ptr 0
-    type-con $ptrset 0
-  attribute inline 1
-    expr-attr
-      true
-fun-decl $in_struct_extent_of 3 1
-    type-con $ptr 0
-    type-con $ptr 0
-    bool
-  attribute inline 1
-    expr-attr
-      true
-fun-decl $volatile_span 3 0
-    type-con $state 0
-    type-con $ptr 0
-    type-con $ptrset 0
-fun-decl $left_split 3 0
-    type-con $ptr 0
-    int
-    type-con $ptr 0
-fun-decl $right_split 3 0
-    type-con $ptr 0
-    int
-    type-con $ptr 0
-fun-decl $joined_array 3 0
-    type-con $ptr 0
-    type-con $ptr 0
-    type-con $ptr 0
-fun-decl $mutable_root 3 1
-    type-con $state 0
-    type-con $ptr 0
-    bool
-  attribute inline 1
-    expr-attr
-      true
-fun-decl $set_in 3 0
-    type-con $ptr 0
-    type-con $ptrset 0
-    bool
-fun-decl $set_empty 1 0
-    type-con $ptrset 0
-fun-decl $set_singleton 2 0
-    type-con $ptr 0
-    type-con $ptrset 0
-fun-decl $non_null_set_singleton 2 0
-    type-con $ptr 0
-    type-con $ptrset 0
-fun-decl $set_union 3 0
-    type-con $ptrset 0
-    type-con $ptrset 0
-    type-con $ptrset 0
-fun-decl $set_difference 3 0
-    type-con $ptrset 0
-    type-con $ptrset 0
-    type-con $ptrset 0
-fun-decl $set_intersection 3 0
-    type-con $ptrset 0
-    type-con $ptrset 0
-    type-con $ptrset 0
-fun-decl $set_subset 3 0
-    type-con $ptrset 0
-    type-con $ptrset 0
-    bool
-fun-decl $set_eq 3 0
-    type-con $ptrset 0
-    type-con $ptrset 0
-    bool
-fun-decl $set_cardinality 2 0
-    type-con $ptrset 0
-    int
-fun-decl $set_universe 1 0
-    type-con $ptrset 0
-fun-decl $set_disjoint 3 0
-    type-con $ptrset 0
-    type-con $ptrset 0
-    bool
-fun-decl $id_set_disjoint 4 0
-    type-con $ptr 0
-    type-con $ptrset 0
-    type-con $ptrset 0
-    int
-fun-decl $set_in3 3 0
-    type-con $ptr 0
-    type-con $ptrset 0
-    bool
-fun-decl $set_in2 3 0
-    type-con $ptr 0
-    type-con $ptrset 0
-    bool
-fun-decl $in_some_owns 2 0
-    type-con $ptr 0
-    bool
-fun-decl $set_in0 3 0
-    type-con $ptr 0
-    type-con $ptrset 0
-    bool
-fun-decl sk_hack 2 0
-    bool
-    bool
-fun-decl $writes_nothing 3 1
-    type-con $state 0
-    type-con $state 0
-    bool
-  attribute inline 1
-    expr-attr
-      true
-fun-decl $array 3 0
-    type-con $ctype 0
-    int
-    type-con $ctype 0
-fun-decl $element_type 2 0
-    type-con $ctype 0
-    type-con $ctype 0
-fun-decl $array_length 2 0
-    type-con $ctype 0
-    int
-fun-decl $is_array_elt 3 1
-    type-con $state 0
-    type-con $ptr 0
-    bool
-  attribute inline 1
-    expr-attr
-      true
-fun-decl $inlined_array 3 1
-    type-con $ptr 0
-    type-con $ctype 0
-    type-con $ptr 0
-  attribute weight 1
-    expr-attr
-      int-num 0
-fun-decl $idx 4 0
-    type-con $ptr 0
-    int
-    type-con $ctype 0
-    type-con $ptr 0
-fun-decl $add.mul 4 2
-    int
-    int
-    int
-    int
-  attribute inline 1
-    expr-attr
-      true
-  attribute expand 1
-    expr-attr
-      true
-fun-decl $add 3 2
-    int
-    int
-    int
-  attribute inline 1
-    expr-attr
-      true
-  attribute expand 1
-    expr-attr
-      true
-fun-decl $is_array_vol_or_nonvol 6 1
-    type-con $state 0
-    type-con $ptr 0
-    type-con $ctype 0
-    int
-    bool
-    bool
-  attribute weight 1
-    expr-attr
-      int-num 0
-fun-decl $is_array 5 1
-    type-con $state 0
-    type-con $ptr 0
-    type-con $ctype 0
-    int
-    bool
-  attribute weight 1
-    expr-attr
-      int-num 0
-fun-decl $is_thread_local_array 5 1
-    type-con $state 0
-    type-con $ptr 0
-    type-con $ctype 0
-    int
-    bool
-  attribute inline 1
-    expr-attr
-      true
-fun-decl $is_mutable_array 5 1
-    type-con $state 0
-    type-con $ptr 0
-    type-con $ctype 0
-    int
-    bool
-  attribute inline 1
-    expr-attr
-      true
-fun-decl $is_array_emb 6 1
-    type-con $state 0
-    type-con $ptr 0
-    type-con $ctype 0
-    int
-    type-con $ptr 0
-    bool
-  attribute inline 1
-    expr-attr
-      true
-fun-decl $is_array_emb_path 8 1
-    type-con $state 0
-    type-con $ptr 0
-    type-con $ctype 0
-    int
-    type-con $ptr 0
-    type-con $field 0
-    bool
-    bool
-  attribute inline 1
-    expr-attr
-      true
-fun-decl $array_field_properties 6 1
-    type-con $field 0
-    type-con $ctype 0
-    int
-    bool
-    bool
-    bool
-  attribute inline 1
-    expr-attr
-      true
-fun-decl $no_inline_array_field_properties 6 1
-    type-con $field 0
-    type-con $ctype 0
-    int
-    bool
-    bool
-    bool
-  attribute inline 1
-    expr-attr
-      true
-fun-decl $array_elt_emb 4 1
-    type-con $state 0
-    type-con $ptr 0
-    type-con $ptr 0
-    bool
-  attribute inline 1
-    expr-attr
-      true
-fun-decl $array_members 4 0
-    type-con $ptr 0
-    type-con $ctype 0
-    int
-    type-con $ptrset 0
-fun-decl $array_range 4 0
-    type-con $ptr 0
-    type-con $ctype 0
-    int
-    type-con $ptrset 0
-fun-decl $non_null_array_range 4 0
-    type-con $ptr 0
-    type-con $ctype 0
-    int
-    type-con $ptrset 0
-fun-decl $non_null_extent 3 0
-    type-con $state 0
-    type-con $ptr 0
-    type-con $ptrset 0
-fun-decl $as_array 4 1
-    type-con $ptr 0
-    type-con $ctype 0
-    int
-    type-con $ptr 0
-  attribute inline 1
-    expr-attr
-      true
-fun-decl $array_eq 6 1
-    type-con $state 0
-    type-con $state 0
-    type-con $ptr 0
-    type-con $ctype 0
-    int
-    bool
-  attribute inline 1
-    expr-attr
-      true
-fun-decl $index_within 3 0
-    type-con $ptr 0
-    type-con $ptr 0
-    int
-fun-decl $in_array 5 1
-    type-con $ptr 0
-    type-con $ptr 0
-    type-con $ctype 0
-    int
-    bool
-  attribute inline 1
-    expr-attr
-      true
-fun-decl $in_array_full_extent_of 5 1
-    type-con $ptr 0
-    type-con $ptr 0
-    type-con $ctype 0
-    int
-    bool
-  attribute inline 1
-    expr-attr
-      true
-fun-decl $in_array_extent_of 6 1
-    type-con $state 0
-    type-con $ptr 0
-    type-con $ptr 0
-    type-con $ctype 0
-    int
-    bool
-  attribute inline 1
-    expr-attr
-      true
-fun-decl $in_range 4 1
-    int
-    int
-    int
-    bool
-  attribute inline 1
-    expr-attr
-      true
-fun-decl $bool_to_int 2 1
-    bool
-    int
-  attribute inline 1
-    expr-attr
-      true
-fun-decl $int_to_bool 2 1
-    int
-    bool
-  attribute inline 1
-    expr-attr
-      true
-fun-decl $read_bool 3 1
-    type-con $state 0
-    type-con $ptr 0
-    bool
-  attribute inline 1
-    expr-attr
-      true
-fun-decl $ite.int 4 3
-    bool
-    int
-    int
-    int
-  attribute external 1
-    string-attr ITE
-  attribute bvz 1
-    string-attr ITE
-  attribute bvint 1
-    string-attr ITE
-fun-decl $ite.bool 4 3
-    bool
-    bool
-    bool
-    bool
-  attribute external 1
-    string-attr ITE
-  attribute bvz 1
-    string-attr ITE
-  attribute bvint 1
-    string-attr ITE
-fun-decl $ite.ptr 4 3
-    bool
-    type-con $ptr 0
-    type-con $ptr 0
-    type-con $ptr 0
-  attribute external 1
-    string-attr ITE
-  attribute bvz 1
-    string-attr ITE
-  attribute bvint 1
-    string-attr ITE
-fun-decl $ite.struct 4 3
-    bool
-    type-con $struct 0
-    type-con $struct 0
-    type-con $struct 0
-  attribute external 1
-    string-attr ITE
-  attribute bvz 1
-    string-attr ITE
-  attribute bvint 1
-    string-attr ITE
-fun-decl $ite.ptrset 4 3
-    bool
-    type-con $ptrset 0
-    type-con $ptrset 0
-    type-con $ptrset 0
-  attribute external 1
-    string-attr ITE
-  attribute bvz 1
-    string-attr ITE
-  attribute bvint 1
-    string-attr ITE
-fun-decl $ite.primitive 4 3
-    bool
-    type-con $primitive 0
-    type-con $primitive 0
-    type-con $primitive 0
-  attribute external 1
-    string-attr ITE
-  attribute bvz 1
-    string-attr ITE
-  attribute bvint 1
-    string-attr ITE
-fun-decl $ite.record 4 3
-    bool
-    type-con $record 0
-    type-con $record 0
-    type-con $record 0
-  attribute external 1
-    string-attr ITE
-  attribute bvz 1
-    string-attr ITE
-  attribute bvint 1
-    string-attr ITE
-fun-decl $bool_id 2 1
-    bool
-    bool
-  attribute weight 1
-    expr-attr
-      int-num 0
-fun-decl $min.i1 1 0
-    int
-fun-decl $max.i1 1 0
-    int
-fun-decl $min.i2 1 0
-    int
-fun-decl $max.i2 1 0
-    int
-fun-decl $min.i4 1 0
-    int
-fun-decl $max.i4 1 0
-    int
-fun-decl $min.i8 1 0
-    int
-fun-decl $max.i8 1 0
-    int
-fun-decl $max.u1 1 0
-    int
-fun-decl $max.u2 1 0
-    int
-fun-decl $max.u4 1 0
-    int
-fun-decl $max.u8 1 0
-    int
-fun-decl $in_range_i1 2 1
-    int
-    bool
-  attribute inline 1
-    expr-attr
-      true
-fun-decl $in_range_i2 2 1
-    int
-    bool
-  attribute inline 1
-    expr-attr
-      true
-fun-decl $in_range_i4 2 1
-    int
-    bool
-  attribute inline 1
-    expr-attr
-      true
-fun-decl $in_range_i8 2 1
-    int
-    bool
-  attribute inline 1
-    expr-attr
-      true
-fun-decl $in_range_u1 2 1
-    int
-    bool
-  attribute inline 1
-    expr-attr
-      true
-fun-decl $in_range_u2 2 1
-    int
-    bool
-  attribute inline 1
-    expr-attr
-      true
-fun-decl $in_range_u4 2 1
-    int
-    bool
-  attribute inline 1
-    expr-attr
-      true
-fun-decl $in_range_u8 2 1
-    int
-    bool
-  attribute inline 1
-    expr-attr
-      true
-fun-decl $in_range_div_i1 3 1
-    int
-    int
-    bool
-  attribute inline 1
-    expr-attr
-      true
-fun-decl $in_range_div_i2 3 1
-    int
-    int
-    bool
-  attribute inline 1
-    expr-attr
-      true
-fun-decl $in_range_div_i4 3 1
-    int
-    int
-    bool
-  attribute inline 1
-    expr-attr
-      true
-fun-decl $in_range_div_i8 3 1
-    int
-    int
-    bool
-  attribute inline 1
-    expr-attr
-      true
-fun-decl $read_i1 3 1
-    type-con $state 0
-    type-con $ptr 0
-    int
-  attribute weight 1
-    expr-attr
-      int-num 0
-fun-decl $read_i2 3 1
-    type-con $state 0
-    type-con $ptr 0
-    int
-  attribute weight 1
-    expr-attr
-      int-num 0
-fun-decl $read_i4 3 1
-    type-con $state 0
-    type-con $ptr 0
-    int
-  attribute weight 1
-    expr-attr
-      int-num 0
-fun-decl $read_i8 3 1
-    type-con $state 0
-    type-con $ptr 0
-    int
-  attribute weight 1
-    expr-attr
-      int-num 0
-fun-decl $read_u1 3 1
-    type-con $state 0
-    type-con $ptr 0
-    int
-  attribute weight 1
-    expr-attr
-      int-num 0
-fun-decl $read_u2 3 1
-    type-con $state 0
-    type-con $ptr 0
-    int
-  attribute weight 1
-    expr-attr
-      int-num 0
-fun-decl $read_u4 3 1
-    type-con $state 0
-    type-con $ptr 0
-    int
-  attribute weight 1
-    expr-attr
-      int-num 0
-fun-decl $read_u8 3 1
-    type-con $state 0
-    type-con $ptr 0
-    int
-  attribute weight 1
-    expr-attr
-      int-num 0
-fun-decl $ptr_to_u8 2 0
-    type-con $ptr 0
-    int
-fun-decl $ptr_to_i8 2 0
-    type-con $ptr 0
-    int
-fun-decl $ptr_to_u4 2 0
-    type-con $ptr 0
-    int
-fun-decl $ptr_to_i4 2 0
-    type-con $ptr 0
-    int
-fun-decl $u8_to_ptr 2 1
-    int
-    type-con $ptr 0
-  attribute inline 1
-    expr-attr
-      true
-fun-decl $i8_to_ptr 2 1
-    int
-    type-con $ptr 0
-  attribute inline 1
-    expr-attr
-      true
-fun-decl $u4_to_ptr 2 1
-    int
-    type-con $ptr 0
-  attribute inline 1
-    expr-attr
-      true
-fun-decl $i4_to_ptr 2 1
-    int
-    type-con $ptr 0
-  attribute inline 1
-    expr-attr
-      true
-fun-decl $byte_ptr_subtraction 3 1
-    type-con $ptr 0
-    type-con $ptr 0
-    int
-  attribute weight 1
-    expr-attr
-      int-num 0
-fun-decl $_pow2 2 0
-    int
-    int
-fun-decl $_or 4 0
-    type-con $ctype 0
-    int
-    int
-    int
-fun-decl $_xor 4 0
-    type-con $ctype 0
-    int
-    int
-    int
-fun-decl $_and 4 0
-    type-con $ctype 0
-    int
-    int
-    int
-fun-decl $_not 3 0
-    type-con $ctype 0
-    int
-    int
-fun-decl $unchk_add 4 1
-    type-con $ctype 0
-    int
-    int
-    int
-  attribute weight 1
-    expr-attr
-      int-num 0
-fun-decl $unchk_sub 4 1
-    type-con $ctype 0
-    int
-    int
-    int
-  attribute weight 1
-    expr-attr
-      int-num 0
-fun-decl $unchk_mul 4 1
-    type-con $ctype 0
-    int
-    int
-    int
-  attribute weight 1
-    expr-attr
-      int-num 0
-fun-decl $unchk_div 4 1
-    type-con $ctype 0
-    int
-    int
-    int
-  attribute weight 1
-    expr-attr
-      int-num 0
-fun-decl $unchk_mod 4 1
-    type-con $ctype 0
-    int
-    int
-    int
-  attribute weight 1
-    expr-attr
-      int-num 0
-fun-decl $_shl 4 1
-    type-con $ctype 0
-    int
-    int
-    int
-  attribute weight 1
-    expr-attr
-      int-num 0
-fun-decl $_shr 3 1
-    int
-    int
-    int
-  attribute weight 1
-    expr-attr
-      int-num 0
-fun-decl $bv_extract_signed 5 0
-    int
-    int
-    int
-    int
-    int
-fun-decl $bv_extract_unsigned 5 0
-    int
-    int
-    int
-    int
-    int
-fun-decl $bv_update 6 0
-    int
-    int
-    int
-    int
-    int
-    int
-fun-decl $unchecked 3 0
-    type-con $ctype 0
-    int
-    int
-fun-decl $in_range_t 3 0
-    type-con $ctype 0
-    int
-    bool
-fun-decl $_mul 3 1
-    int
-    int
-    int
-  attribute weight 1
-    expr-attr
-      int-num 0
-fun-decl $get_string_literal 3 0
-    int
-    int
-    type-con $ptr 0
-fun-decl $get_fnptr 3 0
-    int
-    type-con $ctype 0
-    type-con $ptr 0
-fun-decl $get_fnptr_ref 2 0
-    int
-    int
-fun-decl $get_fnptr_inv 2 0
-    int
-    int
-fun-decl $is_fnptr_type 2 0
-    type-con $ctype 0
-    bool
-fun-decl $is_math_type 2 0
-    type-con $ctype 0
-    bool
-fun-decl $claims_obj 3 0
-    type-con $ptr 0
-    type-con $ptr 0
-    bool
-fun-decl $valid_claim 3 0
-    type-con $state 0
-    type-con $ptr 0
-    bool
-fun-decl $claim_initial_assumptions 4 1
-    type-con $state 0
-    type-con $ptr 0
-    type-con $token 0
-    bool
-  attribute inline 1
-    expr-attr
-      true
-fun-decl $claim_transitivity_assumptions 5 1
-    type-con $state 0
-    type-con $state 0
-    type-con $ptr 0
-    type-con $token 0
-    bool
-  attribute inline 1
-    expr-attr
-      true
-fun-decl $valid_claim_impl 3 1
-    type-con $state 0
-    type-con $state 0
-    bool
-  attribute inline 1
-    expr-attr
-      true
-fun-decl $claims_claim 3 0
-    type-con $ptr 0
-    type-con $ptr 0
-    bool
-fun-decl $not_shared 3 1
-    type-con $state 0
-    type-con $ptr 0
-    bool
-  attribute weight 1
-    expr-attr
-      int-num 0
-fun-decl $claimed_closed 3 1
-    type-con $state 0
-    type-con $ptr 0
-    bool
-  attribute weight 1
-    expr-attr
-      int-num 0
-fun-decl $no_claim 1 1
-    type-con $ptr 0
-  attribute unique 0
-fun-decl $ref_cnt 3 1
-    type-con $state 0
-    type-con $ptr 0
-    int
-  attribute weight 1
-    expr-attr
-      int-num 0
-fun-decl $is_claimable 2 0
-    type-con $ctype 0
-    bool
-fun-decl $is_thread_local_storage 2 0
-    type-con $ctype 0
-    bool
-fun-decl $frame_level 2 0
-    type-con $pure_function 0
-    int
-fun-decl $current_frame_level 1 0
-    int
-fun-decl $can_use_all_frame_axioms 2 1
-    type-con $state 0
-    bool
-  attribute inline 1
-    expr-attr
-      true
-fun-decl $can_use_frame_axiom_of 2 1
-    type-con $pure_function 0
-    bool
-  attribute inline 1
-    expr-attr
-      true
-fun-decl $reads_check_pre 2 0
-    type-con $state 0
-    bool
-fun-decl $reads_check_post 2 0
-    type-con $state 0
-    bool
-fun-decl $start_here 1 0
-    bool
-fun-decl $ptrset_to_int 2 0
-    type-con $ptrset 0
-    int
-fun-decl $int_to_ptrset 2 0
-    int
-    type-con $ptrset 0
-fun-decl $version_to_int 2 0
-    type-con $version 0
-    int
-fun-decl $int_to_version 2 0
-    int
-    type-con $version 0
-fun-decl $vol_version_to_int 2 0
-    type-con $vol_version 0
-    int
-fun-decl $int_to_vol_version 2 0
-    int
-    type-con $vol_version 0
-fun-decl $ptr_to_int 2 0
-    type-con $ptr 0
-    int
-fun-decl $int_to_ptr 2 0
-    int
-    type-con $ptr 0
-fun-decl $precise_test 2 0
-    type-con $ptr 0
-    bool
-fun-decl $updated_only_values 4 0
-    type-con $state 0
-    type-con $state 0
-    type-con $ptrset 0
-    bool
-fun-decl $updated_only_domains 4 0
-    type-con $state 0
-    type-con $state 0
-    type-con $ptrset 0
-    bool
-fun-decl $domain_updated_at 5 0
-    type-con $state 0
-    type-con $state 0
-    type-con $ptr 0
-    type-con $ptrset 0
-    bool
-fun-decl l#public 1 1
-    type-con $label 0
-  attribute unique 0
-fun-decl #tok$1^16.24 1 1
-    type-con $token 0
-  attribute unique 0
-fun-decl #tok$1^24.47 1 1
-    type-con $token 0
-  attribute unique 0
-fun-decl #tok$1^23.7 1 1
-    type-con $token 0
-  attribute unique 0
-fun-decl #tok$1^16.3 1 1
-    type-con $token 0
-  attribute unique 0
-fun-decl #loc.p 1 1
-    type-con $token 0
-  attribute unique 0
-fun-decl #tok$1^16.8 1 1
-    type-con $token 0
-  attribute unique 0
-fun-decl #loc.witness 1 1
-    type-con $token 0
-  attribute unique 0
-fun-decl #tok$1^14.3 1 1
-    type-con $token 0
-  attribute unique 0
-fun-decl #loc.max 1 1
-    type-con $token 0
-  attribute unique 0
-fun-decl #tok$1^12.3 1 1
-    type-con $token 0
-  attribute unique 0
-fun-decl #loc.len 1 1
-    type-con $token 0
-  attribute unique 0
-fun-decl #distTp1 1 1
-    type-con $ctype 0
-  attribute unique 0
-fun-decl #loc.arr 1 1
-    type-con $token 0
-  attribute unique 0
-fun-decl #tok$1^6.1 1 1
-    type-con $token 0
-  attribute unique 0
-fun-decl #file^Z?3A?5CC?5Cmax.c 1 1
-    type-con $token 0
-  attribute unique 0
-axiom 0
-    =
-    fun $sizeof 1
-    fun ^^i1 0
-    int-num 1
-axiom 0
-    =
-    fun $sizeof 1
-    fun ^^i2 0
-    int-num 2
-axiom 0
-    =
-    fun $sizeof 1
-    fun ^^i4 0
-    int-num 4
-axiom 0
-    =
-    fun $sizeof 1
-    fun ^^i8 0
-    int-num 8
-axiom 0
-    =
-    fun $sizeof 1
-    fun ^^u1 0
-    int-num 1
-axiom 0
-    =
-    fun $sizeof 1
-    fun ^^u2 0
-    int-num 2
-axiom 0
-    =
-    fun $sizeof 1
-    fun ^^u4 0
-    int-num 4
-axiom 0
-    =
-    fun $sizeof 1
-    fun ^^u8 0
-    int-num 8
-axiom 0
-    =
-    fun $sizeof 1
-    fun ^^f4 0
-    int-num 4
-axiom 0
-    =
-    fun $sizeof 1
-    fun ^^f8 0
-    int-num 8
-axiom 0
-    =
-    fun $sizeof 1
-    fun ^$#thread_id_t 0
-    int-num 1
-axiom 0
-    =
-    fun $sizeof 1
-    fun ^$#ptrset 0
-    int-num 1
-axiom 0
-    =
-    fun $ptr_level 1
-    fun ^^i1 0
-    int-num 0
-axiom 0
-    =
-    fun $ptr_level 1
-    fun ^^i2 0
-    int-num 0
-axiom 0
-    =
-    fun $ptr_level 1
-    fun ^^i4 0
-    int-num 0
-axiom 0
-    =
-    fun $ptr_level 1
-    fun ^^i8 0
-    int-num 0
-axiom 0
-    =
-    fun $ptr_level 1
-    fun ^^u1 0
-    int-num 0
-axiom 0
-    =
-    fun $ptr_level 1
-    fun ^^u2 0
-    int-num 0
-axiom 0
-    =
-    fun $ptr_level 1
-    fun ^^u4 0
-    int-num 0
-axiom 0
-    =
-    fun $ptr_level 1
-    fun ^^u8 0
-    int-num 0
-axiom 0
-    =
-    fun $ptr_level 1
-    fun ^^f4 0
-    int-num 0
-axiom 0
-    =
-    fun $ptr_level 1
-    fun ^^f8 0
-    int-num 0
-axiom 0
-    =
-    fun $ptr_level 1
-    fun ^^mathint 0
-    int-num 0
-axiom 0
-    =
-    fun $ptr_level 1
-    fun ^^bool 0
-    int-num 0
-axiom 0
-    =
-    fun $ptr_level 1
-    fun ^^void 0
-    int-num 0
-axiom 0
-    =
-    fun $ptr_level 1
-    fun ^^claim 0
-    int-num 0
-axiom 0
-    =
-    fun $ptr_level 1
-    fun ^^root_emb 0
-    int-num 0
-axiom 0
-    =
-    fun $ptr_level 1
-    fun ^$#ptrset 0
-    int-num 0
-axiom 0
-    =
-    fun $ptr_level 1
-    fun ^$#thread_id_t 0
-    int-num 0
-axiom 0
-    =
-    fun $ptr_level 1
-    fun ^$#state_t 0
-    int-num 0
-axiom 0
-    =
-    fun $ptr_level 1
-    fun ^$#struct 0
-    int-num 0
-axiom 0
-    fun $is_composite 1
-    fun ^^claim 0
-axiom 0
-    fun $is_composite 1
-    fun ^^root_emb 0
-axiom 0
-    forall 1 1 3
-      var #n
-        type-con $ctype 0
-      pat 1
-        fun $ptr_to 1
-        var #n
-          type-con $ctype 0
-      attribute qid 1
-        string-attr VccPrelu.145:15
-      attribute uniqueId 1
-        string-attr 4
-      attribute bvZ3Native 1
-        string-attr False
-    =
-    fun $unptr_to 1
-    fun $ptr_to 1
-    var #n
-      type-con $ctype 0
-    var #n
-      type-con $ctype 0
-axiom 0
-    forall 1 1 3
-      var #n
-        type-con $ctype 0
-      pat 1
-        fun $ptr_to 1
-        var #n
-          type-con $ctype 0
-      attribute qid 1
-        string-attr VccPrelu.146:15
-      attribute uniqueId 1
-        string-attr 5
-      attribute bvZ3Native 1
-        string-attr False
-    =
-    fun $sizeof 1
-    fun $ptr_to 1
-    var #n
-      type-con $ctype 0
-    int-num 8
-axiom 0
-    forall 2 1 3
-      var #r
-        type-con $ctype 0
-      var #d
-        type-con $ctype 0
-      pat 1
-        fun $map_t 2
-        var #r
-          type-con $ctype 0
-        var #d
-          type-con $ctype 0
-      attribute qid 1
-        string-attr VccPrelu.152:15
-      attribute uniqueId 1
-        string-attr 6
-      attribute bvZ3Native 1
-        string-attr False
-    =
-    fun $map_domain 1
-    fun $map_t 2
-    var #r
-      type-con $ctype 0
-    var #d
-      type-con $ctype 0
-    var #d
-      type-con $ctype 0
-axiom 0
-    forall 2 1 3
-      var #r
-        type-con $ctype 0
-      var #d
-        type-con $ctype 0
-      pat 1
-        fun $map_t 2
-        var #r
-          type-con $ctype 0
-        var #d
-          type-con $ctype 0
-      attribute qid 1
-        string-attr VccPrelu.153:15
-      attribute uniqueId 1
-        string-attr 7
-      attribute bvZ3Native 1
-        string-attr False
-    =
-    fun $map_range 1
-    fun $map_t 2
-    var #r
-      type-con $ctype 0
-    var #d
-      type-con $ctype 0
-    var #r
-      type-con $ctype 0
-axiom 0
-    forall 1 1 3
-      var #n
-        type-con $ctype 0
-      pat 1
-        fun $ptr_to 1
-        var #n
-          type-con $ctype 0
-      attribute qid 1
-        string-attr VccPrelu.158:15
-      attribute uniqueId 1
-        string-attr 8
-      attribute bvZ3Native 1
-        string-attr False
-    =
-    fun $ptr_level 1
-    fun $ptr_to 1
-    var #n
-      type-con $ctype 0
-    +
-    fun $ptr_level 1
-    var #n
-      type-con $ctype 0
-    int-num 17
-axiom 0
-    forall 2 1 3
-      var #r
-        type-con $ctype 0
-      var #d
-        type-con $ctype 0
-      pat 1
-        fun $map_t 2
-        var #r
-          type-con $ctype 0
-        var #d
-          type-con $ctype 0
-      attribute qid 1
-        string-attr VccPrelu.159:15
-      attribute uniqueId 1
-        string-attr 9
-      attribute bvZ3Native 1
-        string-attr False
-    =
-    fun $ptr_level 1
-    fun $map_t 2
-    var #r
-      type-con $ctype 0
-    var #d
-      type-con $ctype 0
-    +
-    fun $ptr_level 1
-    var #r
-      type-con $ctype 0
-    int-num 23
-axiom 0
-    forall 1 1 4
-      var t
-        type-con $ctype 0
-      pat 1
-        fun $is_primitive 1
-        var t
-          type-con $ctype 0
-      attribute qid 1
-        string-attr VccPrelu.167:36
-      attribute uniqueId 1
-        string-attr 10
-      attribute bvZ3Native 1
-        string-attr False
-      attribute weight 1
-        expr-attr
-          int-num 0
-    =
-    fun $is_primitive 1
-    var t
-      type-con $ctype 0
-    =
-    fun $kind_of 1
-    var t
-      type-con $ctype 0
-    fun $kind_primitive 0
-axiom 0
-    forall 1 1 4
-      var t
-        type-con $ctype 0
-      pat 1
-        fun $is_composite 1
-        var t
-          type-con $ctype 0
-      attribute qid 1
-        string-attr VccPrelu.173:36
-      attribute uniqueId 1
-        string-attr 11
-      attribute bvZ3Native 1
-        string-attr False
-      attribute weight 1
-        expr-attr
-          int-num 0
-    =
-    fun $is_composite 1
-    var t
-      type-con $ctype 0
-    =
-    fun $kind_of 1
-    var t
-      type-con $ctype 0
-    fun $kind_composite 0
-axiom 0
-    forall 1 1 4
-      var t
-        type-con $ctype 0
-      pat 1
-        fun $is_arraytype 1
-        var t
-          type-con $ctype 0
-      attribute qid 1
-        string-attr VccPrelu.179:36
-      attribute uniqueId 1
-        string-attr 12
-      attribute bvZ3Native 1
-        string-attr False
-      attribute weight 1
-        expr-attr
-          int-num 0
-    =
-    fun $is_arraytype 1
-    var t
-      type-con $ctype 0
-    =
-    fun $kind_of 1
-    var t
-      type-con $ctype 0
-    fun $kind_array 0
-axiom 0
-    forall 1 1 4
-      var t
-        type-con $ctype 0
-      pat 1
-        fun $is_threadtype 1
-        var t
-          type-con $ctype 0
-      attribute qid 1
-        string-attr VccPrelu.185:37
-      attribute uniqueId 1
-        string-attr 13
-      attribute bvZ3Native 1
-        string-attr False
-      attribute weight 1
-        expr-attr
-          int-num 0
-    =
-    fun $is_threadtype 1
-    var t
-      type-con $ctype 0
-    =
-    fun $kind_of 1
-    var t
-      type-con $ctype 0
-    fun $kind_thread 0
-axiom 0
-    forall 1 1 4
-      var t
-        type-con $ctype 0
-      pat 1
-        fun $is_composite 1
-        var t
-          type-con $ctype 0
-      attribute qid 1
-        string-attr VccPrelu.198:15
-      attribute uniqueId 1
-        string-attr 14
-      attribute bvZ3Native 1
-        string-attr False
-      attribute weight 1
-        expr-attr
-          int-num 0
-    implies
-    fun $is_composite 1
-    var t
-      type-con $ctype 0
-    fun $is_non_primitive 1
-    var t
-      type-con $ctype 0
-axiom 0
-    forall 1 1 4
-      var t
-        type-con $ctype 0
-      pat 1
-        fun $is_arraytype 1
-        var t
-          type-con $ctype 0
-      attribute qid 1
-        string-attr VccPrelu.199:15
-      attribute uniqueId 1
-        string-attr 15
-      attribute bvZ3Native 1
-        string-attr False
-      attribute weight 1
-        expr-attr
-          int-num 0
-    implies
-    fun $is_arraytype 1
-    var t
-      type-con $ctype 0
-    fun $is_non_primitive 1
-    var t
-      type-con $ctype 0
-axiom 0
-    forall 1 1 4
-      var t
-        type-con $ctype 0
-      pat 1
-        fun $is_threadtype 1
-        var t
-          type-con $ctype 0
-      attribute qid 1
-        string-attr VccPrelu.200:15
-      attribute uniqueId 1
-        string-attr 16
-      attribute bvZ3Native 1
-        string-attr False
-      attribute weight 1
-        expr-attr
-          int-num 0
-    implies
-    fun $is_threadtype 1
-    var t
-      type-con $ctype 0
-    fun $is_non_primitive 1
-    var t
-      type-con $ctype 0
-axiom 0
-    forall 2 1 3
-      var #r
-        type-con $ctype 0
-      var #d
-        type-con $ctype 0
-      pat 1
-        fun $map_t 2
-        var #r
-          type-con $ctype 0
-        var #d
-          type-con $ctype 0
-      attribute qid 1
-        string-attr VccPrelu.208:15
-      attribute uniqueId 1
-        string-attr 17
-      attribute bvZ3Native 1
-        string-attr False
-    fun $is_primitive 1
-    fun $map_t 2
-    var #r
-      type-con $ctype 0
-    var #d
-      type-con $ctype 0
-axiom 0
-    forall 1 1 3
-      var #n
-        type-con $ctype 0
-      pat 1
-        fun $ptr_to 1
-        var #n
-          type-con $ctype 0
-      attribute qid 1
-        string-attr VccPrelu.209:15
-      attribute uniqueId 1
-        string-attr 18
-      attribute bvZ3Native 1
-        string-attr False
-    fun $is_primitive 1
-    fun $ptr_to 1
-    var #n
-      type-con $ctype 0
-axiom 0
-    forall 1 1 3
-      var #n
-        type-con $ctype 0
-      pat 1
-        fun $is_primitive 1
-        var #n
-          type-con $ctype 0
-      attribute qid 1
-        string-attr VccPrelu.210:15
-      attribute uniqueId 1
-        string-attr 19
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    fun $is_primitive 1
-    var #n
-      type-con $ctype 0
-    not
-    fun $is_claimable 1
-    var #n
-      type-con $ctype 0
-axiom 0
-    fun $is_primitive 1
-    fun ^^void 0
-axiom 0
-    fun $is_primitive 1
-    fun ^^bool 0
-axiom 0
-    fun $is_primitive 1
-    fun ^^mathint 0
-axiom 0
-    fun $is_primitive 1
-    fun ^$#ptrset 0
-axiom 0
-    fun $is_primitive 1
-    fun ^$#state_t 0
-axiom 0
-    fun $is_threadtype 1
-    fun ^$#thread_id_t 0
-axiom 0
-    fun $is_primitive 1
-    fun ^^i1 0
-axiom 0
-    fun $is_primitive 1
-    fun ^^i2 0
-axiom 0
-    fun $is_primitive 1
-    fun ^^i4 0
-axiom 0
-    fun $is_primitive 1
-    fun ^^i8 0
-axiom 0
-    fun $is_primitive 1
-    fun ^^u1 0
-axiom 0
-    fun $is_primitive 1
-    fun ^^u2 0
-axiom 0
-    fun $is_primitive 1
-    fun ^^u4 0
-axiom 0
-    fun $is_primitive 1
-    fun ^^u8 0
-axiom 0
-    fun $is_primitive 1
-    fun ^^f4 0
-axiom 0
-    fun $is_primitive 1
-    fun ^^f8 0
-axiom 0
-    =
-    fun $me 0
-    fun $ptr 2
-    fun ^$#thread_id_t 0
-    fun $me_ref 0
-axiom 0
-    forall 3 0 4
-      var M
-        type-con $memory_t 0
-      var p
-        type-con $ptr 0
-      var v
-        int
-      attribute qid 1
-        string-attr VccPrelu.238:15
-      attribute uniqueId 1
-        string-attr 20
-      attribute bvZ3Native 1
-        string-attr False
-      attribute weight 1
-        expr-attr
-          int-num 0
-    =
-    fun $select.mem 2
-    fun $store.mem 3
-    var M
-      type-con $memory_t 0
-    var p
-      type-con $ptr 0
-    var v
-      int
-    var p
-      type-con $ptr 0
-    var v
-      int
-axiom 0
-    forall 4 0 4
-      var M
-        type-con $memory_t 0
-      var p
-        type-con $ptr 0
-      var q
-        type-con $ptr 0
-      var v
-        int
-      attribute qid 1
-        string-attr VccPrelu.240:15
-      attribute uniqueId 1
-        string-attr 21
-      attribute bvZ3Native 1
-        string-attr False
-      attribute weight 1
-        expr-attr
-          int-num 0
-    or 2
-    =
-    var p
-      type-con $ptr 0
-    var q
-      type-con $ptr 0
-    =
-    fun $select.mem 2
-    fun $store.mem 3
-    var M
-      type-con $memory_t 0
-    var p
-      type-con $ptr 0
-    var v
-      int
-    var q
-      type-con $ptr 0
-    fun $select.mem 2
-    var M
-      type-con $memory_t 0
-    var q
-      type-con $ptr 0
-axiom 0
-    forall 3 0 4
-      var M
-        type-con $typemap_t 0
-      var p
-        type-con $ptr 0
-      var v
-        type-con $type_state 0
-      attribute qid 1
-        string-attr VccPrelu.249:15
-      attribute uniqueId 1
-        string-attr 22
-      attribute bvZ3Native 1
-        string-attr False
-      attribute weight 1
-        expr-attr
-          int-num 0
-    =
-    fun $select.tm 2
-    fun $store.tm 3
-    var M
-      type-con $typemap_t 0
-    var p
-      type-con $ptr 0
-    var v
-      type-con $type_state 0
-    var p
-      type-con $ptr 0
-    var v
-      type-con $type_state 0
-axiom 0
-    forall 4 0 4
-      var M
-        type-con $typemap_t 0
-      var p
-        type-con $ptr 0
-      var q
-        type-con $ptr 0
-      var v
-        type-con $type_state 0
-      attribute qid 1
-        string-attr VccPrelu.251:15
-      attribute uniqueId 1
-        string-attr 23
-      attribute bvZ3Native 1
-        string-attr False
-      attribute weight 1
-        expr-attr
-          int-num 0
-    or 2
-    =
-    var p
-      type-con $ptr 0
-    var q
-      type-con $ptr 0
-    =
-    fun $select.tm 2
-    fun $store.tm 3
-    var M
-      type-con $typemap_t 0
-    var p
-      type-con $ptr 0
-    var v
-      type-con $type_state 0
-    var q
-      type-con $ptr 0
-    fun $select.tm 2
-    var M
-      type-con $typemap_t 0
-    var q
-      type-con $ptr 0
-axiom 0
-    forall 3 0 4
-      var M
-        type-con $statusmap_t 0
-      var p
-        type-con $ptr 0
-      var v
-        type-con $status 0
-      attribute qid 1
-        string-attr VccPrelu.260:15
-      attribute uniqueId 1
-        string-attr 24
-      attribute bvZ3Native 1
-        string-attr False
-      attribute weight 1
-        expr-attr
-          int-num 0
-    =
-    fun $select.sm 2
-    fun $store.sm 3
-    var M
-      type-con $statusmap_t 0
-    var p
-      type-con $ptr 0
-    var v
-      type-con $status 0
-    var p
-      type-con $ptr 0
-    var v
-      type-con $status 0
-axiom 0
-    forall 4 0 4
-      var M
-        type-con $statusmap_t 0
-      var p
-        type-con $ptr 0
-      var q
-        type-con $ptr 0
-      var v
-        type-con $status 0
-      attribute qid 1
-        string-attr VccPrelu.262:15
-      attribute uniqueId 1
-        string-attr 25
-      attribute bvZ3Native 1
-        string-attr False
-      attribute weight 1
-        expr-attr
-          int-num 0
-    or 2
-    =
-    var p
-      type-con $ptr 0
-    var q
-      type-con $ptr 0
-    =
-    fun $select.sm 2
-    fun $store.sm 3
-    var M
-      type-con $statusmap_t 0
-    var p
-      type-con $ptr 0
-    var v
-      type-con $status 0
-    var q
-      type-con $ptr 0
-    fun $select.sm 2
-    var M
-      type-con $statusmap_t 0
-    var q
-      type-con $ptr 0
-axiom 0
-    forall 3 1 3
-      var p
-        type-con $ptr 0
-      var q
-        type-con $ptr 0
-      var r
-        type-con $ptr 0
-      pat 2
-        fun $extent_hint 2
-        var p
-          type-con $ptr 0
-        var q
-          type-con $ptr 0
-        fun $extent_hint 2
-        var q
-          type-con $ptr 0
-        var r
-          type-con $ptr 0
-      attribute qid 1
-        string-attr VccPrelu.288:15
-      attribute uniqueId 1
-        string-attr 26
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 2
-    fun $extent_hint 2
-    var p
-      type-con $ptr 0
-    var q
-      type-con $ptr 0
-    fun $extent_hint 2
-    var q
-      type-con $ptr 0
-    var r
-      type-con $ptr 0
-    fun $extent_hint 2
-    var p
-      type-con $ptr 0
-    var r
-      type-con $ptr 0
-axiom 0
-    forall 1 1 3
-      var p
-        type-con $ptr 0
-      pat 1
-        fun $typ 1
-        var p
-          type-con $ptr 0
-      attribute qid 1
-        string-attr VccPrelu.290:15
-      attribute uniqueId 1
-        string-attr 27
-      attribute bvZ3Native 1
-        string-attr False
-    fun $extent_hint 2
-    var p
-      type-con $ptr 0
-    var p
-      type-con $ptr 0
-axiom 0
-    forall 4 1 3
-      var t
-        type-con $ctype 0
-      var s
-        type-con $ctype 0
-      var min
-        int
-      var max
-        int
-      pat 1
-        fun $is_nested_range 4
-        var t
-          type-con $ctype 0
-        var s
-          type-con $ctype 0
-        var min
-          int
-        var max
-          int
-      attribute qid 1
-        string-attr VccPrelu.297:27
-      attribute uniqueId 1
-        string-attr 28
-      attribute bvZ3Native 1
-        string-attr False
-    =
-    fun $is_nested_range 4
-    var t
-      type-con $ctype 0
-    var s
-      type-con $ctype 0
-    var min
-      int
-    var max
-      int
-    and 3
-    fun $is_nested 2
-    var t
-      type-con $ctype 0
-    var s
-      type-con $ctype 0
-    =
-    fun $nesting_min 2
-    var t
-      type-con $ctype 0
-    var s
-      type-con $ctype 0
-    var min
-      int
-    =
-    fun $nesting_max 2
-    var t
-      type-con $ctype 0
-    var s
-      type-con $ctype 0
-    var max
-      int
-axiom 0
-    forall 2 0 4
-      var #t
-        type-con $ctype 0
-      var #b
-        int
-      attribute qid 1
-        string-attr VccPrelu.334:15
-      attribute uniqueId 1
-        string-attr 29
-      attribute bvZ3Native 1
-        string-attr False
-      attribute weight 1
-        expr-attr
-          int-num 0
-    =
-    fun $typ 1
-    fun $ptr 2
-    var #t
-      type-con $ctype 0
-    var #b
-      int
-    var #t
-      type-con $ctype 0
-axiom 0
-    forall 2 0 4
-      var #t
-        type-con $ctype 0
-      var #b
-        int
-      attribute qid 1
-        string-attr VccPrelu.335:15
-      attribute uniqueId 1
-        string-attr 30
-      attribute bvZ3Native 1
-        string-attr False
-      attribute weight 1
-        expr-attr
-          int-num 0
-    =
-    fun $ref 1
-    fun $ptr 2
-    var #t
-      type-con $ctype 0
-    var #b
-      int
-    var #b
-      int
-axiom 0
-    forall 2 1 4
-      var p
-        type-con $ptr 0
-      var f
-        type-con $field 0
-      pat 1
-        fun $ghost_ref 2
-        var p
-          type-con $ptr 0
-        var f
-          type-con $field 0
-      attribute qid 1
-        string-attr VccPrelu.344:15
-      attribute uniqueId 1
-        string-attr 31
-      attribute bvZ3Native 1
-        string-attr False
-      attribute weight 1
-        expr-attr
-          int-num 0
-    and 2
-    =
-    fun $ghost_emb 1
-    fun $ghost_ref 2
-    var p
-      type-con $ptr 0
-    var f
-      type-con $field 0
-    var p
-      type-con $ptr 0
-    =
-    fun $ghost_path 1
-    fun $ghost_ref 2
-    var p
-      type-con $ptr 0
-    var f
-      type-con $field 0
-    var f
-      type-con $field 0
-axiom 0
-    forall 2 1 4
-      var fld
-        type-con $field 0
-      var off
-        int
-      pat 1
-        fun $array_path 2
-        var fld
-          type-con $field 0
-        var off
-          int
-      attribute qid 1
-        string-attr VccPrelu.355:15
-      attribute uniqueId 1
-        string-attr 32
-      attribute bvZ3Native 1
-        string-attr False
-      attribute weight 1
-        expr-attr
-          int-num 0
-    and 3
-    not
-    fun $is_base_field 1
-    fun $array_path 2
-    var fld
-      type-con $field 0
-    var off
-      int
-    =
-    fun $array_path_1 1
-    fun $array_path 2
-    var fld
-      type-con $field 0
-    var off
-      int
-    var fld
-      type-con $field 0
-    =
-    fun $array_path_2 1
-    fun $array_path 2
-    var fld
-      type-con $field 0
-    var off
-      int
-    var off
-      int
-axiom 0
-    =
-    fun $null 0
-    fun $ptr 2
-    fun ^^void 0
-    int-num 0
-axiom 0
-    forall 2 0 4
-      var #p
-        type-con $ptr 0
-      var #t
-        type-con $ctype 0
-      attribute qid 1
-        string-attr VccPrelu.368:15
-      attribute uniqueId 1
-        string-attr 33
-      attribute bvZ3Native 1
-        string-attr False
-      attribute weight 1
-        expr-attr
-          int-num 0
-    =
-    fun $is 2
-    var #p
-      type-con $ptr 0
-    var #t
-      type-con $ctype 0
-    =
-    fun $typ 1
-    var #p
-      type-con $ptr 0
-    var #t
-      type-con $ctype 0
-axiom 0
-    forall 2 1 3
-      var #p
-        type-con $ptr 0
-      var #t
-        type-con $ctype 0
-      pat 1
-        fun $is 2
-        var #p
-          type-con $ptr 0
-        var #t
-          type-con $ctype 0
-      attribute qid 1
-        string-attr VccPrelu.370:15
-      attribute uniqueId 1
-        string-attr 34
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    fun $is 2
-    var #p
-      type-con $ptr 0
-    var #t
-      type-con $ctype 0
-    =
-    var #p
-      type-con $ptr 0
-    fun $ptr 2
-    var #t
-      type-con $ctype 0
-    fun $ref 1
-    var #p
-      type-con $ptr 0
-axiom 0
-    forall 2 1 3
-      var r
-        int
-      var f
-        type-con $field 0
-      pat 1
-        fun $containing_struct 2
-        fun $dot 2
-        fun $ptr 2
-        fun $field_parent_type 1
-        var f
-          type-con $field 0
-        var r
-          int
-        var f
-          type-con $field 0
-        var f
-          type-con $field 0
-      attribute qid 1
-        string-attr VccPrelu.388:15
-      attribute uniqueId 1
-        string-attr 35
-      attribute bvZ3Native 1
-        string-attr False
-    =
-    fun $containing_struct 2
-    fun $dot 2
-    fun $ptr 2
-    fun $field_parent_type 1
-    var f
-      type-con $field 0
-    var r
-      int
-    var f
-      type-con $field 0
-    var f
-      type-con $field 0
-    fun $ptr 2
-    fun $field_parent_type 1
-    var f
-      type-con $field 0
-    var r
-      int
-axiom 0
-    forall 2 1 3
-      var p
-        type-con $ptr 0
-      var f
-        type-con $field 0
-      pat 1
-        fun $containing_struct 2
-        var p
-          type-con $ptr 0
-        var f
-          type-con $field 0
-      attribute qid 1
-        string-attr VccPrelu.392:15
-      attribute uniqueId 1
-        string-attr 36
-      attribute bvZ3Native 1
-        string-attr False
-    =
-    fun $containing_struct 2
-    var p
-      type-con $ptr 0
-    var f
-      type-con $field 0
-    fun $ptr 2
-    fun $field_parent_type 1
-    var f
-      type-con $field 0
-    fun $containing_struct_ref 2
-    var p
-      type-con $ptr 0
-    var f
-      type-con $field 0
-axiom 0
-    forall 2 1 3
-      var p
-        type-con $ptr 0
-      var f
-        type-con $field 0
-      pat 1
-        fun $dot 2
-        fun $containing_struct 2
-        var p
-          type-con $ptr 0
-        var f
-          type-con $field 0
-        var f
-          type-con $field 0
-      attribute qid 1
-        string-attr VccPrelu.396:15
-      attribute uniqueId 1
-        string-attr 37
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    >=
-    fun $field_offset 1
-    var f
-      type-con $field 0
-    int-num 0
-    =
-    fun $ref 1
-    fun $dot 2
-    fun $containing_struct 2
-    var p
-      type-con $ptr 0
-    var f
-      type-con $field 0
-    var f
-      type-con $field 0
-    fun $ref 1
-    var p
-      type-con $ptr 0
-axiom 0
-    forall 1 1 3
-      var ts
-        type-con $type_state 0
-      pat 1
-        fun $ts_emb 1
-        var ts
-          type-con $type_state 0
-      attribute qid 1
-        string-attr VccPrelu.427:15
-      attribute uniqueId 1
-        string-attr 38
-      attribute bvZ3Native 1
-        string-attr False
-    and 2
-    not
-    =
-    fun $kind_of 1
-    fun $typ 1
-    fun $ts_emb 1
-    var ts
-      type-con $type_state 0
-    fun $kind_primitive 0
-    fun $is_non_primitive 1
-    fun $typ 1
-    fun $ts_emb 1
-    var ts
-      type-con $type_state 0
-axiom 0
-    forall 2 1 3
-      var S
-        type-con $state 0
-      var p
-        type-con $ptr 0
-      pat 2
-        fun $typed 2
-        var S
-          type-con $state 0
-        var p
-          type-con $ptr 0
-        fun $select.tm 2
-        fun $typemap 1
-        var S
-          type-con $state 0
-        fun $ts_emb 1
-        fun $select.tm 2
-        fun $typemap 1
-        var S
-          type-con $state 0
-        var p
-          type-con $ptr 0
-      attribute qid 1
-        string-attr VccPrelu.430:15
-      attribute uniqueId 1
-        string-attr 39
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    fun $typed 2
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    fun $typed 2
-    var S
-      type-con $state 0
-    fun $ts_emb 1
-    fun $select.tm 2
-    fun $typemap 1
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-axiom 0
-    forall 2 1 3
-      var S
-        type-con $state 0
-      var p
-        type-con $ptr 0
-      pat 1
-        fun $ts_is_volatile 1
-        fun $select.tm 2
-        fun $typemap 1
-        var S
-          type-con $state 0
-        var p
-          type-con $ptr 0
-      attribute qid 1
-        string-attr VccPrelu.440:15
-      attribute uniqueId 1
-        string-attr 40
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 2
-    fun $good_state 1
-    var S
-      type-con $state 0
-    fun $ts_is_volatile 1
-    fun $select.tm 2
-    fun $typemap 1
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    =
-    fun $kind_of 1
-    fun $typ 1
-    var p
-      type-con $ptr 0
-    fun $kind_primitive 0
-axiom 0
-    forall 2 1 4
-      var S
-        type-con $state 0
-      var p
-        type-con $ptr 0
-      pat 1
-        fun $select.sm 2
-        fun $statusmap 1
-        var S
-          type-con $state 0
-        var p
-          type-con $ptr 0
-      attribute qid 1
-        string-attr VccPrelu.456:15
-      attribute uniqueId 1
-        string-attr 41
-      attribute bvZ3Native 1
-        string-attr False
-      attribute weight 1
-        expr-attr
-          int-num 0
-    or 2
-    <=
-    fun $timestamp 2
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    fun $current_timestamp 1
-    var S
-      type-con $state 0
-    not
-    fun $ts_typed 1
-    fun $select.tm 2
-    fun $typemap 1
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-axiom 0
-    fun $good_state 1
-    fun $vs_state 1
-    fun $struct_zero 0
-axiom 0
-    forall 1 0 3
-      var s
-        type-con $struct 0
-      attribute qid 1
-        string-attr VccPrelu.486:15
-      attribute uniqueId 1
-        string-attr 42
-      attribute bvZ3Native 1
-        string-attr False
-    fun $good_state 1
-    fun $vs_state 1
-    var s
-      type-con $struct 0
-axiom 0
-    forall 2 1 3
-      var S
-        type-con $state 0
-      var p
-        type-con $ptr 0
-      pat 1
-        fun $vs_ctor 2
-        var S
-          type-con $state 0
-        var p
-          type-con $ptr 0
-      attribute qid 1
-        string-attr VccPrelu.489:15
-      attribute uniqueId 1
-        string-attr 43
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    fun $good_state 1
-    var S
-      type-con $state 0
-    and 2
-    =
-    fun $vs_base_ref 1
-    fun $vs_ctor 2
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    fun $ref 1
-    var p
-      type-con $ptr 0
-    =
-    fun $vs_state 1
-    fun $vs_ctor 2
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    var S
-      type-con $state 0
-axiom 0
-    forall 6 1 3
-      var r
-        type-con $record 0
-      var f
-        type-con $field 0
-      var val_bitsize
-        int
-      var from
-        int
-      var to
-        int
-      var repl
-        int
-      pat 1
-        fun $rec_update_bv 6
-        var r
-          type-con $record 0
-        var f
-          type-con $field 0
-        var val_bitsize
-          int
-        var from
-          int
-        var to
-          int
-        var repl
-          int
-      attribute qid 1
-        string-attr VccPrelu.502:25
-      attribute uniqueId 1
-        string-attr 44
-      attribute bvZ3Native 1
-        string-attr False
-    =
-    fun $rec_update_bv 6
-    var r
-      type-con $record 0
-    var f
-      type-con $field 0
-    var val_bitsize
-      int
-    var from
-      int
-    var to
-      int
-    var repl
-      int
-    fun $rec_update 3
-    var r
-      type-con $record 0
-    var f
-      type-con $field 0
-    fun $bv_update 5
-    fun $rec_fetch 2
-    var r
-      type-con $record 0
-    var f
-      type-con $field 0
-    var val_bitsize
-      int
-    var from
-      int
-    var to
-      int
-    var repl
-      int
-axiom 0
-    forall 1 0 3
-      var f
-        type-con $field 0
-      attribute qid 1
-        string-attr VccPrelu.505:15
-      attribute uniqueId 1
-        string-attr 45
-      attribute bvZ3Native 1
-        string-attr False
-    =
-    fun $rec_fetch 2
-    fun $rec_zero 0
-    var f
-      type-con $field 0
-    int-num 0
-axiom 0
-    forall 3 1 3
-      var r
-        type-con $record 0
-      var f
-        type-con $field 0
-      var v
-        int
-      pat 1
-        fun $rec_fetch 2
-        fun $rec_update 3
-        var r
-          type-con $record 0
-        var f
-          type-con $field 0
-        var v
-          int
-        var f
-          type-con $field 0
-      attribute qid 1
-        string-attr VccPrelu.507:15
-      attribute uniqueId 1
-        string-attr 46
-      attribute bvZ3Native 1
-        string-attr False
-    =
-    fun $rec_fetch 2
-    fun $rec_update 3
-    var r
-      type-con $record 0
-    var f
-      type-con $field 0
-    var v
-      int
-    var f
-      type-con $field 0
-    var v
-      int
-axiom 0
-    forall 4 1 3
-      var r
-        type-con $record 0
-      var f1
-        type-con $field 0
-      var f2
-        type-con $field 0
-      var v
-        int
-      pat 1
-        fun $rec_fetch 2
-        fun $rec_update 3
-        var r
-          type-con $record 0
-        var f1
-          type-con $field 0
-        var v
-          int
-        var f2
-          type-con $field 0
-      attribute qid 1
-        string-attr VccPrelu.510:15
-      attribute uniqueId 1
-        string-attr 47
-      attribute bvZ3Native 1
-        string-attr False
-    or 2
-    =
-    fun $rec_fetch 2
-    fun $rec_update 3
-    var r
-      type-con $record 0
-    var f1
-      type-con $field 0
-    var v
-      int
-    var f2
-      type-con $field 0
-    fun $rec_fetch 2
-    var r
-      type-con $record 0
-    var f2
-      type-con $field 0
-    =
-    var f1
-      type-con $field 0
-    var f2
-      type-con $field 0
-axiom 0
-    forall 1 1 3
-      var t
-        type-con $ctype 0
-      pat 1
-        fun $is_record_type 1
-        var t
-          type-con $ctype 0
-      attribute qid 1
-        string-attr VccPrelu.516:15
-      attribute uniqueId 1
-        string-attr 48
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    fun $is_record_type 1
-    var t
-      type-con $ctype 0
-    fun $is_primitive 1
-    var t
-      type-con $ctype 0
-axiom 0
-    forall 3 1 3
-      var p
-        type-con $ctype 0
-      var f
-        type-con $field 0
-      var ft
-        type-con $ctype 0
-      pat 2
-        fun $is_record_field 3
-        var p
-          type-con $ctype 0
-        var f
-          type-con $field 0
-        var ft
-          type-con $ctype 0
-        fun $is_record_type 1
-        var ft
-          type-con $ctype 0
-      attribute qid 1
-        string-attr VccPrelu.519:15
-      attribute uniqueId 1
-        string-attr 49
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 2
-    fun $is_record_field 3
-    var p
-      type-con $ctype 0
-    var f
-      type-con $field 0
-    var ft
-      type-con $ctype 0
-    fun $is_record_type 1
-    var ft
-      type-con $ctype 0
-    =
-    fun $as_record_record_field 1
-    var f
-      type-con $field 0
-    var f
-      type-con $field 0
-axiom 0
-    forall 2 1 3
-      var r1
-        type-con $record 0
-      var r2
-        type-con $record 0
-      pat 1
-        fun $rec_eq 2
-        var r1
-          type-con $record 0
-        var r2
-          type-con $record 0
-      attribute qid 1
-        string-attr VccPrelu.522:18
-      attribute uniqueId 1
-        string-attr 50
-      attribute bvZ3Native 1
-        string-attr False
-    =
-    fun $rec_eq 2
-    var r1
-      type-con $record 0
-    var r2
-      type-con $record 0
-    =
-    var r1
-      type-con $record 0
-    var r2
-      type-con $record 0
-axiom 0
-    forall 2 1 3
-      var x
-        int
-      var y
-        int
-      pat 1
-        fun $rec_base_eq 2
-        var x
-          int
-        var y
-          int
-      attribute qid 1
-        string-attr VccPrelu.524:23
-      attribute uniqueId 1
-        string-attr 51
-      attribute bvZ3Native 1
-        string-attr False
-    =
-    fun $rec_base_eq 2
-    var x
-      int
-    var y
-      int
-    =
-    var x
-      int
-    var y
-      int
-axiom 0
-    forall 1 0 3
-      var r
-        type-con $record 0
-      attribute qid 1
-        string-attr VccPrelu.530:15
-      attribute uniqueId 1
-        string-attr 52
-      attribute bvZ3Native 1
-        string-attr False
-    =
-    fun $int_to_record 1
-    fun $record_to_int 1
-    var r
-      type-con $record 0
-    var r
-      type-con $record 0
-axiom 0
-    forall 2 1 3
-      var r1
-        type-con $record 0
-      var r2
-        type-con $record 0
-      pat 1
-        fun $rec_eq 2
-        var r1
-          type-con $record 0
-        var r2
-          type-con $record 0
-      attribute qid 1
-        string-attr VccPrelu.532:15
-      attribute uniqueId 1
-        string-attr 54
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    forall 1 0 3
-      var f
-        type-con $field 0
-      attribute qid 1
-        string-attr VccPrelu.534:11
-      attribute uniqueId 1
-        string-attr 53
-      attribute bvZ3Native 1
-        string-attr False
-    fun $rec_base_eq 2
-    fun $rec_fetch 2
-    var r1
-      type-con $record 0
-    var f
-      type-con $field 0
-    fun $rec_fetch 2
-    var r2
-      type-con $record 0
-    var f
-      type-con $field 0
-    fun $rec_eq 2
-    var r1
-      type-con $record 0
-    var r2
-      type-con $record 0
-axiom 0
-    forall 3 1 3
-      var r1
-        type-con $record 0
-      var r2
-        type-con $record 0
-      var f
-        type-con $field 0
-      pat 1
-        fun $rec_base_eq 2
-        fun $rec_fetch 2
-        var r1
-          type-con $record 0
-        var f
-          type-con $field 0
-        fun $rec_fetch 2
-        var r2
-          type-con $record 0
-        fun $as_record_record_field 1
-        var f
-          type-con $field 0
-      attribute qid 1
-        string-attr VccPrelu.536:15
-      attribute uniqueId 1
-        string-attr 55
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    fun $rec_eq 2
-    fun $int_to_record 1
-    fun $rec_fetch 2
-    var r1
-      type-con $record 0
-    var f
-      type-con $field 0
-    fun $int_to_record 1
-    fun $rec_fetch 2
-    var r2
-      type-con $record 0
-    var f
-      type-con $field 0
-    fun $rec_base_eq 2
-    fun $rec_fetch 2
-    var r1
-      type-con $record 0
-    var f
-      type-con $field 0
-    fun $rec_fetch 2
-    var r2
-      type-con $record 0
-    var f
-      type-con $field 0
-axiom 0
-    fun $has_volatile_owns_set 1
-    fun ^^claim 0
-axiom 0
-    forall 2 1 3
-      var #p
-        type-con $ptr 0
-      var t
-        type-con $ctype 0
-      pat 1
-        fun $dot 2
-        var #p
-          type-con $ptr 0
-        fun $owns_set_field 1
-        var t
-          type-con $ctype 0
-      attribute qid 1
-        string-attr VccPrelu.555:15
-      attribute uniqueId 1
-        string-attr 56
-      attribute bvZ3Native 1
-        string-attr False
-    =
-    fun $dot 2
-    var #p
-      type-con $ptr 0
-    fun $owns_set_field 1
-    var t
-      type-con $ctype 0
-    fun $ptr 2
-    fun ^$#ptrset 0
-    fun $ghost_ref 2
-    var #p
-      type-con $ptr 0
-    fun $owns_set_field 1
-    var t
-      type-con $ctype 0
-axiom 0
-    forall 2 1 4
-      var S
-        type-con $state 0
-      var p
-        type-con $ptr 0
-      pat 2
-        fun $is_primitive 1
-        fun $typ 1
-        var p
-          type-con $ptr 0
-        fun $owner 2
-        var S
-          type-con $state 0
-        var p
-          type-con $ptr 0
-      attribute qid 1
-        string-attr VccPrelu.567:15
-      attribute uniqueId 1
-        string-attr 57
-      attribute bvZ3Native 1
-        string-attr False
-      attribute weight 1
-        expr-attr
-          int-num 0
-    implies
-    fun $is_primitive 1
-    fun $typ 1
-    var p
-      type-con $ptr 0
-    =
-    fun $owner 2
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    fun $owner 2
-    var S
-      type-con $state 0
-    fun $ts_emb 1
-    fun $select.tm 2
-    fun $typemap 1
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-axiom 0
-    forall 2 1 4
-      var S
-        type-con $state 0
-      var p
-        type-con $ptr 0
-      pat 2
-        fun $is_non_primitive 1
-        fun $typ 1
-        var p
-          type-con $ptr 0
-        fun $owner 2
-        var S
-          type-con $state 0
-        var p
-          type-con $ptr 0
-      attribute qid 1
-        string-attr VccPrelu.569:15
-      attribute uniqueId 1
-        string-attr 58
-      attribute bvZ3Native 1
-        string-attr False
-      attribute weight 1
-        expr-attr
-          int-num 0
-    implies
-    fun $is_non_primitive 1
-    fun $typ 1
-    var p
-      type-con $ptr 0
-    =
-    fun $owner 2
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    fun $st_owner 1
-    fun $select.sm 2
-    fun $statusmap 1
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-axiom 0
-    forall 2 1 4
-      var S
-        type-con $state 0
-      var p
-        type-con $ptr 0
-      pat 2
-        fun $is_primitive 1
-        fun $typ 1
-        var p
-          type-con $ptr 0
-        fun $closed 2
-        var S
-          type-con $state 0
-        var p
-          type-con $ptr 0
-      attribute qid 1
-        string-attr VccPrelu.572:15
-      attribute uniqueId 1
-        string-attr 59
-      attribute bvZ3Native 1
-        string-attr False
-      attribute weight 1
-        expr-attr
-          int-num 0
-    implies
-    fun $is_primitive 1
-    fun $typ 1
-    var p
-      type-con $ptr 0
-    =
-    fun $closed 2
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    fun $st_closed 1
-    fun $select.sm 2
-    fun $statusmap 1
-    var S
-      type-con $state 0
-    fun $ts_emb 1
-    fun $select.tm 2
-    fun $typemap 1
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-axiom 0
-    forall 2 1 4
-      var S
-        type-con $state 0
-      var p
-        type-con $ptr 0
-      pat 2
-        fun $is_non_primitive 1
-        fun $typ 1
-        var p
-          type-con $ptr 0
-        fun $closed 2
-        var S
-          type-con $state 0
-        var p
-          type-con $ptr 0
-      attribute qid 1
-        string-attr VccPrelu.574:15
-      attribute uniqueId 1
-        string-attr 60
-      attribute bvZ3Native 1
-        string-attr False
-      attribute weight 1
-        expr-attr
-          int-num 0
-    implies
-    fun $is_non_primitive 1
-    fun $typ 1
-    var p
-      type-con $ptr 0
-    =
-    fun $closed 2
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    fun $st_closed 1
-    fun $select.sm 2
-    fun $statusmap 1
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-axiom 0
-    forall 2 1 4
-      var S
-        type-con $state 0
-      var p
-        type-con $ptr 0
-      pat 2
-        fun $is_primitive 1
-        fun $typ 1
-        var p
-          type-con $ptr 0
-        fun $timestamp 2
-        var S
-          type-con $state 0
-        var p
-          type-con $ptr 0
-      attribute qid 1
-        string-attr VccPrelu.577:15
-      attribute uniqueId 1
-        string-attr 61
-      attribute bvZ3Native 1
-        string-attr False
-      attribute weight 1
-        expr-attr
-          int-num 0
-    implies
-    fun $is_primitive 1
-    fun $typ 1
-    var p
-      type-con $ptr 0
-    =
-    fun $timestamp 2
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    fun $st_timestamp 1
-    fun $select.sm 2
-    fun $statusmap 1
-    var S
-      type-con $state 0
-    fun $ts_emb 1
-    fun $select.tm 2
-    fun $typemap 1
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-axiom 0
-    forall 2 1 4
-      var S
-        type-con $state 0
-      var p
-        type-con $ptr 0
-      pat 2
-        fun $is_non_primitive 1
-        fun $typ 1
-        var p
-          type-con $ptr 0
-        fun $timestamp 2
-        var S
-          type-con $state 0
-        var p
-          type-con $ptr 0
-      attribute qid 1
-        string-attr VccPrelu.579:15
-      attribute uniqueId 1
-        string-attr 62
-      attribute bvZ3Native 1
-        string-attr False
-      attribute weight 1
-        expr-attr
-          int-num 0
-    implies
-    fun $is_non_primitive 1
-    fun $typ 1
-    var p
-      type-con $ptr 0
-    =
-    fun $timestamp 2
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    fun $st_timestamp 1
-    fun $select.sm 2
-    fun $statusmap 1
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-axiom 0
-    fun $position_marker 0
-axiom 0
-    forall 1 1 3
-      var s
-        type-con $status 0
-      pat 1
-        fun $st_owner 1
-        var s
-          type-con $status 0
-      attribute qid 1
-        string-attr VccPrelu.585:15
-      attribute uniqueId 1
-        string-attr 63
-      attribute bvZ3Native 1
-        string-attr False
-    and 2
-    not
-    =
-    fun $kind_of 1
-    fun $typ 1
-    fun $st_owner 1
-    var s
-      type-con $status 0
-    fun $kind_primitive 0
-    fun $is_non_primitive 1
-    fun $typ 1
-    fun $st_owner 1
-    var s
-      type-con $status 0
-axiom 0
-    forall 2 1 4
-      var S
-        type-con $state 0
-      var #p
-        type-con $ptr 0
-      pat 1
-        fun $owns 2
-        var S
-          type-con $state 0
-        var #p
-          type-con $ptr 0
-      attribute qid 1
-        string-attr VccPrelu.593:28
-      attribute uniqueId 1
-        string-attr 64
-      attribute bvZ3Native 1
-        string-attr False
-      attribute weight 1
-        expr-attr
-          int-num 0
-    =
-    fun $owns 2
-    var S
-      type-con $state 0
-    var #p
-      type-con $ptr 0
-    fun $int_to_ptrset 1
-    fun $select.mem 2
-    fun $memory 1
-    var S
-      type-con $state 0
-    fun $dot 2
-    var #p
-      type-con $ptr 0
-    fun $owns_set_field 1
-    fun $typ 1
-    var #p
-      type-con $ptr 0
-axiom 0
-    forall 2 1 4
-      var S
-        type-con $state 0
-      var p
-        type-con $ptr 0
-      pat 1
-        fun $mutable 2
-        var S
-          type-con $state 0
-        var p
-          type-con $ptr 0
-      attribute qid 1
-        string-attr VccPrelu.608:31
-      attribute uniqueId 1
-        string-attr 65
-      attribute bvZ3Native 1
-        string-attr False
-      attribute weight 1
-        expr-attr
-          int-num 0
-    =
-    fun $mutable 2
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    and 3
-    fun $typed 2
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    =
-    fun $owner 2
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    fun $me 0
-    not
-    fun $closed 2
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-axiom 0
-    forall 2 1 4
-      var S
-        type-con $state 0
-      var #p
-        type-con $ptr 0
-      pat 1
-        fun $typed 2
-        var S
-          type-con $state 0
-        var #p
-          type-con $ptr 0
-      attribute qid 1
-        string-attr VccPrelu.619:11
-      attribute uniqueId 1
-        string-attr 66
-      attribute bvZ3Native 1
-        string-attr False
-      attribute weight 1
-        expr-attr
-          int-num 0
-    implies
-    fun $good_state 1
-    var S
-      type-con $state 0
-    =
-    fun $typed 2
-    var S
-      type-con $state 0
-    var #p
-      type-con $ptr 0
-    fun $ts_typed 1
-    fun $select.tm 2
-    fun $typemap 1
-    var S
-      type-con $state 0
-    var #p
-      type-con $ptr 0
-axiom 0
-    forall 2 1 3
-      var S
-        type-con $state 0
-      var #p
-        type-con $ptr 0
-      pat 1
-        fun $typed 2
-        var S
-          type-con $state 0
-        var #p
-          type-con $ptr 0
-      attribute qid 1
-        string-attr VccPrelu.621:11
-      attribute uniqueId 1
-        string-attr 67
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 2
-    fun $good_state 1
-    var S
-      type-con $state 0
-    fun $typed 2
-    var S
-      type-con $state 0
-    var #p
-      type-con $ptr 0
-    >
-    fun $ref 1
-    var #p
-      type-con $ptr 0
-    int-num 0
-axiom 0
-    forall 3 1 3
-      var S1
-        type-con $state 0
-      var S2
-        type-con $state 0
-      var p
-        type-con $ptr 0
-      pat 2
-        fun $select.sm 2
-        fun $statusmap 1
-        var S2
-          type-con $state 0
-        var p
-          type-con $ptr 0
-        fun $call_transition 2
-        var S1
-          type-con $state 0
-        var S2
-          type-con $state 0
-      attribute qid 1
-        string-attr VccPrelu.685:15
-      attribute uniqueId 1
-        string-attr 68
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    fun $call_transition 2
-    var S1
-      type-con $state 0
-    var S2
-      type-con $state 0
-    fun $instantiate_st 1
-    fun $select.sm 2
-    fun $statusmap 1
-    var S1
-      type-con $state 0
-    var p
-      type-con $ptr 0
-axiom 0
-    forall 2 1 3
-      var S
-        type-con $state 0
-      var p
-        type-con $ptr 0
-      pat 1
-        fun $is_domain_root 2
-        var S
-          type-con $state 0
-        var p
-          type-con $ptr 0
-      attribute qid 1
-        string-attr VccPrelu.711:26
-      attribute uniqueId 1
-        string-attr 69
-      attribute bvZ3Native 1
-        string-attr False
-    =
-    fun $is_domain_root 2
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    true
-axiom 0
-    forall 2 1 3
-      var S
-        type-con $state 0
-      var p
-        type-con $ptr 0
-      pat 1
-        fun $in_wrapped_domain 2
-        var S
-          type-con $state 0
-        var p
-          type-con $ptr 0
-      attribute qid 1
-        string-attr VccPrelu.714:29
-      attribute uniqueId 1
-        string-attr 71
-      attribute bvZ3Native 1
-        string-attr False
-    =
-    fun $in_wrapped_domain 2
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    exists 1 1 3
-      var q
-        type-con $ptr 0
-      pat 1
-        fun $set_in2 2
-        var p
-          type-con $ptr 0
-        fun $ver_domain 1
-        fun $read_version 2
-        var S
-          type-con $state 0
-        var q
-          type-con $ptr 0
-      attribute qid 1
-        string-attr VccPrelu.715:13
-      attribute uniqueId 1
-        string-attr 70
-      attribute bvZ3Native 1
-        string-attr False
-    and 8
-    fun $set_in 2
-    var p
-      type-con $ptr 0
-    fun $ver_domain 1
-    fun $read_version 2
-    var S
-      type-con $state 0
-    var q
-      type-con $ptr 0
-    fun $closed 2
-    var S
-      type-con $state 0
-    var q
-      type-con $ptr 0
-    =
-    fun $owner 2
-    var S
-      type-con $state 0
-    var q
-      type-con $ptr 0
-    fun $me 0
-    fun $is 2
-    var q
-      type-con $ptr 0
-    fun $typ 1
-    var q
-      type-con $ptr 0
-    fun $typed 2
-    var S
-      type-con $state 0
-    var q
-      type-con $ptr 0
-    not
-    =
-    fun $kind_of 1
-    fun $typ 1
-    var q
-      type-con $ptr 0
-    fun $kind_primitive 0
-    fun $is_non_primitive 1
-    fun $typ 1
-    var q
-      type-con $ptr 0
-    fun $is_domain_root 2
-    var S
-      type-con $state 0
-    var q
-      type-con $ptr 0
-axiom 0
-    forall 2 1 3
-      var S
-        type-con $state 0
-      var p
-        type-con $ptr 0
-      pat 1
-        fun $thread_local 2
-        var S
-          type-con $state 0
-        var p
-          type-con $ptr 0
-      attribute qid 1
-        string-attr VccPrelu.728:24
-      attribute uniqueId 1
-        string-attr 72
-      attribute bvZ3Native 1
-        string-attr False
-    =
-    fun $thread_local 2
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    and 2
-    fun $typed 2
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    or 2
-    and 4
-    =
-    fun $kind_of 1
-    fun $typ 1
-    var p
-      type-con $ptr 0
-    fun $kind_primitive 0
-    or 2
-    not
-    fun $ts_is_volatile 1
-    fun $select.tm 2
-    fun $typemap 1
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    not
-    fun $closed 2
-    var S
-      type-con $state 0
-    fun $ts_emb 1
-    fun $select.tm 2
-    fun $typemap 1
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    not
-    =
-    fun $kind_of 1
-    fun $typ 1
-    fun $ts_emb 1
-    fun $select.tm 2
-    fun $typemap 1
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    fun $kind_primitive 0
-    or 2
-    =
-    fun $owner 2
-    var S
-      type-con $state 0
-    fun $ts_emb 1
-    fun $select.tm 2
-    fun $typemap 1
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    fun $me 0
-    fun $in_wrapped_domain 2
-    var S
-      type-con $state 0
-    fun $ts_emb 1
-    fun $select.tm 2
-    fun $typemap 1
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    and 2
-    not
-    =
-    fun $kind_of 1
-    fun $typ 1
-    var p
-      type-con $ptr 0
-    fun $kind_primitive 0
-    or 2
-    =
-    fun $owner 2
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    fun $me 0
-    fun $in_wrapped_domain 2
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-axiom 0
-    forall 3 1 3
-      var #s1
-        type-con $state 0
-      var #p
-        type-con $ptr 0
-      var typ
-        type-con $ctype 0
-      pat 1
-        fun $inv2 4
-        var #s1
-          type-con $state 0
-        var #s1
-          type-con $state 0
-        var #p
-          type-con $ptr 0
-        var typ
-          type-con $ctype 0
-      attribute qid 1
-        string-attr VccPrelu.766:15
-      attribute uniqueId 1
-        string-attr 73
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    fun $imply_inv 3
-    var #s1
-      type-con $state 0
-    var #p
-      type-con $ptr 0
-    var typ
-      type-con $ctype 0
-    fun $inv2 4
-    var #s1
-      type-con $state 0
-    var #s1
-      type-con $state 0
-    var #p
-      type-con $ptr 0
-    var typ
-      type-con $ctype 0
-axiom 0
-    forall 4 1 4
-      var #s1
-        type-con $state 0
-      var #s2
-        type-con $state 0
-      var #p
-        type-con $ptr 0
-      var #t
-        type-con $ctype 0
-      pat 1
-        fun $sequential 4
-        var #s1
-          type-con $state 0
-        var #s2
-          type-con $state 0
-        var #p
-          type-con $ptr 0
-        var #t
-          type-con $ctype 0
-      attribute qid 1
-        string-attr VccPrelu.778:34
-      attribute uniqueId 1
-        string-attr 74
-      attribute bvZ3Native 1
-        string-attr False
-      attribute weight 1
-        expr-attr
-          int-num 0
-    =
-    fun $sequential 4
-    var #s1
-      type-con $state 0
-    var #s2
-      type-con $state 0
-    var #p
-      type-con $ptr 0
-    var #t
-      type-con $ctype 0
-    implies
-    and 2
-    fun $closed 2
-    var #s1
-      type-con $state 0
-    var #p
-      type-con $ptr 0
-    fun $closed 2
-    var #s2
-      type-con $state 0
-    var #p
-      type-con $ptr 0
-    fun $spans_the_same 4
-    var #s1
-      type-con $state 0
-    var #s2
-      type-con $state 0
-    var #p
-      type-con $ptr 0
-    var #t
-      type-con $ctype 0
-axiom 0
-    forall 4 1 4
-      var #s1
-        type-con $state 0
-      var #s2
-        type-con $state 0
-      var #dependant
-        type-con $ptr 0
-      var #this
-        type-con $ptr 0
-      pat 1
-        fun $depends 4
-        var #s1
-          type-con $state 0
-        var #s2
-          type-con $state 0
-        var #dependant
-          type-con $ptr 0
-        var #this
-          type-con $ptr 0
-      attribute qid 1
-        string-attr VccPrelu.781:31
-      attribute uniqueId 1
-        string-attr 75
-      attribute bvZ3Native 1
-        string-attr False
-      attribute weight 1
-        expr-attr
-          int-num 0
-    =
-    fun $depends 4
-    var #s1
-      type-con $state 0
-    var #s2
-      type-con $state 0
-    var #dependant
-      type-con $ptr 0
-    var #this
-      type-con $ptr 0
-    or 4
-    fun $spans_the_same 4
-    var #s1
-      type-con $state 0
-    var #s2
-      type-con $state 0
-    var #this
-      type-con $ptr 0
-    fun $typ 1
-    var #this
-      type-con $ptr 0
-    and 2
-    not
-    fun $closed 2
-    var #s1
-      type-con $state 0
-    var #dependant
-      type-con $ptr 0
-    not
-    fun $closed 2
-    var #s2
-      type-con $state 0
-    var #dependant
-      type-con $ptr 0
-    and 2
-    fun $inv2 4
-    var #s1
-      type-con $state 0
-    var #s2
-      type-con $state 0
-    var #dependant
-      type-con $ptr 0
-    fun $typ 1
-    var #dependant
-      type-con $ptr 0
-    fun $nonvolatile_spans_the_same 4
-    var #s1
-      type-con $state 0
-    var #s2
-      type-con $state 0
-    var #dependant
-      type-con $ptr 0
-    fun $typ 1
-    var #dependant
-      type-con $ptr 0
-    fun $is_threadtype 1
-    fun $typ 1
-    var #dependant
-      type-con $ptr 0
-axiom 0
-    forall 4 1 4
-      var #s1
-        type-con $state 0
-      var #s2
-        type-con $state 0
-      var #p
-        type-con $ptr 0
-      var #t
-        type-con $ctype 0
-      pat 1
-        fun $spans_the_same 4
-        var #s1
-          type-con $state 0
-        var #s2
-          type-con $state 0
-        var #p
-          type-con $ptr 0
-        var #t
-          type-con $ctype 0
-      attribute qid 1
-        string-attr VccPrelu.786:38
-      attribute uniqueId 1
-        string-attr 76
-      attribute bvZ3Native 1
-        string-attr False
-      attribute weight 1
-        expr-attr
-          int-num 0
-    =
-    fun $spans_the_same 4
-    var #s1
-      type-con $state 0
-    var #s2
-      type-con $state 0
-    var #p
-      type-con $ptr 0
-    var #t
-      type-con $ctype 0
-    and 4
-    =
-    fun $read_version 2
-    var #s1
-      type-con $state 0
-    var #p
-      type-con $ptr 0
-    fun $read_version 2
-    var #s2
-      type-con $state 0
-    var #p
-      type-con $ptr 0
-    =
-    fun $owns 2
-    var #s1
-      type-con $state 0
-    var #p
-      type-con $ptr 0
-    fun $owns 2
-    var #s2
-      type-con $state 0
-    var #p
-      type-con $ptr 0
-    =
-    fun $select.tm 2
-    fun $typemap 1
-    var #s1
-      type-con $state 0
-    var #p
-      type-con $ptr 0
-    fun $select.tm 2
-    fun $typemap 1
-    var #s2
-      type-con $state 0
-    var #p
-      type-con $ptr 0
-    fun $state_spans_the_same 4
-    var #s1
-      type-con $state 0
-    var #s2
-      type-con $state 0
-    var #p
-      type-con $ptr 0
-    var #t
-      type-con $ctype 0
-axiom 0
-    forall 4 1 4
-      var #s1
-        type-con $state 0
-      var #s2
-        type-con $state 0
-      var #p
-        type-con $ptr 0
-      var #t
-        type-con $ctype 0
-      pat 1
-        fun $nonvolatile_spans_the_same 4
-        var #s1
-          type-con $state 0
-        var #s2
-          type-con $state 0
-        var #p
-          type-con $ptr 0
-        var #t
-          type-con $ctype 0
-      attribute qid 1
-        string-attr VccPrelu.795:50
-      attribute uniqueId 1
-        string-attr 77
-      attribute bvZ3Native 1
-        string-attr False
-      attribute weight 1
-        expr-attr
-          int-num 0
-    =
-    fun $nonvolatile_spans_the_same 4
-    var #s1
-      type-con $state 0
-    var #s2
-      type-con $state 0
-    var #p
-      type-con $ptr 0
-    var #t
-      type-con $ctype 0
-    and 3
-    =
-    fun $read_version 2
-    var #s1
-      type-con $state 0
-    var #p
-      type-con $ptr 0
-    fun $read_version 2
-    var #s2
-      type-con $state 0
-    var #p
-      type-con $ptr 0
-    =
-    fun $select.tm 2
-    fun $typemap 1
-    var #s1
-      type-con $state 0
-    var #p
-      type-con $ptr 0
-    fun $select.tm 2
-    fun $typemap 1
-    var #s2
-      type-con $state 0
-    var #p
-      type-con $ptr 0
-    fun $state_nonvolatile_spans_the_same 4
-    var #s1
-      type-con $state 0
-    var #s2
-      type-con $state 0
-    var #p
-      type-con $ptr 0
-    var #t
-      type-con $ctype 0
-axiom 0
-    forall 1 1 3
-      var T
-        type-con $ctype 0
-      pat 1
-        fun $is_primitive 1
-        var T
-          type-con $ctype 0
-      attribute qid 1
-        string-attr VccPrelu.813:15
-      attribute uniqueId 1
-        string-attr 79
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    fun $is_primitive 1
-    var T
-      type-con $ctype 0
-    forall 2 1 3
-      var r
-        int
-      var p
-        type-con $ptr 0
-      pat 1
-        fun $set_in 2
-        var p
-          type-con $ptr 0
-        fun $full_extent 1
-        fun $ptr 2
-        var T
-          type-con $ctype 0
-        var r
-          int
-      attribute qid 1
-        string-attr VccPrelu.815:13
-      attribute uniqueId 1
-        string-attr 78
-      attribute bvZ3Native 1
-        string-attr False
-    =
-    fun $set_in 2
-    var p
-      type-con $ptr 0
-    fun $full_extent 1
-    fun $ptr 2
-    var T
-      type-con $ctype 0
-    var r
-      int
-    =
-    var p
-      type-con $ptr 0
-    fun $ptr 2
-    var T
-      type-con $ctype 0
-    var r
-      int
-axiom 0
-    forall 1 1 3
-      var T
-        type-con $ctype 0
-      pat 1
-        fun $is_primitive 1
-        var T
-          type-con $ctype 0
-      attribute qid 1
-        string-attr VccPrelu.818:15
-      attribute uniqueId 1
-        string-attr 81
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    fun $is_primitive 1
-    var T
-      type-con $ctype 0
-    forall 3 1 3
-      var S
-        type-con $state 0
-      var r
-        int
-      var p
-        type-con $ptr 0
-      pat 1
-        fun $set_in 2
-        var p
-          type-con $ptr 0
-        fun $extent 2
-        var S
-          type-con $state 0
-        fun $ptr 2
-        var T
-          type-con $ctype 0
-        var r
-          int
-      attribute qid 1
-        string-attr VccPrelu.820:13
-      attribute uniqueId 1
-        string-attr 80
-      attribute bvZ3Native 1
-        string-attr False
-    =
-    fun $set_in 2
-    var p
-      type-con $ptr 0
-    fun $extent 2
-    var S
-      type-con $state 0
-    fun $ptr 2
-    var T
-      type-con $ctype 0
-    var r
-      int
-    =
-    var p
-      type-con $ptr 0
-    fun $ptr 2
-    var T
-      type-con $ctype 0
-    var r
-      int
-axiom 0
-    forall 1 1 3
-      var S
-        type-con $state 0
-      pat 1
-        fun $function_entry 1
-        var S
-          type-con $state 0
-      attribute qid 1
-        string-attr VccPrelu.835:15
-      attribute uniqueId 1
-        string-attr 83
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    fun $function_entry 1
-    var S
-      type-con $state 0
-    and 2
-    fun $full_stop 1
-    var S
-      type-con $state 0
-    >=
-    fun $current_timestamp 1
-    var S
-      type-con $state 0
-    int-num 0
-axiom 0
-    forall 1 1 3
-      var S
-        type-con $state 0
-      pat 1
-        fun $full_stop 1
-        var S
-          type-con $state 0
-      attribute qid 1
-        string-attr VccPrelu.838:15
-      attribute uniqueId 1
-        string-attr 84
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    fun $full_stop 1
-    var S
-      type-con $state 0
-    and 2
-    fun $good_state 1
-    var S
-      type-con $state 0
-    fun $invok_state 1
-    var S
-      type-con $state 0
-axiom 0
-    forall 1 1 3
-      var S
-        type-con $state 0
-      pat 1
-        fun $invok_state 1
-        var S
-          type-con $state 0
-      attribute qid 1
-        string-attr VccPrelu.841:15
-      attribute uniqueId 1
-        string-attr 85
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    fun $invok_state 1
-    var S
-      type-con $state 0
-    fun $good_state 1
-    var S
-      type-con $state 0
-axiom 0
-    forall 2 1 3
-      var id
-        type-con $token 0
-      var S
-        type-con $state 0
-      pat 1
-        fun $good_state_ext 2
-        var id
-          type-con $token 0
-        var S
-          type-con $state 0
-      attribute qid 1
-        string-attr VccPrelu.860:15
-      attribute uniqueId 1
-        string-attr 87
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    fun $good_state_ext 2
-    var id
-      type-con $token 0
-    var S
-      type-con $state 0
-    fun $good_state 1
-    var S
-      type-con $state 0
-axiom 0
-    forall 3 1 3
-      var S
-        type-con $state 0
-      var r
-        int
-      var t
-        type-con $ctype 0
-      pat 1
-        fun $ptr 2
-        var t
-          type-con $ctype 0
-        fun $select.mem 2
-        fun $memory 1
-        var S
-          type-con $state 0
-        fun $ptr 2
-        fun $ptr_to 1
-        var t
-          type-con $ctype 0
-        var r
-          int
-      attribute qid 1
-        string-attr VccPrelu.872:15
-      attribute uniqueId 1
-        string-attr 88
-      attribute bvZ3Native 1
-        string-attr False
-    =
-    fun $ptr 2
-    var t
-      type-con $ctype 0
-    fun $select.mem 2
-    fun $memory 1
-    var S
-      type-con $state 0
-    fun $ptr 2
-    fun $ptr_to 1
-    var t
-      type-con $ctype 0
-    var r
-      int
-    fun $read_ptr_m 3
-    var S
-      type-con $state 0
-    fun $ptr 2
-    fun $ptr_to 1
-    var t
-      type-con $ctype 0
-    var r
-      int
-    var t
-      type-con $ctype 0
-axiom 0
-    forall 2 1 4
-      var S
-        type-con $state 0
-      var p
-        type-con $ptr 0
-      pat 1
-        fun $read_version 2
-        var S
-          type-con $state 0
-        var p
-          type-con $ptr 0
-      attribute qid 1
-        string-attr VccPrelu.886:36
-      attribute uniqueId 1
-        string-attr 89
-      attribute bvZ3Native 1
-        string-attr False
-      attribute weight 1
-        expr-attr
-          int-num 0
-    =
-    fun $read_version 2
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    fun $int_to_version 1
-    fun $select.mem 2
-    fun $memory 1
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-axiom 0
-    forall 2 1 4
-      var S
-        type-con $state 0
-      var p
-        type-con $ptr 0
-      pat 1
-        fun $domain 2
-        var S
-          type-con $state 0
-        var p
-          type-con $ptr 0
-      attribute qid 1
-        string-attr VccPrelu.889:30
-      attribute uniqueId 1
-        string-attr 90
-      attribute bvZ3Native 1
-        string-attr False
-      attribute weight 1
-        expr-attr
-          int-num 0
-    =
-    fun $domain 2
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    fun $ver_domain 1
-    fun $read_version 2
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-axiom 0
-    forall 4 1 4
-      var S
-        type-con $state 0
-      var p
-        type-con $ptr 0
-      var q
-        type-con $ptr 0
-      var l
-        type-con $label 0
-      pat 1
-        fun $in_domain_lab 4
-        var S
-          type-con $state 0
-        var p
-          type-con $ptr 0
-        var q
-          type-con $ptr 0
-        var l
-          type-con $label 0
-      attribute qid 1
-        string-attr VccPrelu.899:15
-      attribute uniqueId 1
-        string-attr 91
-      attribute bvZ3Native 1
-        string-attr False
-      attribute weight 1
-        expr-attr
-          int-num 0
-    implies
-    fun $in_domain_lab 4
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    var q
-      type-con $ptr 0
-    var l
-      type-con $label 0
-    fun $inv_lab 3
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    var l
-      type-con $label 0
-axiom 0
-    forall 4 1 4
-      var S
-        type-con $state 0
-      var p
-        type-con $ptr 0
-      var q
-        type-con $ptr 0
-      var l
-        type-con $label 0
-      pat 1
-        fun $in_domain_lab 4
-        var S
-          type-con $state 0
-        var p
-          type-con $ptr 0
-        var q
-          type-con $ptr 0
-        var l
-          type-con $label 0
-      attribute qid 1
-        string-attr VccPrelu.902:15
-      attribute uniqueId 1
-        string-attr 92
-      attribute bvZ3Native 1
-        string-attr False
-      attribute weight 1
-        expr-attr
-          int-num 0
-    =
-    fun $in_domain_lab 4
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    var q
-      type-con $ptr 0
-    var l
-      type-con $label 0
-    fun $in_domain 3
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    var q
-      type-con $ptr 0
-axiom 0
-    forall 4 1 4
-      var S
-        type-con $state 0
-      var p
-        type-con $ptr 0
-      var q
-        type-con $ptr 0
-      var l
-        type-con $label 0
-      pat 1
-        fun $in_vdomain_lab 4
-        var S
-          type-con $state 0
-        var p
-          type-con $ptr 0
-        var q
-          type-con $ptr 0
-        var l
-          type-con $label 0
-      attribute qid 1
-        string-attr VccPrelu.905:15
-      attribute uniqueId 1
-        string-attr 93
-      attribute bvZ3Native 1
-        string-attr False
-      attribute weight 1
-        expr-attr
-          int-num 0
-    implies
-    fun $in_vdomain_lab 4
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    var q
-      type-con $ptr 0
-    var l
-      type-con $label 0
-    fun $inv_lab 3
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    var l
-      type-con $label 0
-axiom 0
-    forall 4 1 4
-      var S
-        type-con $state 0
-      var p
-        type-con $ptr 0
-      var q
-        type-con $ptr 0
-      var l
-        type-con $label 0
-      pat 1
-        fun $in_vdomain_lab 4
-        var S
-          type-con $state 0
-        var p
-          type-con $ptr 0
-        var q
-          type-con $ptr 0
-        var l
-          type-con $label 0
-      attribute qid 1
-        string-attr VccPrelu.908:15
-      attribute uniqueId 1
-        string-attr 94
-      attribute bvZ3Native 1
-        string-attr False
-      attribute weight 1
-        expr-attr
-          int-num 0
-    =
-    fun $in_vdomain_lab 4
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    var q
-      type-con $ptr 0
-    var l
-      type-con $label 0
-    fun $in_vdomain 3
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    var q
-      type-con $ptr 0
-axiom 0
-    forall 3 1 4
-      var S
-        type-con $state 0
-      var p
-        type-con $ptr 0
-      var q
-        type-con $ptr 0
-      pat 1
-        fun $in_domain 3
-        var S
-          type-con $state 0
-        var p
-          type-con $ptr 0
-        var q
-          type-con $ptr 0
-      attribute qid 1
-        string-attr VccPrelu.914:15
-      attribute uniqueId 1
-        string-attr 96
-      attribute bvZ3Native 1
-        string-attr False
-      attribute weight 1
-        expr-attr
-          int-num 0
-    implies
-    fun $in_domain 3
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    var q
-      type-con $ptr 0
-    and 3
-    fun $set_in 2
-    var p
-      type-con $ptr 0
-    fun $domain 2
-    var S
-      type-con $state 0
-    var q
-      type-con $ptr 0
-    fun $closed 2
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    forall 1 1 3
-      var r
-        type-con $ptr 0
-      pat 1
-        fun $set_in 2
-        var r
-          type-con $ptr 0
-        fun $owns 2
-        var S
-          type-con $state 0
-        var p
-          type-con $ptr 0
-      attribute qid 1
-        string-attr VccPrelu.918:16
-      attribute uniqueId 1
-        string-attr 95
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 2
-    not
-    fun $has_volatile_owns_set 1
-    fun $typ 1
-    var p
-      type-con $ptr 0
-    fun $set_in 2
-    var r
-      type-con $ptr 0
-    fun $owns 2
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    fun $set_in2 2
-    var r
-      type-con $ptr 0
-    fun $ver_domain 1
-    fun $read_version 2
-    var S
-      type-con $state 0
-    var q
-      type-con $ptr 0
-axiom 0
-    forall 2 1 3
-      var S
-        type-con $state 0
-      var p
-        type-con $ptr 0
-      pat 1
-        fun $in_domain 3
-        var S
-          type-con $state 0
-        var p
-          type-con $ptr 0
-        var p
-          type-con $ptr 0
-      attribute qid 1
-        string-attr VccPrelu.923:15
-      attribute uniqueId 1
-        string-attr 97
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 7
-    fun $full_stop 1
-    var S
-      type-con $state 0
-    fun $closed 2
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    =
-    fun $owner 2
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    fun $me 0
-    fun $is 2
-    var p
-      type-con $ptr 0
-    fun $typ 1
-    var p
-      type-con $ptr 0
-    fun $typed 2
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    not
-    =
-    fun $kind_of 1
-    fun $typ 1
-    var p
-      type-con $ptr 0
-    fun $kind_primitive 0
-    fun $is_non_primitive 1
-    fun $typ 1
-    var p
-      type-con $ptr 0
-    fun $in_domain 3
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    var p
-      type-con $ptr 0
-axiom 0
-    forall 3 1 4
-      var S
-        type-con $state 0
-      var p
-        type-con $ptr 0
-      var q
-        type-con $ptr 0
-      pat 1
-        fun $in_domain 3
-        var S
-          type-con $state 0
-        var q
-          type-con $ptr 0
-        var p
-          type-con $ptr 0
-      attribute qid 1
-        string-attr VccPrelu.932:15
-      attribute uniqueId 1
-        string-attr 98
-      attribute bvZ3Native 1
-        string-attr False
-      attribute weight 1
-        expr-attr
-          int-num 0
-    implies
-    and 2
-    fun $full_stop 1
-    var S
-      type-con $state 0
-    fun $set_in 2
-    var q
-      type-con $ptr 0
-    fun $domain 2
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    fun $in_domain 3
-    var S
-      type-con $state 0
-    var q
-      type-con $ptr 0
-    var p
-      type-con $ptr 0
-axiom 0
-    forall 4 1 4
-      var S
-        type-con $state 0
-      var p
-        type-con $ptr 0
-      var q
-        type-con $ptr 0
-      var r
-        type-con $ptr 0
-      pat 2
-        fun $set_in 2
-        var q
-          type-con $ptr 0
-        fun $domain 2
-        var S
-          type-con $state 0
-        var p
-          type-con $ptr 0
-        fun $in_domain 3
-        var S
-          type-con $state 0
-        var r
-          type-con $ptr 0
-        var p
-          type-con $ptr 0
-      attribute qid 1
-        string-attr VccPrelu.936:15
-      attribute uniqueId 1
-        string-attr 99
-      attribute bvZ3Native 1
-        string-attr False
-      attribute weight 1
-        expr-attr
-          int-num 0
-    implies
-    and 3
-    not
-    fun $has_volatile_owns_set 1
-    fun $typ 1
-    var q
-      type-con $ptr 0
-    fun $set_in 2
-    var q
-      type-con $ptr 0
-    fun $domain 2
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    fun $set_in0 2
-    var r
-      type-con $ptr 0
-    fun $owns 2
-    var S
-      type-con $state 0
-    var q
-      type-con $ptr 0
-    and 2
-    fun $in_domain 3
-    var S
-      type-con $state 0
-    var r
-      type-con $ptr 0
-    var p
-      type-con $ptr 0
-    fun $set_in0 2
-    var r
-      type-con $ptr 0
-    fun $owns 2
-    var S
-      type-con $state 0
-    var q
-      type-con $ptr 0
-axiom 0
-    forall 4 1 4
-      var S
-        type-con $state 0
-      var p
-        type-con $ptr 0
-      var q
-        type-con $ptr 0
-      var r
-        type-con $ptr 0
-      pat 2
-        fun $set_in 2
-        var q
-          type-con $ptr 0
-        fun $domain 2
-        var S
-          type-con $state 0
-        var p
-          type-con $ptr 0
-        fun $in_vdomain 3
-        var S
-          type-con $state 0
-        var r
-          type-con $ptr 0
-        var p
-          type-con $ptr 0
-      attribute qid 1
-        string-attr VccPrelu.941:15
-      attribute uniqueId 1
-        string-attr 101
-      attribute bvZ3Native 1
-        string-attr False
-      attribute weight 1
-        expr-attr
-          int-num 0
-    implies
-    and 3
-    fun $has_volatile_owns_set 1
-    fun $typ 1
-    var q
-      type-con $ptr 0
-    fun $set_in 2
-    var q
-      type-con $ptr 0
-    fun $domain 2
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    forall 1 0 3
-      var S1
-        type-con $state 0
-      attribute qid 1
-        string-attr VccPrelu.945:11
-      attribute uniqueId 1
-        string-attr 100
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 3
-    fun $inv2 4
-    var S1
-      type-con $state 0
-    var S1
-      type-con $state 0
-    var q
-      type-con $ptr 0
-    fun $typ 1
-    var q
-      type-con $ptr 0
-    =
-    fun $read_version 2
-    var S1
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    fun $read_version 2
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    =
-    fun $domain 2
-    var S1
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    fun $domain 2
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    fun $set_in0 2
-    var r
-      type-con $ptr 0
-    fun $owns 2
-    var S1
-      type-con $state 0
-    var q
-      type-con $ptr 0
-    and 2
-    fun $in_vdomain 3
-    var S
-      type-con $state 0
-    var r
-      type-con $ptr 0
-    var p
-      type-con $ptr 0
-    fun $set_in0 2
-    var r
-      type-con $ptr 0
-    fun $owns 2
-    var S
-      type-con $state 0
-    var q
-      type-con $ptr 0
-axiom 0
-    forall 3 1 4
-      var S
-        type-con $state 0
-      var p
-        type-con $ptr 0
-      var q
-        type-con $ptr 0
-      pat 1
-        fun $in_vdomain 3
-        var S
-          type-con $state 0
-        var p
-          type-con $ptr 0
-        var q
-          type-con $ptr 0
-      attribute qid 1
-        string-attr VccPrelu.952:15
-      attribute uniqueId 1
-        string-attr 102
-      attribute bvZ3Native 1
-        string-attr False
-      attribute weight 1
-        expr-attr
-          int-num 0
-    implies
-    fun $in_vdomain 3
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    var q
-      type-con $ptr 0
-    fun $in_domain 3
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    var q
-      type-con $ptr 0
-axiom 0
-    forall 4 1 3
-      var S
-        type-con $state 0
-      var p
-        type-con $ptr 0
-      var d
-        type-con $ptr 0
-      var f
-        type-con $field 0
-      pat 3
-        fun $set_in 2
-        var p
-          type-con $ptr 0
-        fun $domain 2
-        var S
-          type-con $state 0
-        var d
-          type-con $ptr 0
-        fun $is_primitive_non_volatile_field 1
-        var f
-          type-con $field 0
-        fun $select.mem 2
-        fun $memory 1
-        var S
-          type-con $state 0
-        fun $dot 2
-        var p
-          type-con $ptr 0
-        var f
-          type-con $field 0
-      attribute qid 1
-        string-attr VccPrelu.957:15
-      attribute uniqueId 1
-        string-attr 103
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 2
-    fun $set_in 2
-    var p
-      type-con $ptr 0
-    fun $domain 2
-    var S
-      type-con $state 0
-    var d
-      type-con $ptr 0
-    fun $is_primitive_non_volatile_field 1
-    var f
-      type-con $field 0
-    =
-    fun $select.mem 2
-    fun $memory 1
-    var S
-      type-con $state 0
-    fun $dot 2
-    var p
-      type-con $ptr 0
-    var f
-      type-con $field 0
-    fun $fetch_from_domain 2
-    fun $read_version 2
-    var S
-      type-con $state 0
-    var d
-      type-con $ptr 0
-    fun $dot 2
-    var p
-      type-con $ptr 0
-    var f
-      type-con $field 0
-axiom 0
-    forall 3 2 3
-      var S
-        type-con $state 0
-      var p
-        type-con $ptr 0
-      var d
-        type-con $ptr 0
-      pat 3
-        fun $full_stop 1
-        var S
-          type-con $state 0
-        fun $set_in 2
-        var p
-          type-con $ptr 0
-        fun $domain 2
-        var S
-          type-con $state 0
-        var d
-          type-con $ptr 0
-        fun $select.sm 2
-        fun $statusmap 1
-        var S
-          type-con $state 0
-        var p
-          type-con $ptr 0
-      pat 3
-        fun $full_stop 1
-        var S
-          type-con $state 0
-        fun $set_in 2
-        var p
-          type-con $ptr 0
-        fun $domain 2
-        var S
-          type-con $state 0
-        var d
-          type-con $ptr 0
-        fun $select.tm 2
-        fun $typemap 1
-        var S
-          type-con $state 0
-        var p
-          type-con $ptr 0
-      attribute qid 1
-        string-attr VccPrelu.962:15
-      attribute uniqueId 1
-        string-attr 104
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 2
-    fun $full_stop 1
-    var S
-      type-con $state 0
-    fun $set_in 2
-    var p
-      type-con $ptr 0
-    fun $domain 2
-    var S
-      type-con $state 0
-    var d
-      type-con $ptr 0
-    and 2
-    fun $typed 2
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    not
-    fun $ts_is_volatile 1
-    fun $select.tm 2
-    fun $typemap 1
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-axiom 0
-    forall 4 2 3
-      var S
-        type-con $state 0
-      var p
-        type-con $ptr 0
-      var d
-        type-con $ptr 0
-      var f
-        type-con $field 0
-      pat 3
-        fun $set_in 2
-        var p
-          type-con $ptr 0
-        fun $domain 2
-        var S
-          type-con $state 0
-        var d
-          type-con $ptr 0
-        fun $is_primitive_non_volatile_field 1
-        var f
-          type-con $field 0
-        fun $owner 2
-        var S
-          type-con $state 0
-        fun $dot 2
-        var p
-          type-con $ptr 0
-        var f
-          type-con $field 0
-      pat 3
-        fun $set_in 2
-        var p
-          type-con $ptr 0
-        fun $domain 2
-        var S
-          type-con $state 0
-        var d
-          type-con $ptr 0
-        fun $is_primitive_non_volatile_field 1
-        var f
-          type-con $field 0
-        fun $select.tm 2
-        fun $typemap 1
-        var S
-          type-con $state 0
-        fun $dot 2
-        var p
-          type-con $ptr 0
-        var f
-          type-con $field 0
-      attribute qid 1
-        string-attr VccPrelu.968:15
-      attribute uniqueId 1
-        string-attr 105
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 3
-    fun $full_stop 1
-    var S
-      type-con $state 0
-    fun $set_in 2
-    var p
-      type-con $ptr 0
-    fun $domain 2
-    var S
-      type-con $state 0
-    var d
-      type-con $ptr 0
-    fun $is_primitive_non_volatile_field 1
-    var f
-      type-con $field 0
-    and 2
-    fun $typed 2
-    var S
-      type-con $state 0
-    fun $dot 2
-    var p
-      type-con $ptr 0
-    var f
-      type-con $field 0
-    not
-    fun $ts_is_volatile 1
-    fun $select.tm 2
-    fun $typemap 1
-    var S
-      type-con $state 0
-    fun $dot 2
-    var p
-      type-con $ptr 0
-    var f
-      type-con $field 0
-axiom 0
-    forall 7 1 3
-      var S
-        type-con $state 0
-      var p
-        type-con $ptr 0
-      var d
-        type-con $ptr 0
-      var f
-        type-con $field 0
-      var sz
-        int
-      var i
-        int
-      var t
-        type-con $ctype 0
-      pat 3
-        fun $set_in 2
-        var p
-          type-con $ptr 0
-        fun $domain 2
-        var S
-          type-con $state 0
-        var d
-          type-con $ptr 0
-        fun $is_primitive_embedded_array 2
-        var f
-          type-con $field 0
-        var sz
-          int
-        fun $select.mem 2
-        fun $memory 1
-        var S
-          type-con $state 0
-        fun $idx 3
-        fun $dot 2
-        var p
-          type-con $ptr 0
-        var f
-          type-con $field 0
-        var i
-          int
-        var t
-          type-con $ctype 0
-      attribute qid 1
-        string-attr VccPrelu.974:15
-      attribute uniqueId 1
-        string-attr 106
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 5
-    fun $full_stop 1
-    var S
-      type-con $state 0
-    fun $set_in 2
-    var p
-      type-con $ptr 0
-    fun $domain 2
-    var S
-      type-con $state 0
-    var d
-      type-con $ptr 0
-    fun $is_primitive_embedded_array 2
-    var f
-      type-con $field 0
-    var sz
-      int
-    <=
-    int-num 0
-    var i
-      int
-    <
-    var i
-      int
-    var sz
-      int
-    =
-    fun $select.mem 2
-    fun $memory 1
-    var S
-      type-con $state 0
-    fun $idx 3
-    fun $dot 2
-    var p
-      type-con $ptr 0
-    var f
-      type-con $field 0
-    var i
-      int
-    var t
-      type-con $ctype 0
-    fun $fetch_from_domain 2
-    fun $read_version 2
-    var S
-      type-con $state 0
-    var d
-      type-con $ptr 0
-    fun $idx 3
-    fun $dot 2
-    var p
-      type-con $ptr 0
-    var f
-      type-con $field 0
-    var i
-      int
-    var t
-      type-con $ctype 0
-axiom 0
-    forall 7 2 3
-      var S
-        type-con $state 0
-      var p
-        type-con $ptr 0
-      var d
-        type-con $ptr 0
-      var f
-        type-con $field 0
-      var sz
-        int
-      var i
-        int
-      var t
-        type-con $ctype 0
-      pat 3
-        fun $set_in 2
-        var p
-          type-con $ptr 0
-        fun $domain 2
-        var S
-          type-con $state 0
-        var d
-          type-con $ptr 0
-        fun $is_primitive_embedded_array 2
-        var f
-          type-con $field 0
-        var sz
-          int
-        fun $select.tm 2
-        fun $typemap 1
-        var S
-          type-con $state 0
-        fun $idx 3
-        fun $dot 2
-        var p
-          type-con $ptr 0
-        var f
-          type-con $field 0
-        var i
-          int
-        var t
-          type-con $ctype 0
-      pat 3
-        fun $set_in 2
-        var p
-          type-con $ptr 0
-        fun $domain 2
-        var S
-          type-con $state 0
-        var d
-          type-con $ptr 0
-        fun $is_primitive_embedded_array 2
-        var f
-          type-con $field 0
-        var sz
-          int
-        fun $owner 2
-        var S
-          type-con $state 0
-        fun $idx 3
-        fun $dot 2
-        var p
-          type-con $ptr 0
-        var f
-          type-con $field 0
-        var i
-          int
-        var t
-          type-con $ctype 0
-      attribute qid 1
-        string-attr VccPrelu.979:15
-      attribute uniqueId 1
-        string-attr 107
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 5
-    fun $full_stop 1
-    var S
-      type-con $state 0
-    fun $set_in 2
-    var p
-      type-con $ptr 0
-    fun $domain 2
-    var S
-      type-con $state 0
-    var d
-      type-con $ptr 0
-    fun $is_primitive_embedded_array 2
-    var f
-      type-con $field 0
-    var sz
-      int
-    <=
-    int-num 0
-    var i
-      int
-    <
-    var i
-      int
-    var sz
-      int
-    and 2
-    fun $typed 2
-    var S
-      type-con $state 0
-    fun $idx 3
-    fun $dot 2
-    var p
-      type-con $ptr 0
-    var f
-      type-con $field 0
-    var i
-      int
-    var t
-      type-con $ctype 0
-    not
-    fun $ts_is_volatile 1
-    fun $select.tm 2
-    fun $typemap 1
-    var S
-      type-con $state 0
-    fun $idx 3
-    fun $dot 2
-    var p
-      type-con $ptr 0
-    var f
-      type-con $field 0
-    var i
-      int
-    var t
-      type-con $ctype 0
-axiom 0
-    forall 6 2 3
-      var S
-        type-con $state 0
-      var r
-        int
-      var d
-        type-con $ptr 0
-      var sz
-        int
-      var i
-        int
-      var t
-        type-con $ctype 0
-      pat 3
-        fun $set_in 2
-        fun $ptr 2
-        fun $array 2
-        var t
-          type-con $ctype 0
-        var sz
-          int
-        var r
-          int
-        fun $domain 2
-        var S
-          type-con $state 0
-        var d
-          type-con $ptr 0
-        fun $select.tm 2
-        fun $typemap 1
-        var S
-          type-con $state 0
-        fun $idx 3
-        fun $ptr 2
-        var t
-          type-con $ctype 0
-        var r
-          int
-        var i
-          int
-        var t
-          type-con $ctype 0
-        fun $is_primitive 1
-        var t
-          type-con $ctype 0
-      pat 3
-        fun $set_in 2
-        fun $ptr 2
-        fun $array 2
-        var t
-          type-con $ctype 0
-        var sz
-          int
-        var r
-          int
-        fun $domain 2
-        var S
-          type-con $state 0
-        var d
-          type-con $ptr 0
-        fun $owner 2
-        var S
-          type-con $state 0
-        fun $idx 3
-        fun $ptr 2
-        var t
-          type-con $ctype 0
-        var r
-          int
-        var i
-          int
-        var t
-          type-con $ctype 0
-        fun $is_primitive 1
-        var t
-          type-con $ctype 0
-      attribute qid 1
-        string-attr VccPrelu.985:15
-      attribute uniqueId 1
-        string-attr 108
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 5
-    fun $full_stop 1
-    var S
-      type-con $state 0
-    fun $is_primitive 1
-    var t
-      type-con $ctype 0
-    fun $set_in 2
-    fun $ptr 2
-    fun $array 2
-    var t
-      type-con $ctype 0
-    var sz
-      int
-    var r
-      int
-    fun $domain 2
-    var S
-      type-con $state 0
-    var d
-      type-con $ptr 0
-    <=
-    int-num 0
-    var i
-      int
-    <
-    var i
-      int
-    var sz
-      int
-    and 2
-    fun $typed 2
-    var S
-      type-con $state 0
-    fun $idx 3
-    fun $ptr 2
-    var t
-      type-con $ctype 0
-    var r
-      int
-    var i
-      int
-    var t
-      type-con $ctype 0
-    not
-    fun $ts_is_volatile 1
-    fun $select.tm 2
-    fun $typemap 1
-    var S
-      type-con $state 0
-    fun $idx 3
-    fun $ptr 2
-    var t
-      type-con $ctype 0
-    var r
-      int
-    var i
-      int
-    var t
-      type-con $ctype 0
-axiom 0
-    forall 6 1 3
-      var S
-        type-con $state 0
-      var r
-        int
-      var d
-        type-con $ptr 0
-      var sz
-        int
-      var i
-        int
-      var t
-        type-con $ctype 0
-      pat 3
-        fun $set_in 2
-        fun $ptr 2
-        fun $array 2
-        var t
-          type-con $ctype 0
-        var sz
-          int
-        var r
-          int
-        fun $domain 2
-        var S
-          type-con $state 0
-        var d
-          type-con $ptr 0
-        fun $select.mem 2
-        fun $memory 1
-        var S
-          type-con $state 0
-        fun $idx 3
-        fun $ptr 2
-        var t
-          type-con $ctype 0
-        var r
-          int
-        var i
-          int
-        var t
-          type-con $ctype 0
-        fun $is_primitive 1
-        var t
-          type-con $ctype 0
-      attribute qid 1
-        string-attr VccPrelu.994:15
-      attribute uniqueId 1
-        string-attr 109
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 5
-    fun $full_stop 1
-    var S
-      type-con $state 0
-    fun $is_primitive 1
-    var t
-      type-con $ctype 0
-    fun $set_in 2
-    fun $ptr 2
-    fun $array 2
-    var t
-      type-con $ctype 0
-    var sz
-      int
-    var r
-      int
-    fun $domain 2
-    var S
-      type-con $state 0
-    var d
-      type-con $ptr 0
-    <=
-    int-num 0
-    var i
-      int
-    <
-    var i
-      int
-    var sz
-      int
-    =
-    fun $select.mem 2
-    fun $memory 1
-    var S
-      type-con $state 0
-    fun $idx 3
-    fun $ptr 2
-    var t
-      type-con $ctype 0
-    var r
-      int
-    var i
-      int
-    var t
-      type-con $ctype 0
-    fun $fetch_from_domain 2
-    fun $read_version 2
-    var S
-      type-con $state 0
-    var d
-      type-con $ptr 0
-    fun $idx 3
-    fun $ptr 2
-    var t
-      type-con $ctype 0
-    var r
-      int
-    var i
-      int
-    var t
-      type-con $ctype 0
-axiom 0
-    forall 6 1 3
-      var S
-        type-con $state 0
-      var p
-        type-con $ptr 0
-      var f
-        type-con $field 0
-      var sz
-        int
-      var i
-        int
-      var t
-        type-con $ctype 0
-      pat 2
-        fun $is_primitive_embedded_volatile_array 3
-        var f
-          type-con $field 0
-        var sz
-          int
-        var t
-          type-con $ctype 0
-        fun $ts_is_volatile 1
-        fun $select.tm 2
-        fun $typemap 1
-        var S
-          type-con $state 0
-        fun $idx 3
-        fun $dot 2
-        var p
-          type-con $ptr 0
-        var f
-          type-con $field 0
-        var i
-          int
-        var t
-          type-con $ctype 0
-      attribute qid 1
-        string-attr VccPrelu.1002:15
-      attribute uniqueId 1
-        string-attr 110
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 4
-    fun $good_state 1
-    var S
-      type-con $state 0
-    fun $is_primitive_embedded_volatile_array 3
-    var f
-      type-con $field 0
-    var sz
-      int
-    var t
-      type-con $ctype 0
-    <=
-    int-num 0
-    var i
-      int
-    <
-    var i
-      int
-    var sz
-      int
-    fun $ts_is_volatile 1
-    fun $select.tm 2
-    fun $typemap 1
-    var S
-      type-con $state 0
-    fun $idx 3
-    fun $dot 2
-    var p
-      type-con $ptr 0
-    var f
-      type-con $field 0
-    var i
-      int
-    var t
-      type-con $ctype 0
-axiom 0
-    forall 4 1 4
-      var p
-        type-con $ptr 0
-      var S1
-        type-con $state 0
-      var S2
-        type-con $state 0
-      var q
-        type-con $ptr 0
-      pat 2
-        fun $set_in 2
-        var q
-          type-con $ptr 0
-        fun $domain 2
-        var S1
-          type-con $state 0
-        var p
-          type-con $ptr 0
-        fun $call_transition 2
-        var S1
-          type-con $state 0
-        var S2
-          type-con $state 0
-      attribute qid 1
-        string-attr VccPrelu.1013:15
-      attribute uniqueId 1
-        string-attr 111
-      attribute bvZ3Native 1
-        string-attr False
-      attribute weight 1
-        expr-attr
-          int-num 0
-    fun $instantiate_bool 1
-    fun $set_in 2
-    var q
-      type-con $ptr 0
-    fun $domain 2
-    var S2
-      type-con $state 0
-    var p
-      type-con $ptr 0
-axiom 0
-    forall 4 1 4
-      var p
-        type-con $ptr 0
-      var S1
-        type-con $state 0
-      var S2
-        type-con $state 0
-      var q
-        type-con $ptr 0
-      pat 2
-        fun $set_in 2
-        var q
-          type-con $ptr 0
-        fun $ver_domain 1
-        fun $read_version 2
-        var S1
-          type-con $state 0
-        var p
-          type-con $ptr 0
-        fun $call_transition 2
-        var S1
-          type-con $state 0
-        var S2
-          type-con $state 0
-      attribute qid 1
-        string-attr VccPrelu.1017:15
-      attribute uniqueId 1
-        string-attr 112
-      attribute bvZ3Native 1
-        string-attr False
-      attribute weight 1
-        expr-attr
-          int-num 0
-    fun $instantiate_bool 1
-    fun $set_in 2
-    var q
-      type-con $ptr 0
-    fun $ver_domain 1
-    fun $read_version 2
-    var S2
-      type-con $state 0
-    var p
-      type-con $ptr 0
-axiom 0
-    forall 2 1 3
-      var p
-        type-con $ptr 0
-      var c
-        type-con $ptr 0
-      pat 1
-        fun $in_claim_domain 2
-        var p
-          type-con $ptr 0
-        var c
-          type-con $ptr 0
-      attribute qid 1
-        string-attr VccPrelu.1022:15
-      attribute uniqueId 1
-        string-attr 114
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    forall 1 1 3
-      var s
-        type-con $state 0
-      pat 1
-        fun $dont_instantiate_state 1
-        var s
-          type-con $state 0
-      attribute qid 1
-        string-attr VccPrelu.1023:11
-      attribute uniqueId 1
-        string-attr 113
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    fun $valid_claim 2
-    var s
-      type-con $state 0
-    var c
-      type-con $ptr 0
-    fun $closed 2
-    var s
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    fun $in_claim_domain 2
-    var p
-      type-con $ptr 0
-    var c
-      type-con $ptr 0
-axiom 0
-    forall 4 1 4
-      var S
-        type-con $state 0
-      var c
-        type-con $ptr 0
-      var obj
-        type-con $ptr 0
-      var ptr
-        type-con $ptr 0
-      pat 1
-        fun $by_claim 4
-        var S
-          type-con $state 0
-        var c
-          type-con $ptr 0
-        var obj
-          type-con $ptr 0
-        var ptr
-          type-con $ptr 0
-      attribute qid 1
-        string-attr VccPrelu.1026:32
-      attribute uniqueId 1
-        string-attr 115
-      attribute bvZ3Native 1
-        string-attr False
-      attribute weight 1
-        expr-attr
-          int-num 0
-    =
-    fun $by_claim 4
-    var S
-      type-con $state 0
-    var c
-      type-con $ptr 0
-    var obj
-      type-con $ptr 0
-    var ptr
-      type-con $ptr 0
-    var ptr
-      type-con $ptr 0
-axiom 0
-    forall 4 2 3
-      var S
-        type-con $state 0
-      var p
-        type-con $ptr 0
-      var c
-        type-con $ptr 0
-      var f
-        type-con $field 0
-      pat 2
-        fun $in_claim_domain 2
-        var p
-          type-con $ptr 0
-        var c
-          type-con $ptr 0
-        fun $select.mem 2
-        fun $memory 1
-        var S
-          type-con $state 0
-        fun $dot 2
-        var p
-          type-con $ptr 0
-        var f
-          type-con $field 0
-      pat 1
-        fun $by_claim 4
-        var S
-          type-con $state 0
-        var c
-          type-con $ptr 0
-        var p
-          type-con $ptr 0
-        fun $dot 2
-        var p
-          type-con $ptr 0
-        var f
-          type-con $field 0
-      attribute qid 1
-        string-attr VccPrelu.1031:15
-      attribute uniqueId 1
-        string-attr 116
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 4
-    fun $good_state 1
-    var S
-      type-con $state 0
-    fun $closed 2
-    var S
-      type-con $state 0
-    var c
-      type-con $ptr 0
-    fun $in_claim_domain 2
-    var p
-      type-con $ptr 0
-    var c
-      type-con $ptr 0
-    fun $is_primitive_non_volatile_field 1
-    var f
-      type-con $field 0
-    and 2
-    fun $in_claim_domain 2
-    var p
-      type-con $ptr 0
-    var c
-      type-con $ptr 0
-    =
-    fun $select.mem 2
-    fun $memory 1
-    var S
-      type-con $state 0
-    fun $dot 2
-    var p
-      type-con $ptr 0
-    var f
-      type-con $field 0
-    fun $fetch_from_domain 2
-    fun $claim_version 1
-    var c
-      type-con $ptr 0
-    fun $dot 2
-    var p
-      type-con $ptr 0
-    var f
-      type-con $field 0
-axiom 0
-    forall 7 2 3
-      var S
-        type-con $state 0
-      var p
-        type-con $ptr 0
-      var c
-        type-con $ptr 0
-      var f
-        type-con $field 0
-      var i
-        int
-      var sz
-        int
-      var t
-        type-con $ctype 0
-      pat 4
-        fun $valid_claim 2
-        var S
-          type-con $state 0
-        var c
-          type-con $ptr 0
-        fun $in_claim_domain 2
-        var p
-          type-con $ptr 0
-        var c
-          type-con $ptr 0
-        fun $select.mem 2
-        fun $memory 1
-        var S
-          type-con $state 0
-        fun $idx 3
-        fun $dot 2
-        var p
-          type-con $ptr 0
-        var f
-          type-con $field 0
-        var i
-          int
-        var t
-          type-con $ctype 0
-        fun $is_primitive_embedded_array 2
-        var f
-          type-con $field 0
-        var sz
-          int
-      pat 2
-        fun $by_claim 4
-        var S
-          type-con $state 0
-        var c
-          type-con $ptr 0
-        var p
-          type-con $ptr 0
-        fun $idx 3
-        fun $dot 2
-        var p
-          type-con $ptr 0
-        var f
-          type-con $field 0
-        var i
-          int
-        var t
-          type-con $ctype 0
-        fun $is_primitive_embedded_array 2
-        var f
-          type-con $field 0
-        var sz
-          int
-      attribute qid 1
-        string-attr VccPrelu.1040:15
-      attribute uniqueId 1
-        string-attr 117
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 6
-    fun $good_state 1
-    var S
-      type-con $state 0
-    fun $closed 2
-    var S
-      type-con $state 0
-    var c
-      type-con $ptr 0
-    fun $in_claim_domain 2
-    var p
-      type-con $ptr 0
-    var c
-      type-con $ptr 0
-    fun $is_primitive_embedded_array 2
-    var f
-      type-con $field 0
-    var sz
-      int
-    <=
-    int-num 0
-    var i
-      int
-    <
-    var i
-      int
-    var sz
-      int
-    =
-    fun $select.mem 2
-    fun $memory 1
-    var S
-      type-con $state 0
-    fun $idx 3
-    fun $dot 2
-    var p
-      type-con $ptr 0
-    var f
-      type-con $field 0
-    var i
-      int
-    var t
-      type-con $ctype 0
-    fun $fetch_from_domain 2
-    fun $claim_version 1
-    var c
-      type-con $ptr 0
-    fun $idx 3
-    fun $dot 2
-    var p
-      type-con $ptr 0
-    var f
-      type-con $field 0
-    var i
-      int
-    var t
-      type-con $ctype 0
-axiom 0
-    forall 2 1 4
-      var S
-        type-con $state 0
-      var p
-        type-con $ptr 0
-      pat 1
-        fun $read_vol_version 2
-        var S
-          type-con $state 0
-        var p
-          type-con $ptr 0
-      attribute qid 1
-        string-attr VccPrelu.1067:40
-      attribute uniqueId 1
-        string-attr 119
-      attribute bvZ3Native 1
-        string-attr False
-      attribute weight 1
-        expr-attr
-          int-num 0
-    =
-    fun $read_vol_version 2
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    fun $int_to_vol_version 1
-    fun $select.mem 2
-    fun $memory 1
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-axiom 0
-    forall 5 1 3
-      var S
-        type-con $state 0
-      var r
-        int
-      var t
-        type-con $ctype 0
-      var approver
-        type-con $field 0
-      var subject
-        type-con $field 0
-      pat 2
-        fun $is_approved_by 3
-        var t
-          type-con $ctype 0
-        var approver
-          type-con $field 0
-        var subject
-          type-con $field 0
-        fun $select.mem 2
-        fun $memory 1
-        var S
-          type-con $state 0
-        fun $dot 2
-        fun $ptr 2
-        var t
-          type-con $ctype 0
-        var r
-          int
-        var subject
-          type-con $field 0
-      attribute qid 1
-        string-attr VccPrelu.1078:15
-      attribute uniqueId 1
-        string-attr 120
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 4
-    fun $full_stop 1
-    var S
-      type-con $state 0
-    fun $is_approved_by 3
-    var t
-      type-con $ctype 0
-    var approver
-      type-con $field 0
-    var subject
-      type-con $field 0
-    fun $closed 2
-    var S
-      type-con $state 0
-    fun $ptr 2
-    var t
-      type-con $ctype 0
-    var r
-      int
-    or 2
-    =
-    fun $int_to_ptr 1
-    fun $select.mem 2
-    fun $memory 1
-    var S
-      type-con $state 0
-    fun $dot 2
-    fun $ptr 2
-    var t
-      type-con $ctype 0
-    var r
-      int
-    var approver
-      type-con $field 0
-    fun $me 0
-    =
-    fun $int_to_ptr 1
-    fun $fetch_from_vv 2
-    fun $read_vol_version 2
-    var S
-      type-con $state 0
-    fun $ptr 2
-    var t
-      type-con $ctype 0
-    var r
-      int
-    fun $dot 2
-    fun $ptr 2
-    var t
-      type-con $ctype 0
-    var r
-      int
-    var approver
-      type-con $field 0
-    fun $me 0
-    =
-    fun $select.mem 2
-    fun $memory 1
-    var S
-      type-con $state 0
-    fun $dot 2
-    fun $ptr 2
-    var t
-      type-con $ctype 0
-    var r
-      int
-    var subject
-      type-con $field 0
-    fun $fetch_from_vv 2
-    fun $read_vol_version 2
-    var S
-      type-con $state 0
-    fun $ptr 2
-    var t
-      type-con $ctype 0
-    var r
-      int
-    fun $dot 2
-    fun $ptr 2
-    var t
-      type-con $ctype 0
-    var r
-      int
-    var subject
-      type-con $field 0
-axiom 0
-    forall 4 1 3
-      var S
-        type-con $state 0
-      var r
-        int
-      var t
-        type-con $ctype 0
-      var subject
-        type-con $field 0
-      pat 2
-        fun $is_owner_approved 2
-        var t
-          type-con $ctype 0
-        var subject
-          type-con $field 0
-        fun $select.mem 2
-        fun $memory 1
-        var S
-          type-con $state 0
-        fun $dot 2
-        fun $ptr 2
-        var t
-          type-con $ctype 0
-        var r
-          int
-        var subject
-          type-con $field 0
-      attribute qid 1
-        string-attr VccPrelu.1103:15
-      attribute uniqueId 1
-        string-attr 121
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 4
-    fun $full_stop 1
-    var S
-      type-con $state 0
-    fun $closed 2
-    var S
-      type-con $state 0
-    fun $ptr 2
-    var t
-      type-con $ctype 0
-    var r
-      int
-    fun $is_owner_approved 2
-    var t
-      type-con $ctype 0
-    var subject
-      type-con $field 0
-    =
-    fun $owner 2
-    var S
-      type-con $state 0
-    fun $ptr 2
-    var t
-      type-con $ctype 0
-    var r
-      int
-    fun $me 0
-    =
-    fun $select.mem 2
-    fun $memory 1
-    var S
-      type-con $state 0
-    fun $dot 2
-    fun $ptr 2
-    var t
-      type-con $ctype 0
-    var r
-      int
-    var subject
-      type-con $field 0
-    fun $fetch_from_vv 2
-    fun $read_vol_version 2
-    var S
-      type-con $state 0
-    fun $ptr 2
-    var t
-      type-con $ctype 0
-    var r
-      int
-    fun $dot 2
-    fun $ptr 2
-    var t
-      type-con $ctype 0
-    var r
-      int
-    var subject
-      type-con $field 0
-axiom 0
-    forall 5 1 3
-      var S1
-        type-con $state 0
-      var S2
-        type-con $state 0
-      var r
-        int
-      var t
-        type-con $ctype 0
-      var subject
-        type-con $field 0
-      pat 3
-        fun $is_owner_approved 2
-        var t
-          type-con $ctype 0
-        var subject
-          type-con $field 0
-        fun $post_unwrap 2
-        var S1
-          type-con $state 0
-        var S2
-          type-con $state 0
-        fun $select.mem 2
-        fun $memory 1
-        var S1
-          type-con $state 0
-        fun $dot 2
-        fun $ptr 2
-        var t
-          type-con $ctype 0
-        var r
-          int
-        var subject
-          type-con $field 0
-      attribute qid 1
-        string-attr VccPrelu.1111:15
-      attribute uniqueId 1
-        string-attr 122
-      attribute bvZ3Native 1
-        string-attr False
-    fun $instantiate_int 1
-    fun $select.mem 2
-    fun $memory 1
-    var S2
-      type-con $state 0
-    fun $dot 2
-    fun $ptr 2
-    var t
-      type-con $ctype 0
-    var r
-      int
-    var subject
-      type-con $field 0
-axiom 0
-    forall 3 1 3
-      var S
-        type-con $state 0
-      var p
-        type-con $ptr 0
-      var q
-        type-con $ptr 0
-      pat 2
-        fun $set_in 2
-        var p
-          type-con $ptr 0
-        fun $owns 2
-        var S
-          type-con $state 0
-        var q
-          type-con $ptr 0
-        fun $is_non_primitive 1
-        fun $typ 1
-        var p
-          type-con $ptr 0
-      attribute qid 1
-        string-attr VccPrelu.1133:15
-      attribute uniqueId 1
-        string-attr 124
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 3
-    fun $good_state 1
-    var S
-      type-con $state 0
-    fun $closed 2
-    var S
-      type-con $state 0
-    var q
-      type-con $ptr 0
-    fun $is_non_primitive 1
-    fun $typ 1
-    var p
-      type-con $ptr 0
-    =
-    fun $set_in 2
-    var p
-      type-con $ptr 0
-    fun $owns 2
-    var S
-      type-con $state 0
-    var q
-      type-con $ptr 0
-    =
-    fun $owner 2
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    var q
-      type-con $ptr 0
-axiom 0
-    forall 4 1 3
-      var #s1
-        type-con $state 0
-      var #s2
-        type-con $state 0
-      var #p
-        type-con $ptr 0
-      var #t
-        type-con $ctype 0
-      pat 2
-        fun $is_arraytype 1
-        var #t
-          type-con $ctype 0
-        fun $inv2 4
-        var #s1
-          type-con $state 0
-        var #s2
-          type-con $state 0
-        var #p
-          type-con $ptr 0
-        var #t
-          type-con $ctype 0
-      attribute qid 1
-        string-attr VccPrelu.1140:15
-      attribute uniqueId 1
-        string-attr 125
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 2
-    fun $is_arraytype 1
-    var #t
-      type-con $ctype 0
-    =
-    fun $typ 1
-    var #p
-      type-con $ptr 0
-    var #t
-      type-con $ctype 0
-    and 2
-    =
-    fun $inv2 4
-    var #s1
-      type-con $state 0
-    var #s2
-      type-con $state 0
-    var #p
-      type-con $ptr 0
-    var #t
-      type-con $ctype 0
-    fun $typed 2
-    var #s2
-      type-con $state 0
-    var #p
-      type-con $ptr 0
-    fun $sequential 4
-    var #s1
-      type-con $state 0
-    var #s2
-      type-con $state 0
-    var #p
-      type-con $ptr 0
-    var #t
-      type-con $ctype 0
-axiom 0
-    forall 3 1 3
-      var S
-        type-con $state 0
-      var #r
-        int
-      var #t
-        type-con $ctype 0
-      pat 2
-        fun $owns 2
-        var S
-          type-con $state 0
-        fun $ptr 2
-        var #t
-          type-con $ctype 0
-        var #r
-          int
-        fun $is_arraytype 1
-        var #t
-          type-con $ctype 0
-      attribute qid 1
-        string-attr VccPrelu.1145:15
-      attribute uniqueId 1
-        string-attr 126
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    fun $good_state 1
-    var S
-      type-con $state 0
-    implies
-    fun $is_arraytype 1
-    var #t
-      type-con $ctype 0
-    =
-    fun $owns 2
-    var S
-      type-con $state 0
-    fun $ptr 2
-    var #t
-      type-con $ctype 0
-    var #r
-      int
-    fun $set_empty 0
-axiom 0
-    forall 3 1 3
-      var S
-        type-con $state 0
-      var #p
-        type-con $ptr 0
-      var #t
-        type-con $ctype 0
-      pat 1
-        fun $inv2 4
-        var S
-          type-con $state 0
-        var S
-          type-con $state 0
-        var #p
-          type-con $ptr 0
-        var #t
-          type-con $ctype 0
-      attribute qid 1
-        string-attr VccPrelu.1149:15
-      attribute uniqueId 1
-        string-attr 127
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 2
-    fun $invok_state 1
-    var S
-      type-con $state 0
-    fun $closed 2
-    var S
-      type-con $state 0
-    var #p
-      type-con $ptr 0
-    fun $inv2 4
-    var S
-      type-con $state 0
-    var S
-      type-con $state 0
-    var #p
-      type-con $ptr 0
-    var #t
-      type-con $ctype 0
-axiom 0
-    forall 1 1 3
-      var S
-        type-con $state 0
-      pat 1
-        fun $good_state 1
-        var S
-          type-con $state 0
-      attribute qid 1
-        string-attr VccPrelu.1152:15
-      attribute uniqueId 1
-        string-attr 128
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    fun $good_state 1
-    var S
-      type-con $state 0
-    forall 2 1 3
-      var #p
-        type-con $ptr 0
-      var #q
-        type-con $ptr 0
-      pat 1
-        fun $set_in 2
-        var #p
-          type-con $ptr 0
-        fun $owns 2
-        var S
-          type-con $state 0
-        var #q
-          type-con $ptr 0
-      attribute qid 1
-        string-attr VccPrelu.846:13
-      attribute uniqueId 1
-        string-attr 86
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 3
-    fun $good_state 1
-    var S
-      type-con $state 0
-    fun $set_in 2
-    var #p
-      type-con $ptr 0
-    fun $owns 2
-    var S
-      type-con $state 0
-    var #q
-      type-con $ptr 0
-    fun $closed 2
-    var S
-      type-con $state 0
-    var #q
-      type-con $ptr 0
-    and 2
-    fun $closed 2
-    var S
-      type-con $state 0
-    var #p
-      type-con $ptr 0
-    not
-    =
-    fun $ref 1
-    var #p
-      type-con $ptr 0
-    int-num 0
-axiom 0
-    forall 3 1 3
-      var S
-        type-con $state 0
-      var p
-        type-con $ptr 0
-      var v
-        int
-      pat 1
-        fun $update_int 3
-        var S
-          type-con $state 0
-        var p
-          type-con $ptr 0
-        var v
-          int
-      attribute qid 1
-        string-attr VccPrelu.1260:15
-      attribute uniqueId 1
-        string-attr 138
-      attribute bvZ3Native 1
-        string-attr False
-    and 6
-    =
-    fun $typemap 1
-    fun $update_int 3
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    var v
-      int
-    fun $typemap 1
-    var S
-      type-con $state 0
-    =
-    fun $statusmap 1
-    fun $update_int 3
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    var v
-      int
-    fun $statusmap 1
-    var S
-      type-con $state 0
-    =
-    fun $memory 1
-    fun $update_int 3
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    var v
-      int
-    fun $store.mem 3
-    fun $memory 1
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    var v
-      int
-    <
-    fun $current_timestamp 1
-    var S
-      type-con $state 0
-    fun $current_timestamp 1
-    fun $update_int 3
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    var v
-      int
-    forall 1 1 4
-      var p
-        type-con $ptr 0
-      pat 1
-        fun $timestamp 2
-        fun $update_int 3
-        var S
-          type-con $state 0
-        var p
-          type-con $ptr 0
-        var v
-          int
-        var p
-          type-con $ptr 0
-      attribute qid 1
-        string-attr VccPrelu.1280:13
-      attribute uniqueId 1
-        string-attr 140
-      attribute bvZ3Native 1
-        string-attr False
-      attribute weight 1
-        expr-attr
-          int-num 0
-    <=
-    fun $timestamp 2
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    fun $timestamp 2
-    fun $update_int 3
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    var v
-      int
-    var p
-      type-con $ptr 0
-    fun $call_transition 2
-    var S
-      type-con $state 0
-    fun $update_int 3
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    var v
-      int
-axiom 0
-    forall 3 1 3
-      var S
-        type-con $state 0
-      var l
-        type-con $ptr 0
-      var p
-        type-con $ptr 0
-      pat 1
-        fun $take_over 3
-        var S
-          type-con $state 0
-        var l
-          type-con $ptr 0
-        var p
-          type-con $ptr 0
-      attribute qid 1
-        string-attr VccPrelu.1309:15
-      attribute uniqueId 1
-        string-attr 141
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    not
-    =
-    fun $kind_of 1
-    fun $typ 1
-    var l
-      type-con $ptr 0
-    fun $kind_primitive 0
-    and 5
-    =
-    fun $statusmap 1
-    fun $take_over 3
-    var S
-      type-con $state 0
-    var l
-      type-con $ptr 0
-    var p
-      type-con $ptr 0
-    fun $store.sm 3
-    fun $statusmap 1
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    fun $taken_over 3
-    var S
-      type-con $state 0
-    var l
-      type-con $ptr 0
-    var p
-      type-con $ptr 0
-    fun $closed 2
-    fun $take_over 3
-    var S
-      type-con $state 0
-    var l
-      type-con $ptr 0
-    var p
-      type-con $ptr 0
-    var p
-      type-con $ptr 0
-    =
-    fun $owner 2
-    fun $take_over 3
-    var S
-      type-con $state 0
-    var l
-      type-con $ptr 0
-    var p
-      type-con $ptr 0
-    var p
-      type-con $ptr 0
-    var l
-      type-con $ptr 0
-    =
-    fun $ref_cnt 2
-    fun $take_over 3
-    var S
-      type-con $state 0
-    var l
-      type-con $ptr 0
-    var p
-      type-con $ptr 0
-    var p
-      type-con $ptr 0
-    fun $ref_cnt 2
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    true
-axiom 0
-    forall 4 1 3
-      var S0
-        type-con $state 0
-      var S
-        type-con $state 0
-      var l
-        type-con $ptr 0
-      var p
-        type-con $ptr 0
-      pat 1
-        fun $release 4
-        var S0
-          type-con $state 0
-        var S
-          type-con $state 0
-        var l
-          type-con $ptr 0
-        var p
-          type-con $ptr 0
-      attribute qid 1
-        string-attr VccPrelu.1325:15
-      attribute uniqueId 1
-        string-attr 142
-      attribute bvZ3Native 1
-        string-attr False
-    and 6
-    =
-    fun $statusmap 1
-    fun $release 4
-    var S0
-      type-con $state 0
-    var S
-      type-con $state 0
-    var l
-      type-con $ptr 0
-    var p
-      type-con $ptr 0
-    fun $store.sm 3
-    fun $statusmap 1
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    fun $released 3
-    var S
-      type-con $state 0
-    var l
-      type-con $ptr 0
-    var p
-      type-con $ptr 0
-    fun $closed 2
-    fun $release 4
-    var S0
-      type-con $state 0
-    var S
-      type-con $state 0
-    var l
-      type-con $ptr 0
-    var p
-      type-con $ptr 0
-    var p
-      type-con $ptr 0
-    =
-    fun $owner 2
-    fun $release 4
-    var S0
-      type-con $state 0
-    var S
-      type-con $state 0
-    var l
-      type-con $ptr 0
-    var p
-      type-con $ptr 0
-    var p
-      type-con $ptr 0
-    fun $me 0
-    =
-    fun $ref_cnt 2
-    fun $release 4
-    var S0
-      type-con $state 0
-    var S
-      type-con $state 0
-    var l
-      type-con $ptr 0
-    var p
-      type-con $ptr 0
-    var p
-      type-con $ptr 0
-    fun $ref_cnt 2
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    =
-    fun $timestamp 2
-    fun $release 4
-    var S0
-      type-con $state 0
-    var S
-      type-con $state 0
-    var l
-      type-con $ptr 0
-    var p
-      type-con $ptr 0
-    var p
-      type-con $ptr 0
-    fun $current_timestamp 1
-    var S0
-      type-con $state 0
-    true
-axiom 0
-    =
-    fun $get_memory_allocator 0
-    fun $ptr 2
-    fun $memory_allocator_type 0
-    fun $memory_allocator_ref 0
-axiom 0
-    =
-    fun $ptr_level 1
-    fun $memory_allocator_type 0
-    int-num 0
-axiom 0
-    forall 1 1 3
-      var S
-        type-con $state 0
-      pat 1
-        fun $program_entry_point 1
-        var S
-          type-con $state 0
-      attribute qid 1
-        string-attr VccPrelu.1661:15
-      attribute uniqueId 1
-        string-attr 175
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    fun $program_entry_point 1
-    var S
-      type-con $state 0
-    fun $program_entry_point_ch 1
-    var S
-      type-con $state 0
-axiom 0
-    forall 3 1 3
-      var S
-        type-con $state 0
-      var p
-        type-con $ptr 0
-      var q
-        type-con $ptr 0
-      pat 1
-        fun $set_in 2
-        var p
-          type-con $ptr 0
-        fun $volatile_span 2
-        var S
-          type-con $state 0
-        var q
-          type-con $ptr 0
-      attribute qid 1
-        string-attr VccPrelu.1745:15
-      attribute uniqueId 1
-        string-attr 186
-      attribute bvZ3Native 1
-        string-attr False
-    =
-    fun $set_in 2
-    var p
-      type-con $ptr 0
-    fun $volatile_span 2
-    var S
-      type-con $state 0
-    var q
-      type-con $ptr 0
-    or 2
-    =
-    var p
-      type-con $ptr 0
-    var q
-      type-con $ptr 0
-    and 2
-    fun $ts_is_volatile 1
-    fun $select.tm 2
-    fun $typemap 1
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    fun $set_in 2
-    var p
-      type-con $ptr 0
-    fun $span 1
-    var q
-      type-con $ptr 0
-axiom 0
-    forall 2 1 3
-      var a
-        type-con $ptr 0
-      var i
-        int
-      pat 1
-        fun $left_split 2
-        var a
-          type-con $ptr 0
-        var i
-          int
-      attribute qid 1
-        string-attr VccPrelu.1752:22
-      attribute uniqueId 1
-        string-attr 187
-      attribute bvZ3Native 1
-        string-attr False
-    =
-    fun $left_split 2
-    var a
-      type-con $ptr 0
-    var i
-      int
-    fun $ptr 2
-    fun $array 2
-    fun $element_type 1
-    fun $typ 1
-    var a
-      type-con $ptr 0
-    var i
-      int
-    fun $ref 1
-    var a
-      type-con $ptr 0
-axiom 0
-    forall 2 1 3
-      var a
-        type-con $ptr 0
-      var i
-        int
-      pat 1
-        fun $right_split 2
-        var a
-          type-con $ptr 0
-        var i
-          int
-      attribute qid 1
-        string-attr VccPrelu.1754:23
-      attribute uniqueId 1
-        string-attr 188
-      attribute bvZ3Native 1
-        string-attr False
-    =
-    fun $right_split 2
-    var a
-      type-con $ptr 0
-    var i
-      int
-    fun $ptr 2
-    fun $array 2
-    fun $element_type 1
-    fun $typ 1
-    var a
-      type-con $ptr 0
-    -
-    fun $array_length 1
-    fun $typ 1
-    var a
-      type-con $ptr 0
-    var i
-      int
-    fun $ref 1
-    fun $idx 3
-    fun $ptr 2
-    fun $element_type 1
-    fun $typ 1
-    var a
-      type-con $ptr 0
-    fun $ref 1
-    var a
-      type-con $ptr 0
-    var i
-      int
-    fun $element_type 1
-    fun $typ 1
-    var a
-      type-con $ptr 0
-axiom 0
-    forall 2 1 3
-      var a1
-        type-con $ptr 0
-      var a2
-        type-con $ptr 0
-      pat 1
-        fun $joined_array 2
-        var a1
-          type-con $ptr 0
-        var a2
-          type-con $ptr 0
-      attribute qid 1
-        string-attr VccPrelu.1757:24
-      attribute uniqueId 1
-        string-attr 189
-      attribute bvZ3Native 1
-        string-attr False
-    =
-    fun $joined_array 2
-    var a1
-      type-con $ptr 0
-    var a2
-      type-con $ptr 0
-    fun $ptr 2
-    fun $array 2
-    fun $element_type 1
-    fun $typ 1
-    var a1
-      type-con $ptr 0
-    +
-    fun $array_length 1
-    fun $typ 1
-    var a1
-      type-con $ptr 0
-    fun $array_length 1
-    fun $typ 1
-    var a2
-      type-con $ptr 0
-    fun $ref 1
-    var a1
-      type-con $ptr 0
-axiom 0
-    forall 1 1 4
-      var #o
-        type-con $ptr 0
-      pat 1
-        fun $set_in 2
-        var #o
-          type-con $ptr 0
-        fun $set_empty 0
-      attribute qid 1
-        string-attr VccPrelu.1854:15
-      attribute uniqueId 1
-        string-attr 198
-      attribute bvZ3Native 1
-        string-attr False
-      attribute weight 1
-        expr-attr
-          int-num 0
-    not
-    fun $set_in 2
-    var #o
-      type-con $ptr 0
-    fun $set_empty 0
-axiom 0
-    forall 2 1 4
-      var #r
-        type-con $ptr 0
-      var #o
-        type-con $ptr 0
-      pat 1
-        fun $set_in 2
-        var #o
-          type-con $ptr 0
-        fun $set_singleton 1
-        var #r
-          type-con $ptr 0
-      attribute qid 1
-        string-attr VccPrelu.1857:15
-      attribute uniqueId 1
-        string-attr 199
-      attribute bvZ3Native 1
-        string-attr False
-      attribute weight 1
-        expr-attr
-          int-num 0
-    =
-    fun $set_in 2
-    var #o
-      type-con $ptr 0
-    fun $set_singleton 1
-    var #r
-      type-con $ptr 0
-    =
-    var #r
-      type-con $ptr 0
-    var #o
-      type-con $ptr 0
-axiom 0
-    forall 2 1 4
-      var #r
-        type-con $ptr 0
-      var #o
-        type-con $ptr 0
-      pat 1
-        fun $set_in 2
-        var #o
-          type-con $ptr 0
-        fun $non_null_set_singleton 1
-        var #r
-          type-con $ptr 0
-      attribute qid 1
-        string-attr VccPrelu.1860:15
-      attribute uniqueId 1
-        string-attr 200
-      attribute bvZ3Native 1
-        string-attr False
-      attribute weight 1
-        expr-attr
-          int-num 0
-    =
-    fun $set_in 2
-    var #o
-      type-con $ptr 0
-    fun $non_null_set_singleton 1
-    var #r
-      type-con $ptr 0
-    and 2
-    =
-    var #r
-      type-con $ptr 0
-    var #o
-      type-con $ptr 0
-    not
-    =
-    fun $ref 1
-    var #r
-      type-con $ptr 0
-    fun $ref 1
-    fun $null 0
-axiom 0
-    forall 3 1 4
-      var #a
-        type-con $ptrset 0
-      var #b
-        type-con $ptrset 0
-      var #o
-        type-con $ptr 0
-      pat 1
-        fun $set_in 2
-        var #o
-          type-con $ptr 0
-        fun $set_union 2
-        var #a
-          type-con $ptrset 0
-        var #b
-          type-con $ptrset 0
-      attribute qid 1
-        string-attr VccPrelu.1863:15
-      attribute uniqueId 1
-        string-attr 201
-      attribute bvZ3Native 1
-        string-attr False
-      attribute weight 1
-        expr-attr
-          int-num 0
-    =
-    fun $set_in 2
-    var #o
-      type-con $ptr 0
-    fun $set_union 2
-    var #a
-      type-con $ptrset 0
-    var #b
-      type-con $ptrset 0
-    or 2
-    fun $set_in 2
-    var #o
-      type-con $ptr 0
-    var #a
-      type-con $ptrset 0
-    fun $set_in 2
-    var #o
-      type-con $ptr 0
-    var #b
-      type-con $ptrset 0
-axiom 0
-    forall 3 1 4
-      var #a
-        type-con $ptrset 0
-      var #b
-        type-con $ptrset 0
-      var #o
-        type-con $ptr 0
-      pat 1
-        fun $set_in 2
-        var #o
-          type-con $ptr 0
-        fun $set_difference 2
-        var #a
-          type-con $ptrset 0
-        var #b
-          type-con $ptrset 0
-      attribute qid 1
-        string-attr VccPrelu.1867:15
-      attribute uniqueId 1
-        string-attr 202
-      attribute bvZ3Native 1
-        string-attr False
-      attribute weight 1
-        expr-attr
-          int-num 0
-    =
-    fun $set_in 2
-    var #o
-      type-con $ptr 0
-    fun $set_difference 2
-    var #a
-      type-con $ptrset 0
-    var #b
-      type-con $ptrset 0
-    and 2
-    fun $set_in 2
-    var #o
-      type-con $ptr 0
-    var #a
-      type-con $ptrset 0
-    not
-    fun $set_in 2
-    var #o
-      type-con $ptr 0
-    var #b
-      type-con $ptrset 0
-axiom 0
-    forall 3 1 4
-      var #a
-        type-con $ptrset 0
-      var #b
-        type-con $ptrset 0
-      var #o
-        type-con $ptr 0
-      pat 1
-        fun $set_in 2
-        var #o
-          type-con $ptr 0
-        fun $set_intersection 2
-        var #a
-          type-con $ptrset 0
-        var #b
-          type-con $ptrset 0
-      attribute qid 1
-        string-attr VccPrelu.1871:15
-      attribute uniqueId 1
-        string-attr 203
-      attribute bvZ3Native 1
-        string-attr False
-      attribute weight 1
-        expr-attr
-          int-num 0
-    =
-    fun $set_in 2
-    var #o
-      type-con $ptr 0
-    fun $set_intersection 2
-    var #a
-      type-con $ptrset 0
-    var #b
-      type-con $ptrset 0
-    and 2
-    fun $set_in 2
-    var #o
-      type-con $ptr 0
-    var #a
-      type-con $ptrset 0
-    fun $set_in 2
-    var #o
-      type-con $ptr 0
-    var #b
-      type-con $ptrset 0
-axiom 0
-    forall 2 1 4
-      var #a
-        type-con $ptrset 0
-      var #b
-        type-con $ptrset 0
-      pat 1
-        fun $set_subset 2
-        var #a
-          type-con $ptrset 0
-        var #b
-          type-con $ptrset 0
-      attribute qid 1
-        string-attr VccPrelu.1875:14
-      attribute uniqueId 1
-        string-attr 205
-      attribute bvZ3Native 1
-        string-attr False
-      attribute weight 1
-        expr-attr
-          int-num 0
-    =
-    fun $set_subset 2
-    var #a
-      type-con $ptrset 0
-    var #b
-      type-con $ptrset 0
-    forall 1 2 4
-      var #o
-        type-con $ptr 0
-      pat 1
-        fun $set_in 2
-        var #o
-          type-con $ptr 0
-        var #a
-          type-con $ptrset 0
-      pat 1
-        fun $set_in 2
-        var #o
-          type-con $ptr 0
-        var #b
-          type-con $ptrset 0
-      attribute qid 1
-        string-attr VccPrelu.1876:35
-      attribute uniqueId 1
-        string-attr 204
-      attribute bvZ3Native 1
-        string-attr False
-      attribute weight 1
-        expr-attr
-          int-num 0
-    implies
-    fun $set_in 2
-    var #o
-      type-con $ptr 0
-    var #a
-      type-con $ptrset 0
-    fun $set_in 2
-    var #o
-      type-con $ptr 0
-    var #b
-      type-con $ptrset 0
-axiom 0
-    forall 2 1 4
-      var #a
-        type-con $ptrset 0
-      var #b
-        type-con $ptrset 0
-      pat 1
-        fun $set_eq 2
-        var #a
-          type-con $ptrset 0
-        var #b
-          type-con $ptrset 0
-      attribute qid 1
-        string-attr VccPrelu.1880:15
-      attribute uniqueId 1
-        string-attr 207
-      attribute bvZ3Native 1
-        string-attr False
-      attribute weight 1
-        expr-attr
-          int-num 0
-    implies
-    forall 1 1 4
-      var #o
-        type-con $ptr 0
-      pat 1
-        fun $dont_instantiate 1
-        var #o
-          type-con $ptr 0
-      attribute qid 1
-        string-attr VccPrelu.1881:11
-      attribute uniqueId 1
-        string-attr 206
-      attribute bvZ3Native 1
-        string-attr False
-      attribute weight 1
-        expr-attr
-          int-num 0
-    =
-    fun $set_in 2
-    var #o
-      type-con $ptr 0
-    var #a
-      type-con $ptrset 0
-    fun $set_in 2
-    var #o
-      type-con $ptr 0
-    var #b
-      type-con $ptrset 0
-    fun $set_eq 2
-    var #a
-      type-con $ptrset 0
-    var #b
-      type-con $ptrset 0
-axiom 0
-    forall 2 1 4
-      var #a
-        type-con $ptrset 0
-      var #b
-        type-con $ptrset 0
-      pat 1
-        fun $set_eq 2
-        var #a
-          type-con $ptrset 0
-        var #b
-          type-con $ptrset 0
-      attribute qid 1
-        string-attr VccPrelu.1882:15
-      attribute uniqueId 1
-        string-attr 208
-      attribute bvZ3Native 1
-        string-attr False
-      attribute weight 1
-        expr-attr
-          int-num 0
-    implies
-    fun $set_eq 2
-    var #a
-      type-con $ptrset 0
-    var #b
-      type-con $ptrset 0
-    =
-    var #a
-      type-con $ptrset 0
-    var #b
-      type-con $ptrset 0
-axiom 0
-    =
-    fun $set_cardinality 1
-    fun $set_empty 0
-    int-num 0
-axiom 0
-    forall 1 0 4
-      var p
-        type-con $ptr 0
-      attribute qid 1
-        string-attr VccPrelu.1888:15
-      attribute uniqueId 1
-        string-attr 209
-      attribute bvZ3Native 1
-        string-attr False
-      attribute weight 1
-        expr-attr
-          int-num 0
-    =
-    fun $set_cardinality 1
-    fun $set_singleton 1
-    var p
-      type-con $ptr 0
-    int-num 1
-axiom 0
-    forall 1 1 4
-      var #o
-        type-con $ptr 0
-      pat 1
-        fun $set_in 2
-        var #o
-          type-con $ptr 0
-        fun $set_universe 0
-      attribute qid 1
-        string-attr VccPrelu.1891:15
-      attribute uniqueId 1
-        string-attr 210
-      attribute bvZ3Native 1
-        string-attr False
-      attribute weight 1
-        expr-attr
-          int-num 0
-    fun $set_in 2
-    var #o
-      type-con $ptr 0
-    fun $set_universe 0
-axiom 0
-    forall 3 1 4
-      var p
-        type-con $ptr 0
-      var s1
-        type-con $ptrset 0
-      var s2
-        type-con $ptrset 0
-      pat 2
-        fun $set_disjoint 2
-        var s1
-          type-con $ptrset 0
-        var s2
-          type-con $ptrset 0
-        fun $set_in 2
-        var p
-          type-con $ptr 0
-        var s1
-          type-con $ptrset 0
-      attribute qid 1
-        string-attr VccPrelu.1896:15
-      attribute uniqueId 1
-        string-attr 211
-      attribute bvZ3Native 1
-        string-attr False
-      attribute weight 1
-        expr-attr
-          int-num 0
-    implies
-    and 2
-    fun $set_disjoint 2
-    var s1
-      type-con $ptrset 0
-    var s2
-      type-con $ptrset 0
-    fun $set_in 2
-    var p
-      type-con $ptr 0
-    var s1
-      type-con $ptrset 0
-    =
-    fun $id_set_disjoint 3
-    var p
-      type-con $ptr 0
-    var s1
-      type-con $ptrset 0
-    var s2
-      type-con $ptrset 0
-    int-num 1
-axiom 0
-    forall 3 1 4
-      var p
-        type-con $ptr 0
-      var s1
-        type-con $ptrset 0
-      var s2
-        type-con $ptrset 0
-      pat 2
-        fun $set_disjoint 2
-        var s1
-          type-con $ptrset 0
-        var s2
-          type-con $ptrset 0
-        fun $set_in 2
-        var p
-          type-con $ptr 0
-        var s2
-          type-con $ptrset 0
-      attribute qid 1
-        string-attr VccPrelu.1899:15
-      attribute uniqueId 1
-        string-attr 212
-      attribute bvZ3Native 1
-        string-attr False
-      attribute weight 1
-        expr-attr
-          int-num 0
-    implies
-    and 2
-    fun $set_disjoint 2
-    var s1
-      type-con $ptrset 0
-    var s2
-      type-con $ptrset 0
-    fun $set_in 2
-    var p
-      type-con $ptr 0
-    var s2
-      type-con $ptrset 0
-    =
-    fun $id_set_disjoint 3
-    var p
-      type-con $ptr 0
-    var s1
-      type-con $ptrset 0
-    var s2
-      type-con $ptrset 0
-    int-num 2
-axiom 0
-    forall 2 1 4
-      var s1
-        type-con $ptrset 0
-      var s2
-        type-con $ptrset 0
-      pat 1
-        fun $set_disjoint 2
-        var s1
-          type-con $ptrset 0
-        var s2
-          type-con $ptrset 0
-      attribute qid 1
-        string-attr VccPrelu.1903:15
-      attribute uniqueId 1
-        string-attr 214
-      attribute bvZ3Native 1
-        string-attr False
-      attribute weight 1
-        expr-attr
-          int-num 0
-    implies
-    forall 1 1 3
-      var p
-        type-con $ptr 0
-      pat 1
-        fun $dont_instantiate 1
-        var p
-          type-con $ptr 0
-      attribute qid 1
-        string-attr VccPrelu.1904:11
-      attribute uniqueId 1
-        string-attr 213
-      attribute bvZ3Native 1
-        string-attr False
-    and 2
-    implies
-    fun $set_in 2
-    var p
-      type-con $ptr 0
-    var s1
-      type-con $ptrset 0
-    not
-    fun $set_in 2
-    var p
-      type-con $ptr 0
-    var s2
-      type-con $ptrset 0
-    implies
-    fun $set_in 2
-    var p
-      type-con $ptr 0
-    var s2
-      type-con $ptrset 0
-    not
-    fun $set_in 2
-    var p
-      type-con $ptr 0
-    var s1
-      type-con $ptrset 0
-    fun $set_disjoint 2
-    var s1
-      type-con $ptrset 0
-    var s2
-      type-con $ptrset 0
-axiom 0
-    forall 3 1 4
-      var p
-        type-con $ptr 0
-      var S1
-        type-con $state 0
-      var p1
-        type-con $ptr 0
-      pat 1
-        fun $set_in 2
-        var p
-          type-con $ptr 0
-        fun $owns 2
-        var S1
-          type-con $state 0
-        var p1
-          type-con $ptr 0
-      attribute qid 1
-        string-attr VccPrelu.1914:15
-      attribute uniqueId 1
-        string-attr 215
-      attribute bvZ3Native 1
-        string-attr False
-      attribute weight 1
-        expr-attr
-          int-num 0
-    implies
-    fun $set_in 2
-    var p
-      type-con $ptr 0
-    fun $owns 2
-    var S1
-      type-con $state 0
-    var p1
-      type-con $ptr 0
-    fun $in_some_owns 1
-    var p
-      type-con $ptr 0
-axiom 0
-    forall 3 1 4
-      var p
-        type-con $ptr 0
-      var S1
-        type-con $state 0
-      var p1
-        type-con $ptr 0
-      pat 2
-        fun $set_in2 2
-        var p
-          type-con $ptr 0
-        fun $owns 2
-        var S1
-          type-con $state 0
-        var p1
-          type-con $ptr 0
-        fun $in_some_owns 1
-        var p
-          type-con $ptr 0
-      attribute qid 1
-        string-attr VccPrelu.1918:15
-      attribute uniqueId 1
-        string-attr 216
-      attribute bvZ3Native 1
-        string-attr False
-      attribute weight 1
-        expr-attr
-          int-num 0
-    =
-    fun $set_in 2
-    var p
-      type-con $ptr 0
-    fun $owns 2
-    var S1
-      type-con $state 0
-    var p1
-      type-con $ptr 0
-    fun $set_in2 2
-    var p
-      type-con $ptr 0
-    fun $owns 2
-    var S1
-      type-con $state 0
-    var p1
-      type-con $ptr 0
-axiom 0
-    forall 2 1 4
-      var p
-        type-con $ptr 0
-      var s
-        type-con $ptrset 0
-      pat 1
-        fun $set_in 2
-        var p
-          type-con $ptr 0
-        var s
-          type-con $ptrset 0
-      attribute qid 1
-        string-attr VccPrelu.1922:15
-      attribute uniqueId 1
-        string-attr 217
-      attribute bvZ3Native 1
-        string-attr False
-      attribute weight 1
-        expr-attr
-          int-num 0
-    =
-    fun $set_in 2
-    var p
-      type-con $ptr 0
-    var s
-      type-con $ptrset 0
-    fun $set_in2 2
-    var p
-      type-con $ptr 0
-    var s
-      type-con $ptrset 0
-axiom 0
-    forall 2 1 4
-      var p
-        type-con $ptr 0
-      var s
-        type-con $ptrset 0
-      pat 1
-        fun $set_in 2
-        var p
-          type-con $ptr 0
-        var s
-          type-con $ptrset 0
-      attribute qid 1
-        string-attr VccPrelu.1924:15
-      attribute uniqueId 1
-        string-attr 218
-      attribute bvZ3Native 1
-        string-attr False
-      attribute weight 1
-        expr-attr
-          int-num 0
-    =
-    fun $set_in 2
-    var p
-      type-con $ptr 0
-    var s
-      type-con $ptrset 0
-    fun $set_in3 2
-    var p
-      type-con $ptr 0
-    var s
-      type-con $ptrset 0
-axiom 0
-    forall 2 1 4
-      var p
-        type-con $ptr 0
-      var s
-        type-con $ptrset 0
-      pat 1
-        fun $set_in0 2
-        var p
-          type-con $ptr 0
-        var s
-          type-con $ptrset 0
-      attribute qid 1
-        string-attr VccPrelu.1928:15
-      attribute uniqueId 1
-        string-attr 219
-      attribute bvZ3Native 1
-        string-attr False
-      attribute weight 1
-        expr-attr
-          int-num 0
-    =
-    fun $set_in 2
-    var p
-      type-con $ptr 0
-    var s
-      type-con $ptrset 0
-    fun $set_in0 2
-    var p
-      type-con $ptr 0
-    var s
-      type-con $ptrset 0
-axiom 0
-    forall 2 1 3
-      var T
-        type-con $ctype 0
-      var s
-        int
-      pat 1
-        fun $array 2
-        var T
-          type-con $ctype 0
-        var s
-          int
-      attribute qid 1
-        string-attr VccPrelu.1989:15
-      attribute uniqueId 1
-        string-attr 224
-      attribute bvZ3Native 1
-        string-attr False
-    =
-    fun $element_type 1
-    fun $array 2
-    var T
-      type-con $ctype 0
-    var s
-      int
-    var T
-      type-con $ctype 0
-axiom 0
-    forall 2 1 3
-      var T
-        type-con $ctype 0
-      var s
-        int
-      pat 1
-        fun $array 2
-        var T
-          type-con $ctype 0
-        var s
-          int
-      attribute qid 1
-        string-attr VccPrelu.1990:15
-      attribute uniqueId 1
-        string-attr 225
-      attribute bvZ3Native 1
-        string-attr False
-    =
-    fun $array_length 1
-    fun $array 2
-    var T
-      type-con $ctype 0
-    var s
-      int
-    var s
-      int
-axiom 0
-    forall 2 1 3
-      var T
-        type-con $ctype 0
-      var s
-        int
-      pat 1
-        fun $array 2
-        var T
-          type-con $ctype 0
-        var s
-          int
-      attribute qid 1
-        string-attr VccPrelu.1991:15
-      attribute uniqueId 1
-        string-attr 226
-      attribute bvZ3Native 1
-        string-attr False
-    =
-    fun $ptr_level 1
-    fun $array 2
-    var T
-      type-con $ctype 0
-    var s
-      int
-    int-num 0
-axiom 0
-    forall 2 1 3
-      var T
-        type-con $ctype 0
-      var s
-        int
-      pat 1
-        fun $array 2
-        var T
-          type-con $ctype 0
-        var s
-          int
-      attribute qid 1
-        string-attr VccPrelu.1992:15
-      attribute uniqueId 1
-        string-attr 227
-      attribute bvZ3Native 1
-        string-attr False
-    fun $is_arraytype 1
-    fun $array 2
-    var T
-      type-con $ctype 0
-    var s
-      int
-axiom 0
-    forall 2 1 3
-      var T
-        type-con $ctype 0
-      var s
-        int
-      pat 1
-        fun $array 2
-        var T
-          type-con $ctype 0
-        var s
-          int
-      attribute qid 1
-        string-attr VccPrelu.1993:15
-      attribute uniqueId 1
-        string-attr 228
-      attribute bvZ3Native 1
-        string-attr False
-    not
-    fun $is_claimable 1
-    fun $array 2
-    var T
-      type-con $ctype 0
-    var s
-      int
-axiom 0
-    forall 2 1 4
-      var p
-        type-con $ptr 0
-      var T
-        type-con $ctype 0
-      pat 1
-        fun $inlined_array 2
-        var p
-          type-con $ptr 0
-        var T
-          type-con $ctype 0
-      attribute qid 1
-        string-attr VccPrelu.1998:37
-      attribute uniqueId 1
-        string-attr 229
-      attribute bvZ3Native 1
-        string-attr False
-      attribute weight 1
-        expr-attr
-          int-num 0
-    =
-    fun $inlined_array 2
-    var p
-      type-con $ptr 0
-    var T
-      type-con $ctype 0
-    var p
-      type-con $ptr 0
-axiom 0
-    forall 3 1 3
-      var #p
-        type-con $ptr 0
-      var #i
-        int
-      var #t
-        type-con $ctype 0
-      pat 1
-        fun $idx 3
-        var #p
-          type-con $ptr 0
-        var #i
-          int
-        var #t
-          type-con $ctype 0
-      attribute qid 1
-        string-attr VccPrelu.2002:15
-      attribute uniqueId 1
-        string-attr 230
-      attribute bvZ3Native 1
-        string-attr False
-    and 2
-    fun $extent_hint 2
-    fun $idx 3
-    var #p
-      type-con $ptr 0
-    var #i
-      int
-    var #t
-      type-con $ctype 0
-    var #p
-      type-con $ptr 0
-    =
-    fun $idx 3
-    var #p
-      type-con $ptr 0
-    var #i
-      int
-    var #t
-      type-con $ctype 0
-    fun $ptr 2
-    var #t
-      type-con $ctype 0
-    +
-    fun $ref 1
-    var #p
-      type-con $ptr 0
-    *
-    var #i
-      int
-    fun $sizeof 1
-    var #t
-      type-con $ctype 0
-axiom 0
-    forall 4 1 3
-      var p
-        type-con $ptr 0
-      var i
-        int
-      var j
-        int
-      var T
-        type-con $ctype 0
-      pat 1
-        fun $idx 3
-        fun $idx 3
-        var p
-          type-con $ptr 0
-        var i
-          int
-        var T
-          type-con $ctype 0
-        var j
-          int
-        var T
-          type-con $ctype 0
-      attribute qid 1
-        string-attr VccPrelu.2016:15
-      attribute uniqueId 1
-        string-attr 231
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 2
-    not
-    =
-    var i
-      int
-    int-num 0
-    not
-    =
-    var j
-      int
-    int-num 0
-    =
-    fun $idx 3
-    fun $idx 3
-    var p
-      type-con $ptr 0
-    var i
-      int
-    var T
-      type-con $ctype 0
-    var j
-      int
-    var T
-      type-con $ctype 0
-    fun $idx 3
-    var p
-      type-con $ptr 0
-    +
-    var i
-      int
-    var j
-      int
-    var T
-      type-con $ctype 0
-axiom 0
-    forall 5 1 4
-      var S
-        type-con $state 0
-      var p
-        type-con $ptr 0
-      var T
-        type-con $ctype 0
-      var sz
-        int
-      var vol
-        bool
-      pat 1
-        fun $is_array_vol_or_nonvol 5
-        var S
-          type-con $state 0
-        var p
-          type-con $ptr 0
-        var T
-          type-con $ctype 0
-        var sz
-          int
-        var vol
-          bool
-      attribute qid 1
-        string-attr VccPrelu.2020:46
-      attribute uniqueId 1
-        string-attr 233
-      attribute bvZ3Native 1
-        string-attr False
-      attribute weight 1
-        expr-attr
-          int-num 0
-    =
-    fun $is_array_vol_or_nonvol 5
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    var T
-      type-con $ctype 0
-    var sz
-      int
-    var vol
-      bool
-    and 2
-    fun $is 2
-    var p
-      type-con $ptr 0
-    var T
-      type-con $ctype 0
-    forall 1 3 3
-      var i
-        int
-      pat 1
-        fun $select.sm 2
-        fun $statusmap 1
-        var S
-          type-con $state 0
-        fun $idx 3
-        var p
-          type-con $ptr 0
-        var i
-          int
-        var T
-          type-con $ctype 0
-      pat 1
-        fun $select.tm 2
-        fun $typemap 1
-        var S
-          type-con $state 0
-        fun $idx 3
-        var p
-          type-con $ptr 0
-        var i
-          int
-        var T
-          type-con $ctype 0
-      pat 1
-        fun $select.mem 2
-        fun $memory 1
-        var S
-          type-con $state 0
-        fun $idx 3
-        var p
-          type-con $ptr 0
-        var i
-          int
-        var T
-          type-con $ctype 0
-      attribute qid 1
-        string-attr VccPrelu.2022:13
-      attribute uniqueId 1
-        string-attr 232
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 2
-    <=
-    int-num 0
-    var i
-      int
-    <
-    var i
-      int
-    var sz
-      int
-    and 3
-    =
-    fun $ts_is_volatile 1
-    fun $select.tm 2
-    fun $typemap 1
-    var S
-      type-con $state 0
-    fun $idx 3
-    var p
-      type-con $ptr 0
-    var i
-      int
-    var T
-      type-con $ctype 0
-    var vol
-      bool
-    fun $ts_is_array_elt 1
-    fun $select.tm 2
-    fun $typemap 1
-    var S
-      type-con $state 0
-    fun $idx 3
-    var p
-      type-con $ptr 0
-    var i
-      int
-    var T
-      type-con $ctype 0
-    fun $typed 2
-    var S
-      type-con $state 0
-    fun $idx 3
-    var p
-      type-con $ptr 0
-    var i
-      int
-    var T
-      type-con $ctype 0
-axiom 0
-    forall 4 1 4
-      var S
-        type-con $state 0
-      var p
-        type-con $ptr 0
-      var T
-        type-con $ctype 0
-      var sz
-        int
-      pat 1
-        fun $is_array 4
-        var S
-          type-con $state 0
-        var p
-          type-con $ptr 0
-        var T
-          type-con $ctype 0
-        var sz
-          int
-      attribute qid 1
-        string-attr VccPrelu.2026:32
-      attribute uniqueId 1
-        string-attr 235
-      attribute bvZ3Native 1
-        string-attr False
-      attribute weight 1
-        expr-attr
-          int-num 0
-    =
-    fun $is_array 4
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    var T
-      type-con $ctype 0
-    var sz
-      int
-    and 2
-    fun $is 2
-    var p
-      type-con $ptr 0
-    var T
-      type-con $ctype 0
-    forall 1 3 3
-      var i
-        int
-      pat 1
-        fun $select.sm 2
-        fun $statusmap 1
-        var S
-          type-con $state 0
-        fun $idx 3
-        var p
-          type-con $ptr 0
-        var i
-          int
-        var T
-          type-con $ctype 0
-      pat 1
-        fun $select.tm 2
-        fun $typemap 1
-        var S
-          type-con $state 0
-        fun $idx 3
-        var p
-          type-con $ptr 0
-        var i
-          int
-        var T
-          type-con $ctype 0
-      pat 1
-        fun $select.mem 2
-        fun $memory 1
-        var S
-          type-con $state 0
-        fun $idx 3
-        var p
-          type-con $ptr 0
-        var i
-          int
-        var T
-          type-con $ctype 0
-      attribute qid 1
-        string-attr VccPrelu.2028:13
-      attribute uniqueId 1
-        string-attr 234
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 2
-    <=
-    int-num 0
-    var i
-      int
-    <
-    var i
-      int
-    var sz
-      int
-    and 2
-    fun $ts_is_array_elt 1
-    fun $select.tm 2
-    fun $typemap 1
-    var S
-      type-con $state 0
-    fun $idx 3
-    var p
-      type-con $ptr 0
-    var i
-      int
-    var T
-      type-con $ctype 0
-    fun $typed 2
-    var S
-      type-con $state 0
-    fun $idx 3
-    var p
-      type-con $ptr 0
-    var i
-      int
-    var T
-      type-con $ctype 0
-axiom 0
-    forall 4 1 3
-      var p
-        type-con $ptr 0
-      var #r
-        int
-      var T
-        type-con $ctype 0
-      var sz
-        int
-      pat 1
-        fun $set_in 2
-        var p
-          type-con $ptr 0
-        fun $full_extent 1
-        fun $ptr 2
-        fun $array 2
-        var T
-          type-con $ctype 0
-        var sz
-          int
-        var #r
-          int
-      attribute qid 1
-        string-attr VccPrelu.2094:15
-      attribute uniqueId 1
-        string-attr 243
-      attribute bvZ3Native 1
-        string-attr False
-    =
-    fun $set_in 2
-    var p
-      type-con $ptr 0
-    fun $full_extent 1
-    fun $ptr 2
-    fun $array 2
-    var T
-      type-con $ctype 0
-    var sz
-      int
-    var #r
-      int
-    or 2
-    =
-    var p
-      type-con $ptr 0
-    fun $ptr 2
-    fun $array 2
-    var T
-      type-con $ctype 0
-    var sz
-      int
-    var #r
-      int
-    and 3
-    <=
-    int-num 0
-    fun $index_within 2
-    var p
-      type-con $ptr 0
-    fun $ptr 2
-    var T
-      type-con $ctype 0
-    var #r
-      int
-    <=
-    fun $index_within 2
-    var p
-      type-con $ptr 0
-    fun $ptr 2
-    var T
-      type-con $ctype 0
-    var #r
-      int
-    -
-    var sz
-      int
-    int-num 1
-    fun $set_in 2
-    var p
-      type-con $ptr 0
-    fun $full_extent 1
-    fun $idx 3
-    fun $ptr 2
-    var T
-      type-con $ctype 0
-    var #r
-      int
-    fun $index_within 2
-    var p
-      type-con $ptr 0
-    fun $ptr 2
-    var T
-      type-con $ctype 0
-    var #r
-      int
-    var T
-      type-con $ctype 0
-axiom 0
-    forall 5 1 3
-      var S
-        type-con $state 0
-      var p
-        type-con $ptr 0
-      var #r
-        int
-      var T
-        type-con $ctype 0
-      var sz
-        int
-      pat 1
-        fun $set_in 2
-        var p
-          type-con $ptr 0
-        fun $extent 2
-        var S
-          type-con $state 0
-        fun $ptr 2
-        fun $array 2
-        var T
-          type-con $ctype 0
-        var sz
-          int
-        var #r
-          int
-      attribute qid 1
-        string-attr VccPrelu.2099:15
-      attribute uniqueId 1
-        string-attr 244
-      attribute bvZ3Native 1
-        string-attr False
-    =
-    fun $set_in 2
-    var p
-      type-con $ptr 0
-    fun $extent 2
-    var S
-      type-con $state 0
-    fun $ptr 2
-    fun $array 2
-    var T
-      type-con $ctype 0
-    var sz
-      int
-    var #r
-      int
-    or 2
-    =
-    var p
-      type-con $ptr 0
-    fun $ptr 2
-    fun $array 2
-    var T
-      type-con $ctype 0
-    var sz
-      int
-    var #r
-      int
-    and 3
-    <=
-    int-num 0
-    fun $index_within 2
-    var p
-      type-con $ptr 0
-    fun $ptr 2
-    var T
-      type-con $ctype 0
-    var #r
-      int
-    <=
-    fun $index_within 2
-    var p
-      type-con $ptr 0
-    fun $ptr 2
-    var T
-      type-con $ctype 0
-    var #r
-      int
-    -
-    var sz
-      int
-    int-num 1
-    fun $set_in 2
-    var p
-      type-con $ptr 0
-    fun $extent 2
-    var S
-      type-con $state 0
-    fun $idx 3
-    fun $ptr 2
-    var T
-      type-con $ctype 0
-    var #r
-      int
-    fun $index_within 2
-    var p
-      type-con $ptr 0
-    fun $ptr 2
-    var T
-      type-con $ctype 0
-    var #r
-      int
-    var T
-      type-con $ctype 0
-axiom 0
-    forall 5 2 3
-      var S
-        type-con $state 0
-      var #r
-        int
-      var T
-        type-con $ctype 0
-      var sz
-        int
-      var i
-        int
-      pat 2
-        fun $select.sm 2
-        fun $statusmap 1
-        var S
-          type-con $state 0
-        fun $idx 3
-        fun $ptr 2
-        var T
-          type-con $ctype 0
-        var #r
-          int
-        var i
-          int
-        var T
-          type-con $ctype 0
-        fun $ptr 2
-        fun $array 2
-        var T
-          type-con $ctype 0
-        var sz
-          int
-        var #r
-          int
-      pat 2
-        fun $select.tm 2
-        fun $typemap 1
-        var S
-          type-con $state 0
-        fun $idx 3
-        fun $ptr 2
-        var T
-          type-con $ctype 0
-        var #r
-          int
-        var i
-          int
-        var T
-          type-con $ctype 0
-        fun $ptr 2
-        fun $array 2
-        var T
-          type-con $ctype 0
-        var sz
-          int
-        var #r
-          int
-      attribute qid 1
-        string-attr VccPrelu.2107:15
-      attribute uniqueId 1
-        string-attr 245
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    fun $typed 2
-    var S
-      type-con $state 0
-    fun $ptr 2
-    fun $array 2
-    var T
-      type-con $ctype 0
-    var sz
-      int
-    var #r
-      int
-    implies
-    and 2
-    <=
-    int-num 0
-    var i
-      int
-    <
-    var i
-      int
-    var sz
-      int
-    and 4
-    =
-    fun $ts_emb 1
-    fun $select.tm 2
-    fun $typemap 1
-    var S
-      type-con $state 0
-    fun $idx 3
-    fun $ptr 2
-    fun $array 2
-    var T
-      type-con $ctype 0
-    var sz
-      int
-    var #r
-      int
-    var i
-      int
-    var T
-      type-con $ctype 0
-    fun $ptr 2
-    fun $array 2
-    var T
-      type-con $ctype 0
-    var sz
-      int
-    var #r
-      int
-    not
-    fun $ts_is_volatile 1
-    fun $select.tm 2
-    fun $typemap 1
-    var S
-      type-con $state 0
-    fun $idx 3
-    fun $ptr 2
-    fun $array 2
-    var T
-      type-con $ctype 0
-    var sz
-      int
-    var #r
-      int
-    var i
-      int
-    var T
-      type-con $ctype 0
-    fun $ts_is_array_elt 1
-    fun $select.tm 2
-    fun $typemap 1
-    var S
-      type-con $state 0
-    fun $idx 3
-    fun $ptr 2
-    fun $array 2
-    var T
-      type-con $ctype 0
-    var sz
-      int
-    var #r
-      int
-    var i
-      int
-    var T
-      type-con $ctype 0
-    fun $typed 2
-    var S
-      type-con $state 0
-    fun $idx 3
-    fun $ptr 2
-    fun $array 2
-    var T
-      type-con $ctype 0
-    var sz
-      int
-    var #r
-      int
-    var i
-      int
-    var T
-      type-con $ctype 0
-axiom 0
-    forall 4 1 3
-      var p
-        type-con $ptr 0
-      var T
-        type-con $ctype 0
-      var sz
-        int
-      var elem
-        type-con $ptr 0
-      pat 1
-        fun $set_in 2
-        var elem
-          type-con $ptr 0
-        fun $array_members 3
-        var p
-          type-con $ptr 0
-        var T
-          type-con $ctype 0
-        var sz
-          int
-      attribute qid 1
-        string-attr VccPrelu.2116:15
-      attribute uniqueId 1
-        string-attr 246
-      attribute bvZ3Native 1
-        string-attr False
-    =
-    fun $set_in 2
-    var elem
-      type-con $ptr 0
-    fun $array_members 3
-    var p
-      type-con $ptr 0
-    var T
-      type-con $ctype 0
-    var sz
-      int
-    and 3
-    <=
-    int-num 0
-    fun $index_within 2
-    var elem
-      type-con $ptr 0
-    var p
-      type-con $ptr 0
-    <=
-    fun $index_within 2
-    var elem
-      type-con $ptr 0
-    var p
-      type-con $ptr 0
-    -
-    var sz
-      int
-    int-num 1
-    =
-    var elem
-      type-con $ptr 0
-    fun $idx 3
-    var p
-      type-con $ptr 0
-    fun $index_within 2
-    var elem
-      type-con $ptr 0
-    var p
-      type-con $ptr 0
-    var T
-      type-con $ctype 0
-axiom 0
-    forall 4 1 3
-      var p
-        type-con $ptr 0
-      var #r
-        int
-      var T
-        type-con $ctype 0
-      var sz
-        int
-      pat 1
-        fun $set_in 2
-        var p
-          type-con $ptr 0
-        fun $array_range 3
-        fun $ptr 2
-        var T
-          type-con $ctype 0
-        var #r
-          int
-        var T
-          type-con $ctype 0
-        var sz
-          int
-      attribute qid 1
-        string-attr VccPrelu.2122:15
-      attribute uniqueId 1
-        string-attr 247
-      attribute bvZ3Native 1
-        string-attr False
-    =
-    fun $set_in 2
-    var p
-      type-con $ptr 0
-    fun $array_range 3
-    fun $ptr 2
-    var T
-      type-con $ctype 0
-    var #r
-      int
-    var T
-      type-con $ctype 0
-    var sz
-      int
-    and 3
-    <=
-    int-num 0
-    fun $index_within 2
-    var p
-      type-con $ptr 0
-    fun $ptr 2
-    var T
-      type-con $ctype 0
-    var #r
-      int
-    <=
-    fun $index_within 2
-    var p
-      type-con $ptr 0
-    fun $ptr 2
-    var T
-      type-con $ctype 0
-    var #r
-      int
-    -
-    var sz
-      int
-    int-num 1
-    fun $set_in 2
-    var p
-      type-con $ptr 0
-    fun $full_extent 1
-    fun $idx 3
-    fun $ptr 2
-    var T
-      type-con $ctype 0
-    var #r
-      int
-    fun $index_within 2
-    var p
-      type-con $ptr 0
-    fun $ptr 2
-    var T
-      type-con $ctype 0
-    var #r
-      int
-    var T
-      type-con $ctype 0
-axiom 0
-    forall 5 1 3
-      var p
-        type-con $ptr 0
-      var T
-        type-con $ctype 0
-      var sz
-        int
-      var idx
-        int
-      var S
-        type-con $ptrset 0
-      pat 2
-        fun $idx 3
-        var p
-          type-con $ptr 0
-        var idx
-          int
-        var T
-          type-con $ctype 0
-        fun $set_disjoint 2
-        fun $array_range 3
-        var p
-          type-con $ptr 0
-        var T
-          type-con $ctype 0
-        var sz
-          int
-        var S
-          type-con $ptrset 0
-      attribute qid 1
-        string-attr VccPrelu.2126:15
-      attribute uniqueId 1
-        string-attr 248
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 2
-    <=
-    int-num 0
-    var idx
-      int
-    <
-    var idx
-      int
-    var sz
-      int
-    =
-    fun $id_set_disjoint 3
-    fun $idx 3
-    var p
-      type-con $ptr 0
-    var idx
-      int
-    var T
-      type-con $ctype 0
-    fun $array_range 3
-    var p
-      type-con $ptr 0
-    var T
-      type-con $ctype 0
-    var sz
-      int
-    var S
-      type-con $ptrset 0
-    int-num 1
-axiom 0
-    forall 5 1 3
-      var p
-        type-con $ptr 0
-      var T
-        type-con $ctype 0
-      var sz
-        int
-      var idx
-        int
-      var S
-        type-con $ptrset 0
-      pat 2
-        fun $idx 3
-        var p
-          type-con $ptr 0
-        var idx
-          int
-        var T
-          type-con $ctype 0
-        fun $set_disjoint 2
-        var S
-          type-con $ptrset 0
-        fun $array_range 3
-        var p
-          type-con $ptr 0
-        var T
-          type-con $ctype 0
-        var sz
-          int
-      attribute qid 1
-        string-attr VccPrelu.2130:15
-      attribute uniqueId 1
-        string-attr 249
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 2
-    <=
-    int-num 0
-    var idx
-      int
-    <
-    var idx
-      int
-    var sz
-      int
-    =
-    fun $id_set_disjoint 3
-    fun $idx 3
-    var p
-      type-con $ptr 0
-    var idx
-      int
-    var T
-      type-con $ctype 0
-    var S
-      type-con $ptrset 0
-    fun $array_range 3
-    var p
-      type-con $ptr 0
-    var T
-      type-con $ctype 0
-    var sz
-      int
-    int-num 2
-axiom 0
-    forall 4 1 3
-      var p
-        type-con $ptr 0
-      var #r
-        int
-      var T
-        type-con $ctype 0
-      var sz
-        int
-      pat 1
-        fun $set_in 2
-        var p
-          type-con $ptr 0
-        fun $non_null_array_range 3
-        fun $ptr 2
-        var T
-          type-con $ctype 0
-        var #r
-          int
-        var T
-          type-con $ctype 0
-        var sz
-          int
-      attribute qid 1
-        string-attr VccPrelu.2135:15
-      attribute uniqueId 1
-        string-attr 250
-      attribute bvZ3Native 1
-        string-attr False
-    =
-    fun $set_in 2
-    var p
-      type-con $ptr 0
-    fun $non_null_array_range 3
-    fun $ptr 2
-    var T
-      type-con $ctype 0
-    var #r
-      int
-    var T
-      type-con $ctype 0
-    var sz
-      int
-    and 4
-    not
-    =
-    var #r
-      int
-    int-num 0
-    <=
-    int-num 0
-    fun $index_within 2
-    var p
-      type-con $ptr 0
-    fun $ptr 2
-    var T
-      type-con $ctype 0
-    var #r
-      int
-    <=
-    fun $index_within 2
-    var p
-      type-con $ptr 0
-    fun $ptr 2
-    var T
-      type-con $ctype 0
-    var #r
-      int
-    -
-    var sz
-      int
-    int-num 1
-    fun $set_in 2
-    var p
-      type-con $ptr 0
-    fun $full_extent 1
-    fun $idx 3
-    fun $ptr 2
-    var T
-      type-con $ctype 0
-    var #r
-      int
-    fun $index_within 2
-    var p
-      type-con $ptr 0
-    fun $ptr 2
-    var T
-      type-con $ctype 0
-    var #r
-      int
-    var T
-      type-con $ctype 0
-axiom 0
-    forall 3 1 3
-      var q
-        type-con $ptr 0
-      var S
-        type-con $state 0
-      var p
-        type-con $ptr 0
-      pat 1
-        fun $set_in 2
-        var q
-          type-con $ptr 0
-        fun $non_null_extent 2
-        var S
-          type-con $state 0
-        var p
-          type-con $ptr 0
-      attribute qid 1
-        string-attr VccPrelu.2142:15
-      attribute uniqueId 1
-        string-attr 251
-      attribute bvZ3Native 1
-        string-attr False
-    =
-    fun $set_in 2
-    var q
-      type-con $ptr 0
-    fun $non_null_extent 2
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    and 2
-    not
-    =
-    fun $ref 1
-    var p
-      type-con $ptr 0
-    fun $ref 1
-    fun $null 0
-    fun $set_in 2
-    var q
-      type-con $ptr 0
-    fun $extent 2
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-axiom 0
-    forall 2 1 3
-      var p
-        type-con $ptr 0
-      var k
-        int
-      pat 1
-        fun $idx 3
-        var p
-          type-con $ptr 0
-        var k
-          int
-        fun $typ 1
-        var p
-          type-con $ptr 0
-      attribute qid 1
-        string-attr VccPrelu.2154:15
-      attribute uniqueId 1
-        string-attr 253
-      attribute bvZ3Native 1
-        string-attr False
-    =
-    fun $index_within 2
-    fun $idx 3
-    var p
-      type-con $ptr 0
-    var k
-      int
-    fun $typ 1
-    var p
-      type-con $ptr 0
-    var p
-      type-con $ptr 0
-    var k
-      int
-axiom 0
-    forall 3 1 3
-      var p
-        type-con $ptr 0
-      var k
-        int
-      var f
-        type-con $field 0
-      pat 1
-        fun $index_within 2
-        fun $dot 2
-        fun $idx 3
-        var p
-          type-con $ptr 0
-        var k
-          int
-        fun $typ 1
-        var p
-          type-con $ptr 0
-        var f
-          type-con $field 0
-        var p
-          type-con $ptr 0
-      attribute qid 1
-        string-attr VccPrelu.2156:15
-      attribute uniqueId 1
-        string-attr 254
-      attribute bvZ3Native 1
-        string-attr False
-    =
-    fun $index_within 2
-    fun $dot 2
-    fun $idx 3
-    var p
-      type-con $ptr 0
-    var k
-      int
-    fun $typ 1
-    var p
-      type-con $ptr 0
-    var f
-      type-con $field 0
-    var p
-      type-con $ptr 0
-    var k
-      int
-axiom 0
-    forall 5 1 3
-      var s1
-        type-con $state 0
-      var s2
-        type-con $state 0
-      var p
-        type-con $ptr 0
-      var t
-        type-con $ctype 0
-      var sz
-        int
-      pat 2
-        fun $state_spans_the_same 4
-        var s1
-          type-con $state 0
-        var s2
-          type-con $state 0
-        var p
-          type-con $ptr 0
-        fun $array 2
-        var t
-          type-con $ctype 0
-        var sz
-          int
-        fun $is_primitive 1
-        var t
-          type-con $ctype 0
-      attribute qid 1
-        string-attr VccPrelu.2168:15
-      attribute uniqueId 1
-        string-attr 256
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    fun $is_primitive 1
-    var t
-      type-con $ctype 0
-    implies
-    fun $state_spans_the_same 4
-    var s1
-      type-con $state 0
-    var s2
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    fun $array 2
-    var t
-      type-con $ctype 0
-    var sz
-      int
-    forall 1 1 3
-      var i
-        int
-      pat 1
-        fun $select.mem 2
-        fun $memory 1
-        var s2
-          type-con $state 0
-        fun $idx 3
-        fun $ptr 2
-        var t
-          type-con $ctype 0
-        fun $ref 1
-        var p
-          type-con $ptr 0
-        var i
-          int
-        var t
-          type-con $ctype 0
-      attribute qid 1
-        string-attr VccPrelu.2171:15
-      attribute uniqueId 1
-        string-attr 255
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 2
-    <=
-    int-num 0
-    var i
-      int
-    <
-    var i
-      int
-    var sz
-      int
-    =
-    fun $select.mem 2
-    fun $memory 1
-    var s1
-      type-con $state 0
-    fun $idx 3
-    fun $ptr 2
-    var t
-      type-con $ctype 0
-    fun $ref 1
-    var p
-      type-con $ptr 0
-    var i
-      int
-    var t
-      type-con $ctype 0
-    fun $select.mem 2
-    fun $memory 1
-    var s2
-      type-con $state 0
-    fun $idx 3
-    fun $ptr 2
-    var t
-      type-con $ctype 0
-    fun $ref 1
-    var p
-      type-con $ptr 0
-    var i
-      int
-    var t
-      type-con $ctype 0
-axiom 0
-    forall 1 1 4
-      var x
-        bool
-      pat 1
-        fun $bool_id 1
-        var x
-          bool
-      attribute qid 1
-        string-attr VccPrelu.2211:31
-      attribute uniqueId 1
-        string-attr 257
-      attribute bvZ3Native 1
-        string-attr False
-      attribute weight 1
-        expr-attr
-          int-num 0
-    =
-    fun $bool_id 1
-    var x
-      bool
-    var x
-      bool
-axiom 0
-    =
-    fun $min.i1 0
-    -
-    int-num 0
-    int-num 128
-axiom 0
-    =
-    fun $max.i1 0
-    int-num 127
-axiom 0
-    =
-    fun $min.i2 0
-    -
-    int-num 0
-    int-num 32768
-axiom 0
-    =
-    fun $max.i2 0
-    int-num 32767
-axiom 0
-    =
-    fun $min.i4 0
-    -
-    int-num 0
-    *
-    int-num 65536
-    int-num 32768
-axiom 0
-    =
-    fun $max.i4 0
-    -
-    *
-    int-num 65536
-    int-num 32768
-    int-num 1
-axiom 0
-    =
-    fun $min.i8 0
-    -
-    int-num 0
-    *
-    *
-    *
-    int-num 65536
-    int-num 65536
-    int-num 65536
-    int-num 32768
-axiom 0
-    =
-    fun $max.i8 0
-    -
-    *
-    *
-    *
-    int-num 65536
-    int-num 65536
-    int-num 65536
-    int-num 32768
-    int-num 1
-axiom 0
-    =
-    fun $max.u1 0
-    int-num 255
-axiom 0
-    =
-    fun $max.u2 0
-    int-num 65535
-axiom 0
-    =
-    fun $max.u4 0
-    -
-    *
-    int-num 65536
-    int-num 65536
-    int-num 1
-axiom 0
-    =
-    fun $max.u8 0
-    -
-    *
-    *
-    *
-    int-num 65536
-    int-num 65536
-    int-num 65536
-    int-num 65536
-    int-num 1
-axiom 0
-    forall 2 1 4
-      var S
-        type-con $state 0
-      var p
-        type-con $ptr 0
-      pat 1
-        fun $read_i1 2
-        var S
-          type-con $state 0
-        var p
-          type-con $ptr 0
-      attribute qid 1
-        string-attr VccPrelu.2253:31
-      attribute uniqueId 1
-        string-attr 258
-      attribute bvZ3Native 1
-        string-attr False
-      attribute weight 1
-        expr-attr
-          int-num 0
-    =
-    fun $read_i1 2
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    fun $select.mem 2
-    fun $memory 1
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-axiom 0
-    forall 2 1 4
-      var S
-        type-con $state 0
-      var p
-        type-con $ptr 0
-      pat 1
-        fun $read_i2 2
-        var S
-          type-con $state 0
-        var p
-          type-con $ptr 0
-      attribute qid 1
-        string-attr VccPrelu.2254:31
-      attribute uniqueId 1
-        string-attr 259
-      attribute bvZ3Native 1
-        string-attr False
-      attribute weight 1
-        expr-attr
-          int-num 0
-    =
-    fun $read_i2 2
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    fun $select.mem 2
-    fun $memory 1
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-axiom 0
-    forall 2 1 4
-      var S
-        type-con $state 0
-      var p
-        type-con $ptr 0
-      pat 1
-        fun $read_i4 2
-        var S
-          type-con $state 0
-        var p
-          type-con $ptr 0
-      attribute qid 1
-        string-attr VccPrelu.2255:31
-      attribute uniqueId 1
-        string-attr 260
-      attribute bvZ3Native 1
-        string-attr False
-      attribute weight 1
-        expr-attr
-          int-num 0
-    =
-    fun $read_i4 2
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    fun $select.mem 2
-    fun $memory 1
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-axiom 0
-    forall 2 1 4
-      var S
-        type-con $state 0
-      var p
-        type-con $ptr 0
-      pat 1
-        fun $read_i8 2
-        var S
-          type-con $state 0
-        var p
-          type-con $ptr 0
-      attribute qid 1
-        string-attr VccPrelu.2256:31
-      attribute uniqueId 1
-        string-attr 261
-      attribute bvZ3Native 1
-        string-attr False
-      attribute weight 1
-        expr-attr
-          int-num 0
-    =
-    fun $read_i8 2
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    fun $select.mem 2
-    fun $memory 1
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-axiom 0
-    forall 2 1 4
-      var S
-        type-con $state 0
-      var p
-        type-con $ptr 0
-      pat 1
-        fun $read_u1 2
-        var S
-          type-con $state 0
-        var p
-          type-con $ptr 0
-      attribute qid 1
-        string-attr VccPrelu.2257:31
-      attribute uniqueId 1
-        string-attr 262
-      attribute bvZ3Native 1
-        string-attr False
-      attribute weight 1
-        expr-attr
-          int-num 0
-    =
-    fun $read_u1 2
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    fun $select.mem 2
-    fun $memory 1
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-axiom 0
-    forall 2 1 4
-      var S
-        type-con $state 0
-      var p
-        type-con $ptr 0
-      pat 1
-        fun $read_u2 2
-        var S
-          type-con $state 0
-        var p
-          type-con $ptr 0
-      attribute qid 1
-        string-attr VccPrelu.2258:31
-      attribute uniqueId 1
-        string-attr 263
-      attribute bvZ3Native 1
-        string-attr False
-      attribute weight 1
-        expr-attr
-          int-num 0
-    =
-    fun $read_u2 2
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    fun $select.mem 2
-    fun $memory 1
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-axiom 0
-    forall 2 1 4
-      var S
-        type-con $state 0
-      var p
-        type-con $ptr 0
-      pat 1
-        fun $read_u4 2
-        var S
-          type-con $state 0
-        var p
-          type-con $ptr 0
-      attribute qid 1
-        string-attr VccPrelu.2259:31
-      attribute uniqueId 1
-        string-attr 264
-      attribute bvZ3Native 1
-        string-attr False
-      attribute weight 1
-        expr-attr
-          int-num 0
-    =
-    fun $read_u4 2
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    fun $select.mem 2
-    fun $memory 1
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-axiom 0
-    forall 2 1 4
-      var S
-        type-con $state 0
-      var p
-        type-con $ptr 0
-      pat 1
-        fun $read_u8 2
-        var S
-          type-con $state 0
-        var p
-          type-con $ptr 0
-      attribute qid 1
-        string-attr VccPrelu.2260:31
-      attribute uniqueId 1
-        string-attr 265
-      attribute bvZ3Native 1
-        string-attr False
-      attribute weight 1
-        expr-attr
-          int-num 0
-    =
-    fun $read_u8 2
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    fun $select.mem 2
-    fun $memory 1
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-axiom 0
-    =
-    fun $ptr_to_u8 1
-    fun $null 0
-    int-num 0
-axiom 0
-    =
-    fun $ptr_to_i8 1
-    fun $null 0
-    int-num 0
-axiom 0
-    =
-    fun $ptr_to_u4 1
-    fun $null 0
-    int-num 0
-axiom 0
-    =
-    fun $ptr_to_i4 1
-    fun $null 0
-    int-num 0
-axiom 0
-    forall 1 1 3
-      var p
-        type-con $ptr 0
-      pat 1
-        fun $ptr_to_u8 1
-        var p
-          type-con $ptr 0
-      attribute qid 1
-        string-attr VccPrelu.2277:15
-      attribute uniqueId 1
-        string-attr 266
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 2
-    <=
-    int-num 0
-    fun $ref 1
-    var p
-      type-con $ptr 0
-    <=
-    fun $ref 1
-    var p
-      type-con $ptr 0
-    fun $max.u8 0
-    =
-    fun $ptr_to_u8 1
-    var p
-      type-con $ptr 0
-    fun $ref 1
-    var p
-      type-con $ptr 0
-axiom 0
-    forall 1 1 3
-      var p
-        type-con $ptr 0
-      pat 1
-        fun $ptr_to_i8 1
-        var p
-          type-con $ptr 0
-      attribute qid 1
-        string-attr VccPrelu.2278:15
-      attribute uniqueId 1
-        string-attr 267
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 2
-    <=
-    fun $min.i8 0
-    fun $ref 1
-    var p
-      type-con $ptr 0
-    <=
-    fun $ref 1
-    var p
-      type-con $ptr 0
-    fun $max.i8 0
-    =
-    fun $ptr_to_i8 1
-    var p
-      type-con $ptr 0
-    fun $ref 1
-    var p
-      type-con $ptr 0
-axiom 0
-    forall 1 1 3
-      var p
-        type-con $ptr 0
-      pat 1
-        fun $ptr_to_u4 1
-        var p
-          type-con $ptr 0
-      attribute qid 1
-        string-attr VccPrelu.2279:15
-      attribute uniqueId 1
-        string-attr 268
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 2
-    <=
-    int-num 0
-    fun $ref 1
-    var p
-      type-con $ptr 0
-    <=
-    fun $ref 1
-    var p
-      type-con $ptr 0
-    fun $max.u4 0
-    =
-    fun $ptr_to_u4 1
-    var p
-      type-con $ptr 0
-    fun $ref 1
-    var p
-      type-con $ptr 0
-axiom 0
-    forall 1 1 3
-      var p
-        type-con $ptr 0
-      pat 1
-        fun $ptr_to_i4 1
-        var p
-          type-con $ptr 0
-      attribute qid 1
-        string-attr VccPrelu.2280:15
-      attribute uniqueId 1
-        string-attr 269
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 2
-    <=
-    fun $min.i4 0
-    fun $ref 1
-    var p
-      type-con $ptr 0
-    <=
-    fun $ref 1
-    var p
-      type-con $ptr 0
-    fun $max.i4 0
-    =
-    fun $ptr_to_i4 1
-    var p
-      type-con $ptr 0
-    fun $ref 1
-    var p
-      type-con $ptr 0
-axiom 0
-    forall 2 1 4
-      var p1
-        type-con $ptr 0
-      var p2
-        type-con $ptr 0
-      pat 1
-        fun $byte_ptr_subtraction 2
-        var p1
-          type-con $ptr 0
-        var p2
-          type-con $ptr 0
-      attribute qid 1
-        string-attr VccPrelu.2282:44
-      attribute uniqueId 1
-        string-attr 270
-      attribute bvZ3Native 1
-        string-attr False
-      attribute weight 1
-        expr-attr
-          int-num 0
-    =
-    fun $byte_ptr_subtraction 2
-    var p1
-      type-con $ptr 0
-    var p2
-      type-con $ptr 0
-    -
-    fun $ref 1
-    var p1
-      type-con $ptr 0
-    fun $ref 1
-    var p2
-      type-con $ptr 0
-axiom 0
-    forall 2 1 3
-      var S
-        type-con $state 0
-      var p
-        type-con $ptr 0
-      pat 1
-        fun $read_i1 2
-        var S
-          type-con $state 0
-        var p
-          type-con $ptr 0
-      attribute qid 1
-        string-attr VccPrelu.2285:15
-      attribute uniqueId 1
-        string-attr 271
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    fun $good_state 1
-    var S
-      type-con $state 0
-    and 2
-    <=
-    fun $min.i1 0
-    fun $read_i1 2
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    <=
-    fun $read_i1 2
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    fun $max.i1 0
-axiom 0
-    forall 2 1 3
-      var S
-        type-con $state 0
-      var p
-        type-con $ptr 0
-      pat 1
-        fun $read_i2 2
-        var S
-          type-con $state 0
-        var p
-          type-con $ptr 0
-      attribute qid 1
-        string-attr VccPrelu.2286:15
-      attribute uniqueId 1
-        string-attr 272
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    fun $good_state 1
-    var S
-      type-con $state 0
-    and 2
-    <=
-    fun $min.i2 0
-    fun $read_i2 2
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    <=
-    fun $read_i2 2
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    fun $max.i2 0
-axiom 0
-    forall 2 1 3
-      var S
-        type-con $state 0
-      var p
-        type-con $ptr 0
-      pat 1
-        fun $read_i4 2
-        var S
-          type-con $state 0
-        var p
-          type-con $ptr 0
-      attribute qid 1
-        string-attr VccPrelu.2287:15
-      attribute uniqueId 1
-        string-attr 273
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    fun $good_state 1
-    var S
-      type-con $state 0
-    and 2
-    <=
-    fun $min.i4 0
-    fun $read_i4 2
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    <=
-    fun $read_i4 2
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    fun $max.i4 0
-axiom 0
-    forall 2 1 3
-      var S
-        type-con $state 0
-      var p
-        type-con $ptr 0
-      pat 1
-        fun $read_i8 2
-        var S
-          type-con $state 0
-        var p
-          type-con $ptr 0
-      attribute qid 1
-        string-attr VccPrelu.2288:15
-      attribute uniqueId 1
-        string-attr 274
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    fun $good_state 1
-    var S
-      type-con $state 0
-    and 2
-    <=
-    fun $min.i8 0
-    fun $read_i8 2
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    <=
-    fun $read_i8 2
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    fun $max.i8 0
-axiom 0
-    forall 2 1 3
-      var S
-        type-con $state 0
-      var p
-        type-con $ptr 0
-      pat 1
-        fun $read_u1 2
-        var S
-          type-con $state 0
-        var p
-          type-con $ptr 0
-      attribute qid 1
-        string-attr VccPrelu.2289:15
-      attribute uniqueId 1
-        string-attr 275
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    fun $good_state 1
-    var S
-      type-con $state 0
-    and 2
-    <=
-    int-num 0
-    fun $read_u1 2
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    <=
-    fun $read_u1 2
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    fun $max.u1 0
-axiom 0
-    forall 2 1 3
-      var S
-        type-con $state 0
-      var p
-        type-con $ptr 0
-      pat 1
-        fun $read_u2 2
-        var S
-          type-con $state 0
-        var p
-          type-con $ptr 0
-      attribute qid 1
-        string-attr VccPrelu.2290:15
-      attribute uniqueId 1
-        string-attr 276
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    fun $good_state 1
-    var S
-      type-con $state 0
-    and 2
-    <=
-    int-num 0
-    fun $read_u2 2
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    <=
-    fun $read_u2 2
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    fun $max.u2 0
-axiom 0
-    forall 2 1 3
-      var S
-        type-con $state 0
-      var p
-        type-con $ptr 0
-      pat 1
-        fun $read_u4 2
-        var S
-          type-con $state 0
-        var p
-          type-con $ptr 0
-      attribute qid 1
-        string-attr VccPrelu.2291:15
-      attribute uniqueId 1
-        string-attr 277
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    fun $good_state 1
-    var S
-      type-con $state 0
-    and 2
-    <=
-    int-num 0
-    fun $read_u4 2
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    <=
-    fun $read_u4 2
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    fun $max.u4 0
-axiom 0
-    forall 2 1 3
-      var S
-        type-con $state 0
-      var p
-        type-con $ptr 0
-      pat 1
-        fun $read_u8 2
-        var S
-          type-con $state 0
-        var p
-          type-con $ptr 0
-      attribute qid 1
-        string-attr VccPrelu.2292:15
-      attribute uniqueId 1
-        string-attr 278
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    fun $good_state 1
-    var S
-      type-con $state 0
-    and 2
-    <=
-    int-num 0
-    fun $read_u8 2
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    <=
-    fun $read_u8 2
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    fun $max.u8 0
-axiom 0
-    and 64
-    =
-    fun $_pow2 1
-    int-num 0
-    int-num 1
-    =
-    fun $_pow2 1
-    int-num 1
-    int-num 2
-    =
-    fun $_pow2 1
-    int-num 2
-    int-num 4
-    =
-    fun $_pow2 1
-    int-num 3
-    int-num 8
-    =
-    fun $_pow2 1
-    int-num 4
-    int-num 16
-    =
-    fun $_pow2 1
-    int-num 5
-    int-num 32
-    =
-    fun $_pow2 1
-    int-num 6
-    int-num 64
-    =
-    fun $_pow2 1
-    int-num 7
-    int-num 128
-    =
-    fun $_pow2 1
-    int-num 8
-    int-num 256
-    =
-    fun $_pow2 1
-    int-num 9
-    int-num 512
-    =
-    fun $_pow2 1
-    int-num 10
-    int-num 1024
-    =
-    fun $_pow2 1
-    int-num 11
-    int-num 2048
-    =
-    fun $_pow2 1
-    int-num 12
-    int-num 4096
-    =
-    fun $_pow2 1
-    int-num 13
-    int-num 8192
-    =
-    fun $_pow2 1
-    int-num 14
-    int-num 16384
-    =
-    fun $_pow2 1
-    int-num 15
-    int-num 32768
-    =
-    fun $_pow2 1
-    int-num 16
-    int-num 65536
-    =
-    fun $_pow2 1
-    int-num 17
-    int-num 131072
-    =
-    fun $_pow2 1
-    int-num 18
-    int-num 262144
-    =
-    fun $_pow2 1
-    int-num 19
-    int-num 524288
-    =
-    fun $_pow2 1
-    int-num 20
-    int-num 1048576
-    =
-    fun $_pow2 1
-    int-num 21
-    int-num 2097152
-    =
-    fun $_pow2 1
-    int-num 22
-    int-num 4194304
-    =
-    fun $_pow2 1
-    int-num 23
-    int-num 8388608
-    =
-    fun $_pow2 1
-    int-num 24
-    int-num 16777216
-    =
-    fun $_pow2 1
-    int-num 25
-    int-num 33554432
-    =
-    fun $_pow2 1
-    int-num 26
-    int-num 67108864
-    =
-    fun $_pow2 1
-    int-num 27
-    int-num 134217728
-    =
-    fun $_pow2 1
-    int-num 28
-    int-num 268435456
-    =
-    fun $_pow2 1
-    int-num 29
-    int-num 536870912
-    =
-    fun $_pow2 1
-    int-num 30
-    int-num 1073741824
-    =
-    fun $_pow2 1
-    int-num 31
-    int-num 2147483648
-    =
-    fun $_pow2 1
-    int-num 32
-    int-num 4294967296
-    =
-    fun $_pow2 1
-    int-num 33
-    int-num 8589934592
-    =
-    fun $_pow2 1
-    int-num 34
-    int-num 17179869184
-    =
-    fun $_pow2 1
-    int-num 35
-    int-num 34359738368
-    =
-    fun $_pow2 1
-    int-num 36
-    int-num 68719476736
-    =
-    fun $_pow2 1
-    int-num 37
-    int-num 137438953472
-    =
-    fun $_pow2 1
-    int-num 38
-    int-num 274877906944
-    =
-    fun $_pow2 1
-    int-num 39
-    int-num 549755813888
-    =
-    fun $_pow2 1
-    int-num 40
-    int-num 1099511627776
-    =
-    fun $_pow2 1
-    int-num 41
-    int-num 2199023255552
-    =
-    fun $_pow2 1
-    int-num 42
-    int-num 4398046511104
-    =
-    fun $_pow2 1
-    int-num 43
-    int-num 8796093022208
-    =
-    fun $_pow2 1
-    int-num 44
-    int-num 17592186044416
-    =
-    fun $_pow2 1
-    int-num 45
-    int-num 35184372088832
-    =
-    fun $_pow2 1
-    int-num 46
-    int-num 70368744177664
-    =
-    fun $_pow2 1
-    int-num 47
-    int-num 140737488355328
-    =
-    fun $_pow2 1
-    int-num 48
-    int-num 281474976710656
-    =
-    fun $_pow2 1
-    int-num 49
-    int-num 562949953421312
-    =
-    fun $_pow2 1
-    int-num 50
-    int-num 1125899906842624
-    =
-    fun $_pow2 1
-    int-num 51
-    int-num 2251799813685248
-    =
-    fun $_pow2 1
-    int-num 52
-    int-num 4503599627370496
-    =
-    fun $_pow2 1
-    int-num 53
-    int-num 9007199254740992
-    =
-    fun $_pow2 1
-    int-num 54
-    int-num 18014398509481984
-    =
-    fun $_pow2 1
-    int-num 55
-    int-num 36028797018963968
-    =
-    fun $_pow2 1
-    int-num 56
-    int-num 72057594037927936
-    =
-    fun $_pow2 1
-    int-num 57
-    int-num 144115188075855872
-    =
-    fun $_pow2 1
-    int-num 58
-    int-num 288230376151711744
-    =
-    fun $_pow2 1
-    int-num 59
-    int-num 576460752303423488
-    =
-    fun $_pow2 1
-    int-num 60
-    int-num 1152921504606846976
-    =
-    fun $_pow2 1
-    int-num 61
-    int-num 2305843009213693952
-    =
-    fun $_pow2 1
-    int-num 62
-    int-num 4611686018427387904
-    =
-    fun $_pow2 1
-    int-num 63
-    int-num 9223372036854775808
-axiom 0
-    forall 3 1 4
-      var t
-        type-con $ctype 0
-      var x
-        int
-      var y
-        int
-      pat 1
-        fun $unchk_add 3
-        var t
-          type-con $ctype 0
-        var x
-          int
-        var y
-          int
-      attribute qid 1
-        string-attr VccPrelu.2318:33
-      attribute uniqueId 1
-        string-attr 279
-      attribute bvZ3Native 1
-        string-attr False
-      attribute weight 1
-        expr-attr
-          int-num 0
-    =
-    fun $unchk_add 3
-    var t
-      type-con $ctype 0
-    var x
-      int
-    var y
-      int
-    fun $unchecked 2
-    var t
-      type-con $ctype 0
-    +
-    var x
-      int
-    var y
-      int
-axiom 0
-    forall 3 1 4
-      var t
-        type-con $ctype 0
-      var x
-        int
-      var y
-        int
-      pat 1
-        fun $unchk_sub 3
-        var t
-          type-con $ctype 0
-        var x
-          int
-        var y
-          int
-      attribute qid 1
-        string-attr VccPrelu.2319:33
-      attribute uniqueId 1
-        string-attr 280
-      attribute bvZ3Native 1
-        string-attr False
-      attribute weight 1
-        expr-attr
-          int-num 0
-    =
-    fun $unchk_sub 3
-    var t
-      type-con $ctype 0
-    var x
-      int
-    var y
-      int
-    fun $unchecked 2
-    var t
-      type-con $ctype 0
-    -
-    var x
-      int
-    var y
-      int
-axiom 0
-    forall 3 1 4
-      var t
-        type-con $ctype 0
-      var x
-        int
-      var y
-        int
-      pat 1
-        fun $unchk_mul 3
-        var t
-          type-con $ctype 0
-        var x
-          int
-        var y
-          int
-      attribute qid 1
-        string-attr VccPrelu.2320:33
-      attribute uniqueId 1
-        string-attr 281
-      attribute bvZ3Native 1
-        string-attr False
-      attribute weight 1
-        expr-attr
-          int-num 0
-    =
-    fun $unchk_mul 3
-    var t
-      type-con $ctype 0
-    var x
-      int
-    var y
-      int
-    fun $unchecked 2
-    var t
-      type-con $ctype 0
-    *
-    var x
-      int
-    var y
-      int
-axiom 0
-    forall 3 1 4
-      var t
-        type-con $ctype 0
-      var x
-        int
-      var y
-        int
-      pat 1
-        fun $unchk_div 3
-        var t
-          type-con $ctype 0
-        var x
-          int
-        var y
-          int
-      attribute qid 1
-        string-attr VccPrelu.2321:33
-      attribute uniqueId 1
-        string-attr 282
-      attribute bvZ3Native 1
-        string-attr False
-      attribute weight 1
-        expr-attr
-          int-num 0
-    =
-    fun $unchk_div 3
-    var t
-      type-con $ctype 0
-    var x
-      int
-    var y
-      int
-    fun $unchecked 2
-    var t
-      type-con $ctype 0
-    /
-    var x
-      int
-    var y
-      int
-axiom 0
-    forall 3 1 4
-      var t
-        type-con $ctype 0
-      var x
-        int
-      var y
-        int
-      pat 1
-        fun $unchk_mod 3
-        var t
-          type-con $ctype 0
-        var x
-          int
-        var y
-          int
-      attribute qid 1
-        string-attr VccPrelu.2322:33
-      attribute uniqueId 1
-        string-attr 283
-      attribute bvZ3Native 1
-        string-attr False
-      attribute weight 1
-        expr-attr
-          int-num 0
-    =
-    fun $unchk_mod 3
-    var t
-      type-con $ctype 0
-    var x
-      int
-    var y
-      int
-    fun $unchecked 2
-    var t
-      type-con $ctype 0
-    %
-    var x
-      int
-    var y
-      int
-axiom 0
-    forall 3 1 4
-      var t
-        type-con $ctype 0
-      var x
-        int
-      var y
-        int
-      pat 1
-        fun $_shl 3
-        var t
-          type-con $ctype 0
-        var x
-          int
-        var y
-          int
-      attribute qid 1
-        string-attr VccPrelu.2324:28
-      attribute uniqueId 1
-        string-attr 284
-      attribute bvZ3Native 1
-        string-attr False
-      attribute weight 1
-        expr-attr
-          int-num 0
-    =
-    fun $_shl 3
-    var t
-      type-con $ctype 0
-    var x
-      int
-    var y
-      int
-    fun $unchecked 2
-    var t
-      type-con $ctype 0
-    *
-    var x
-      int
-    fun $_pow2 1
-    var y
-      int
-axiom 0
-    forall 2 1 4
-      var x
-        int
-      var y
-        int
-      pat 1
-        fun $_shr 2
-        var x
-          int
-        var y
-          int
-      attribute qid 1
-        string-attr VccPrelu.2326:28
-      attribute uniqueId 1
-        string-attr 285
-      attribute bvZ3Native 1
-        string-attr False
-      attribute weight 1
-        expr-attr
-          int-num 0
-    =
-    fun $_shr 2
-    var x
-      int
-    var y
-      int
-    /
-    var x
-      int
-    fun $_pow2 1
-    var y
-      int
-axiom 0
-    forall 5 1 3
-      var x
-        int
-      var from
-        int
-      var to
-        int
-      var xs
-        int
-      var val
-        int
-      pat 1
-        fun $bv_update 5
-        var x
-          int
-        var xs
-          int
-        var from
-          int
-        var to
-          int
-        var val
-          int
-      attribute qid 1
-        string-attr VccPrelu.2333:15
-      attribute uniqueId 1
-        string-attr 286
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 3
-    <=
-    int-num 0
-    var from
-      int
-    <
-    var from
-      int
-    var to
-      int
-    <=
-    var to
-      int
-    var xs
-      int
-    implies
-    and 2
-    <=
-    int-num 0
-    var val
-      int
-    <
-    var val
-      int
-    fun $_pow2 1
-    -
-    var to
-      int
-    var from
-      int
-    and 2
-    <=
-    int-num 0
-    fun $bv_update 5
-    var x
-      int
-    var xs
-      int
-    var from
-      int
-    var to
-      int
-    var val
-      int
-    <
-    fun $bv_update 5
-    var x
-      int
-    var xs
-      int
-    var from
-      int
-    var to
-      int
-    var val
-      int
-    fun $_pow2 1
-    var xs
-      int
-axiom 0
-    forall 3 1 3
-      var from
-        int
-      var to
-        int
-      var xs
-        int
-      pat 1
-        fun $bv_update 5
-        int-num 0
-        var xs
-          int
-        var from
-          int
-        var to
-          int
-        int-num 0
-      attribute qid 1
-        string-attr VccPrelu.2339:15
-      attribute uniqueId 1
-        string-attr 287
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 3
-    <=
-    int-num 0
-    var from
-      int
-    <
-    var from
-      int
-    var to
-      int
-    <=
-    var to
-      int
-    var xs
-      int
-    =
-    fun $bv_update 5
-    int-num 0
-    var xs
-      int
-    var from
-      int
-    var to
-      int
-    int-num 0
-    int-num 0
-axiom 0
-    forall 5 1 3
-      var from
-        int
-      var to
-        int
-      var val
-        int
-      var x
-        int
-      var xs
-        int
-      pat 1
-        fun $bv_extract_signed 4
-        fun $bv_update 5
-        var x
-          int
-        var xs
-          int
-        var from
-          int
-        var to
-          int
-        var val
-          int
-        var xs
-          int
-        var from
-          int
-        var to
-          int
-      attribute qid 1
-        string-attr VccPrelu.2343:15
-      attribute uniqueId 1
-        string-attr 288
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 3
-    <=
-    int-num 0
-    var from
-      int
-    <
-    var from
-      int
-    var to
-      int
-    <=
-    var to
-      int
-    var xs
-      int
-    implies
-    and 2
-    <=
-    -
-    int-num 0
-    fun $_pow2 1
-    -
-    -
-    var to
-      int
-    var from
-      int
-    int-num 1
-    var val
-      int
-    <
-    var val
-      int
-    fun $_pow2 1
-    -
-    -
-    var to
-      int
-    var from
-      int
-    int-num 1
-    =
-    fun $bv_extract_signed 4
-    fun $bv_update 5
-    var x
-      int
-    var xs
-      int
-    var from
-      int
-    var to
-      int
-    var val
-      int
-    var xs
-      int
-    var from
-      int
-    var to
-      int
-    var val
-      int
-axiom 0
-    forall 5 1 3
-      var from
-        int
-      var to
-        int
-      var val
-        int
-      var x
-        int
-      var xs
-        int
-      pat 1
-        fun $bv_extract_unsigned 4
-        fun $bv_update 5
-        var x
-          int
-        var xs
-          int
-        var from
-          int
-        var to
-          int
-        var val
-          int
-        var xs
-          int
-        var from
-          int
-        var to
-          int
-      attribute qid 1
-        string-attr VccPrelu.2349:15
-      attribute uniqueId 1
-        string-attr 289
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 3
-    <=
-    int-num 0
-    var from
-      int
-    <
-    var from
-      int
-    var to
-      int
-    <=
-    var to
-      int
-    var xs
-      int
-    implies
-    and 2
-    <=
-    int-num 0
-    var val
-      int
-    <
-    var val
-      int
-    fun $_pow2 1
-    -
-    var to
-      int
-    var from
-      int
-    =
-    fun $bv_extract_unsigned 4
-    fun $bv_update 5
-    var x
-      int
-    var xs
-      int
-    var from
-      int
-    var to
-      int
-    var val
-      int
-    var xs
-      int
-    var from
-      int
-    var to
-      int
-    var val
-      int
-axiom 0
-    forall 4 1 3
-      var from
-        int
-      var to
-        int
-      var x
-        int
-      var xs
-        int
-      pat 1
-        fun $bv_extract_signed 4
-        var x
-          int
-        var xs
-          int
-        var from
-          int
-        var to
-          int
-      attribute qid 1
-        string-attr VccPrelu.2355:15
-      attribute uniqueId 1
-        string-attr 290
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 3
-    <=
-    int-num 0
-    var from
-      int
-    <
-    var from
-      int
-    var to
-      int
-    <=
-    var to
-      int
-    var xs
-      int
-    and 2
-    <=
-    -
-    int-num 0
-    fun $_pow2 1
-    -
-    -
-    var to
-      int
-    var from
-      int
-    int-num 1
-    fun $bv_extract_signed 4
-    var x
-      int
-    var xs
-      int
-    var from
-      int
-    var to
-      int
-    <=
-    fun $bv_extract_signed 4
-    var x
-      int
-    var xs
-      int
-    var from
-      int
-    var to
-      int
-    -
-    fun $_pow2 1
-    -
-    -
-    var to
-      int
-    var from
-      int
-    int-num 1
-    int-num 1
-axiom 0
-    forall 4 1 3
-      var from
-        int
-      var to
-        int
-      var x
-        int
-      var xs
-        int
-      pat 1
-        fun $bv_extract_unsigned 4
-        var x
-          int
-        var xs
-          int
-        var from
-          int
-        var to
-          int
-      attribute qid 1
-        string-attr VccPrelu.2360:15
-      attribute uniqueId 1
-        string-attr 291
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 3
-    <=
-    int-num 0
-    var from
-      int
-    <
-    var from
-      int
-    var to
-      int
-    <=
-    var to
-      int
-    var xs
-      int
-    and 2
-    <=
-    int-num 0
-    fun $bv_extract_unsigned 4
-    var x
-      int
-    var xs
-      int
-    var from
-      int
-    var to
-      int
-    <=
-    fun $bv_extract_unsigned 4
-    var x
-      int
-    var xs
-      int
-    var from
-      int
-    var to
-      int
-    -
-    fun $_pow2 1
-    -
-    var to
-      int
-    var from
-      int
-    int-num 1
-axiom 0
-    forall 7 1 3
-      var from
-        int
-      var to
-        int
-      var val
-        int
-      var x
-        int
-      var xs
-        int
-      var from2
-        int
-      var to2
-        int
-      pat 1
-        fun $bv_extract_signed 4
-        fun $bv_update 5
-        var x
-          int
-        var xs
-          int
-        var from
-          int
-        var to
-          int
-        var val
-          int
-        var xs
-          int
-        var from2
-          int
-        var to2
-          int
-      attribute qid 1
-        string-attr VccPrelu.2365:15
-      attribute uniqueId 1
-        string-attr 292
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 3
-    <=
-    int-num 0
-    var from
-      int
-    <
-    var from
-      int
-    var to
-      int
-    <=
-    var to
-      int
-    var xs
-      int
-    implies
-    and 3
-    <=
-    int-num 0
-    var from2
-      int
-    <
-    var from2
-      int
-    var to2
-      int
-    <=
-    var to2
-      int
-    var xs
-      int
-    implies
-    or 2
-    <=
-    var to2
-      int
-    var from
-      int
-    <=
-    var to
-      int
-    var from2
-      int
-    =
-    fun $bv_extract_signed 4
-    fun $bv_update 5
-    var x
-      int
-    var xs
-      int
-    var from
-      int
-    var to
-      int
-    var val
-      int
-    var xs
-      int
-    var from2
-      int
-    var to2
-      int
-    fun $bv_extract_signed 4
-    var x
-      int
-    var xs
-      int
-    var from2
-      int
-    var to2
-      int
-axiom 0
-    forall 7 1 3
-      var from
-        int
-      var to
-        int
-      var val
-        int
-      var x
-        int
-      var xs
-        int
-      var from2
-        int
-      var to2
-        int
-      pat 1
-        fun $bv_extract_unsigned 4
-        fun $bv_update 5
-        var x
-          int
-        var xs
-          int
-        var from
-          int
-        var to
-          int
-        var val
-          int
-        var xs
-          int
-        var from2
-          int
-        var to2
-          int
-      attribute qid 1
-        string-attr VccPrelu.2372:15
-      attribute uniqueId 1
-        string-attr 293
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 3
-    <=
-    int-num 0
-    var from
-      int
-    <
-    var from
-      int
-    var to
-      int
-    <=
-    var to
-      int
-    var xs
-      int
-    implies
-    and 3
-    <=
-    int-num 0
-    var from2
-      int
-    <
-    var from2
-      int
-    var to2
-      int
-    <=
-    var to2
-      int
-    var xs
-      int
-    implies
-    or 2
-    <=
-    var to2
-      int
-    var from
-      int
-    <=
-    var to
-      int
-    var from2
-      int
-    =
-    fun $bv_extract_unsigned 4
-    fun $bv_update 5
-    var x
-      int
-    var xs
-      int
-    var from
-      int
-    var to
-      int
-    var val
-      int
-    var xs
-      int
-    var from2
-      int
-    var to2
-      int
-    fun $bv_extract_unsigned 4
-    var x
-      int
-    var xs
-      int
-    var from2
-      int
-    var to2
-      int
-axiom 0
-    forall 3 1 3
-      var from
-        int
-      var to
-        int
-      var xs
-        int
-      pat 1
-        fun $bv_extract_signed 4
-        int-num 0
-        var xs
-          int
-        var from
-          int
-        var to
-          int
-      attribute qid 1
-        string-attr VccPrelu.2379:15
-      attribute uniqueId 1
-        string-attr 294
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 3
-    <=
-    int-num 0
-    var from
-      int
-    <
-    var from
-      int
-    var to
-      int
-    <=
-    var to
-      int
-    var xs
-      int
-    =
-    fun $bv_extract_signed 4
-    int-num 0
-    var xs
-      int
-    var from
-      int
-    var to
-      int
-    int-num 0
-axiom 0
-    forall 3 1 3
-      var from
-        int
-      var to
-        int
-      var xs
-        int
-      pat 1
-        fun $bv_extract_unsigned 4
-        int-num 0
-        var xs
-          int
-        var from
-          int
-        var to
-          int
-      attribute qid 1
-        string-attr VccPrelu.2384:15
-      attribute uniqueId 1
-        string-attr 295
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 3
-    <=
-    int-num 0
-    var from
-      int
-    <
-    var from
-      int
-    var to
-      int
-    <=
-    var to
-      int
-    var xs
-      int
-    =
-    fun $bv_extract_unsigned 4
-    int-num 0
-    var xs
-      int
-    var from
-      int
-    var to
-      int
-    int-num 0
-axiom 0
-    forall 4 1 3
-      var from
-        int
-      var to
-        int
-      var val
-        int
-      var xs
-        int
-      pat 1
-        fun $bv_extract_unsigned 4
-        var val
-          int
-        var xs
-          int
-        var from
-          int
-        var to
-          int
-      attribute qid 1
-        string-attr VccPrelu.2389:15
-      attribute uniqueId 1
-        string-attr 296
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 4
-    <=
-    int-num 0
-    var from
-      int
-    <
-    var from
-      int
-    var to
-      int
-    <=
-    var to
-      int
-    var xs
-      int
-    <=
-    int-num 0
-    var val
-      int
-    =
-    fun $bv_extract_unsigned 4
-    var val
-      int
-    var xs
-      int
-    var from
-      int
-    var to
-      int
-    %
-    /
-    var val
-      int
-    fun $_pow2 1
-    var from
-      int
-    fun $_pow2 1
-    -
-    var to
-      int
-    var from
-      int
-axiom 0
-    forall 4 1 3
-      var from
-        int
-      var to
-        int
-      var val
-        int
-      var xs
-        int
-      pat 1
-        fun $bv_extract_signed 4
-        var val
-          int
-        var xs
-          int
-        var from
-          int
-        var to
-          int
-      attribute qid 1
-        string-attr VccPrelu.2394:15
-      attribute uniqueId 1
-        string-attr 297
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 5
-    <=
-    int-num 0
-    var from
-      int
-    <
-    var from
-      int
-    var to
-      int
-    <=
-    var to
-      int
-    var xs
-      int
-    <=
-    int-num 0
-    var val
-      int
-    <
-    %
-    /
-    var val
-      int
-    fun $_pow2 1
-    var from
-      int
-    fun $_pow2 1
-    -
-    var to
-      int
-    var from
-      int
-    fun $_pow2 1
-    -
-    -
-    var to
-      int
-    var from
-      int
-    int-num 1
-    =
-    fun $bv_extract_signed 4
-    var val
-      int
-    var xs
-      int
-    var from
-      int
-    var to
-      int
-    %
-    /
-    var val
-      int
-    fun $_pow2 1
-    var from
-      int
-    fun $_pow2 1
-    -
-    var to
-      int
-    var from
-      int
-axiom 0
-    forall 4 1 3
-      var from
-        int
-      var to
-        int
-      var val
-        int
-      var xs
-        int
-      pat 1
-        fun $bv_extract_signed 4
-        var val
-          int
-        var xs
-          int
-        var from
-          int
-        var to
-          int
-      attribute qid 1
-        string-attr VccPrelu.2399:15
-      attribute uniqueId 1
-        string-attr 298
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 5
-    <=
-    int-num 0
-    var from
-      int
-    <
-    var from
-      int
-    var to
-      int
-    <=
-    var to
-      int
-    var xs
-      int
-    <=
-    int-num 0
-    var val
-      int
-    >=
-    %
-    /
-    var val
-      int
-    fun $_pow2 1
-    var from
-      int
-    fun $_pow2 1
-    -
-    var to
-      int
-    var from
-      int
-    fun $_pow2 1
-    -
-    -
-    var to
-      int
-    var from
-      int
-    int-num 1
-    =
-    fun $bv_extract_signed 4
-    var val
-      int
-    var xs
-      int
-    var from
-      int
-    var to
-      int
-    -
-    fun $_pow2 1
-    -
-    -
-    var to
-      int
-    var from
-      int
-    int-num 1
-    %
-    /
-    var val
-      int
-    fun $_pow2 1
-    var from
-      int
-    fun $_pow2 1
-    -
-    var to
-      int
-    var from
-      int
-axiom 0
-    forall 1 1 3
-      var val
-        int
-      pat 1
-        fun $in_range_t 2
-        fun ^^i1 0
-        var val
-          int
-      attribute qid 1
-        string-attr VccPrelu.2427:15
-      attribute uniqueId 1
-        string-attr 299
-      attribute bvZ3Native 1
-        string-attr False
-    =
-    fun $in_range_t 2
-    fun ^^i1 0
-    var val
-      int
-    and 2
-    <=
-    fun $min.i1 0
-    var val
-      int
-    <=
-    var val
-      int
-    fun $max.i1 0
-axiom 0
-    forall 1 1 3
-      var val
-        int
-      pat 1
-        fun $in_range_t 2
-        fun ^^i2 0
-        var val
-          int
-      attribute qid 1
-        string-attr VccPrelu.2428:15
-      attribute uniqueId 1
-        string-attr 300
-      attribute bvZ3Native 1
-        string-attr False
-    =
-    fun $in_range_t 2
-    fun ^^i2 0
-    var val
-      int
-    and 2
-    <=
-    fun $min.i2 0
-    var val
-      int
-    <=
-    var val
-      int
-    fun $max.i2 0
-axiom 0
-    forall 1 1 3
-      var val
-        int
-      pat 1
-        fun $in_range_t 2
-        fun ^^i4 0
-        var val
-          int
-      attribute qid 1
-        string-attr VccPrelu.2429:15
-      attribute uniqueId 1
-        string-attr 301
-      attribute bvZ3Native 1
-        string-attr False
-    =
-    fun $in_range_t 2
-    fun ^^i4 0
-    var val
-      int
-    and 2
-    <=
-    fun $min.i4 0
-    var val
-      int
-    <=
-    var val
-      int
-    fun $max.i4 0
-axiom 0
-    forall 1 1 3
-      var val
-        int
-      pat 1
-        fun $in_range_t 2
-        fun ^^i8 0
-        var val
-          int
-      attribute qid 1
-        string-attr VccPrelu.2430:15
-      attribute uniqueId 1
-        string-attr 302
-      attribute bvZ3Native 1
-        string-attr False
-    =
-    fun $in_range_t 2
-    fun ^^i8 0
-    var val
-      int
-    and 2
-    <=
-    fun $min.i8 0
-    var val
-      int
-    <=
-    var val
-      int
-    fun $max.i8 0
-axiom 0
-    forall 1 1 3
-      var val
-        int
-      pat 1
-        fun $in_range_t 2
-        fun ^^u1 0
-        var val
-          int
-      attribute qid 1
-        string-attr VccPrelu.2431:15
-      attribute uniqueId 1
-        string-attr 303
-      attribute bvZ3Native 1
-        string-attr False
-    =
-    fun $in_range_t 2
-    fun ^^u1 0
-    var val
-      int
-    and 2
-    <=
-    int-num 0
-    var val
-      int
-    <=
-    var val
-      int
-    fun $max.u1 0
-axiom 0
-    forall 1 1 3
-      var val
-        int
-      pat 1
-        fun $in_range_t 2
-        fun ^^u2 0
-        var val
-          int
-      attribute qid 1
-        string-attr VccPrelu.2432:15
-      attribute uniqueId 1
-        string-attr 304
-      attribute bvZ3Native 1
-        string-attr False
-    =
-    fun $in_range_t 2
-    fun ^^u2 0
-    var val
-      int
-    and 2
-    <=
-    int-num 0
-    var val
-      int
-    <=
-    var val
-      int
-    fun $max.u2 0
-axiom 0
-    forall 1 1 3
-      var val
-        int
-      pat 1
-        fun $in_range_t 2
-        fun ^^u4 0
-        var val
-          int
-      attribute qid 1
-        string-attr VccPrelu.2433:15
-      attribute uniqueId 1
-        string-attr 305
-      attribute bvZ3Native 1
-        string-attr False
-    =
-    fun $in_range_t 2
-    fun ^^u4 0
-    var val
-      int
-    and 2
-    <=
-    int-num 0
-    var val
-      int
-    <=
-    var val
-      int
-    fun $max.u4 0
-axiom 0
-    forall 1 1 3
-      var val
-        int
-      pat 1
-        fun $in_range_t 2
-        fun ^^u8 0
-        var val
-          int
-      attribute qid 1
-        string-attr VccPrelu.2434:15
-      attribute uniqueId 1
-        string-attr 306
-      attribute bvZ3Native 1
-        string-attr False
-    =
-    fun $in_range_t 2
-    fun ^^u8 0
-    var val
-      int
-    and 2
-    <=
-    int-num 0
-    var val
-      int
-    <=
-    var val
-      int
-    fun $max.u8 0
-axiom 0
-    forall 2 1 3
-      var t
-        type-con $ctype 0
-      var val
-        int
-      pat 1
-        fun $unchecked 2
-        var t
-          type-con $ctype 0
-        var val
-          int
-      attribute qid 1
-        string-attr VccPrelu.2436:15
-      attribute uniqueId 1
-        string-attr 307
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    fun $in_range_t 2
-    var t
-      type-con $ctype 0
-    var val
-      int
-    =
-    fun $unchecked 2
-    var t
-      type-con $ctype 0
-    var val
-      int
-    var val
-      int
-axiom 0
-    forall 2 1 3
-      var t
-        type-con $ctype 0
-      var val
-        int
-      pat 1
-        fun $unchecked 2
-        var t
-          type-con $ctype 0
-        var val
-          int
-      attribute qid 1
-        string-attr VccPrelu.2437:15
-      attribute uniqueId 1
-        string-attr 308
-      attribute bvZ3Native 1
-        string-attr False
-    fun $in_range_t 2
-    var t
-      type-con $ctype 0
-    fun $unchecked 2
-    var t
-      type-con $ctype 0
-    var val
-      int
-axiom 0
-    forall 1 1 3
-      var val
-        int
-      pat 1
-        fun $unchecked 2
-        fun ^^u1 0
-        fun $unchecked 2
-        fun ^^i1 0
-        var val
-          int
-      attribute qid 1
-        string-attr VccPrelu.2439:15
-      attribute uniqueId 1
-        string-attr 309
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 2
-    <=
-    int-num 0
-    var val
-      int
-    <=
-    var val
-      int
-    fun $max.u1 0
-    =
-    fun $unchecked 2
-    fun ^^u1 0
-    fun $unchecked 2
-    fun ^^i1 0
-    var val
-      int
-    var val
-      int
-axiom 0
-    forall 1 1 3
-      var val
-        int
-      pat 1
-        fun $unchecked 2
-        fun ^^u2 0
-        fun $unchecked 2
-        fun ^^i2 0
-        var val
-          int
-      attribute qid 1
-        string-attr VccPrelu.2440:15
-      attribute uniqueId 1
-        string-attr 310
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 2
-    <=
-    int-num 0
-    var val
-      int
-    <=
-    var val
-      int
-    fun $max.u2 0
-    =
-    fun $unchecked 2
-    fun ^^u2 0
-    fun $unchecked 2
-    fun ^^i2 0
-    var val
-      int
-    var val
-      int
-axiom 0
-    forall 1 1 3
-      var val
-        int
-      pat 1
-        fun $unchecked 2
-        fun ^^u4 0
-        fun $unchecked 2
-        fun ^^i4 0
-        var val
-          int
-      attribute qid 1
-        string-attr VccPrelu.2441:15
-      attribute uniqueId 1
-        string-attr 311
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 2
-    <=
-    int-num 0
-    var val
-      int
-    <=
-    var val
-      int
-    fun $max.u4 0
-    =
-    fun $unchecked 2
-    fun ^^u4 0
-    fun $unchecked 2
-    fun ^^i4 0
-    var val
-      int
-    var val
-      int
-axiom 0
-    forall 1 1 3
-      var val
-        int
-      pat 1
-        fun $unchecked 2
-        fun ^^u8 0
-        fun $unchecked 2
-        fun ^^i8 0
-        var val
-          int
-      attribute qid 1
-        string-attr VccPrelu.2442:15
-      attribute uniqueId 1
-        string-attr 312
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 2
-    <=
-    int-num 0
-    var val
-      int
-    <=
-    var val
-      int
-    fun $max.u8 0
-    =
-    fun $unchecked 2
-    fun ^^u8 0
-    fun $unchecked 2
-    fun ^^i8 0
-    var val
-      int
-    var val
-      int
-axiom 0
-    forall 1 1 3
-      var val
-        int
-      pat 1
-        fun $unchecked 2
-        fun ^^i1 0
-        fun $unchecked 2
-        fun ^^u1 0
-        var val
-          int
-      attribute qid 1
-        string-attr VccPrelu.2443:15
-      attribute uniqueId 1
-        string-attr 313
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 2
-    <=
-    fun $min.i1 0
-    var val
-      int
-    <=
-    var val
-      int
-    fun $max.i1 0
-    =
-    fun $unchecked 2
-    fun ^^i1 0
-    fun $unchecked 2
-    fun ^^u1 0
-    var val
-      int
-    var val
-      int
-axiom 0
-    forall 1 1 3
-      var val
-        int
-      pat 1
-        fun $unchecked 2
-        fun ^^i2 0
-        fun $unchecked 2
-        fun ^^u2 0
-        var val
-          int
-      attribute qid 1
-        string-attr VccPrelu.2444:15
-      attribute uniqueId 1
-        string-attr 314
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 2
-    <=
-    fun $min.i2 0
-    var val
-      int
-    <=
-    var val
-      int
-    fun $max.i2 0
-    =
-    fun $unchecked 2
-    fun ^^i2 0
-    fun $unchecked 2
-    fun ^^u2 0
-    var val
-      int
-    var val
-      int
-axiom 0
-    forall 1 1 3
-      var val
-        int
-      pat 1
-        fun $unchecked 2
-        fun ^^i4 0
-        fun $unchecked 2
-        fun ^^u4 0
-        var val
-          int
-      attribute qid 1
-        string-attr VccPrelu.2445:15
-      attribute uniqueId 1
-        string-attr 315
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 2
-    <=
-    fun $min.i4 0
-    var val
-      int
-    <=
-    var val
-      int
-    fun $max.i4 0
-    =
-    fun $unchecked 2
-    fun ^^i4 0
-    fun $unchecked 2
-    fun ^^u4 0
-    var val
-      int
-    var val
-      int
-axiom 0
-    forall 1 1 3
-      var val
-        int
-      pat 1
-        fun $unchecked 2
-        fun ^^i8 0
-        fun $unchecked 2
-        fun ^^u8 0
-        var val
-          int
-      attribute qid 1
-        string-attr VccPrelu.2446:15
-      attribute uniqueId 1
-        string-attr 316
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 2
-    <=
-    fun $min.i8 0
-    var val
-      int
-    <=
-    var val
-      int
-    fun $max.i8 0
-    =
-    fun $unchecked 2
-    fun ^^i8 0
-    fun $unchecked 2
-    fun ^^u8 0
-    var val
-      int
-    var val
-      int
-axiom 0
-    forall 4 1 3
-      var t
-        type-con $ctype 0
-      var x
-        int
-      var y
-        int
-      var z
-        int
-      pat 2
-        %
-        var x
-          int
-        fun $_pow2 1
-        var y
-          int
-        fun $_and 3
-        var t
-          type-con $ctype 0
-        var x
-          int
-        var z
-          int
-      attribute qid 1
-        string-attr VccPrelu.2452:15
-      attribute uniqueId 1
-        string-attr 317
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 3
-    fun $in_range_t 2
-    var t
-      type-con $ctype 0
-    var x
-      int
-    fun $in_range_t 2
-    var t
-      type-con $ctype 0
-    -
-    fun $_pow2 1
-    var y
-      int
-    int-num 1
-    >=
-    var x
-      int
-    int-num 0
-    =
-    %
-    var x
-      int
-    fun $_pow2 1
-    var y
-      int
-    fun $_and 3
-    var t
-      type-con $ctype 0
-    var x
-      int
-    -
-    fun $_pow2 1
-    var y
-      int
-    int-num 1
-axiom 0
-    forall 2 1 3
-      var i
-        int
-      var j
-        int
-      pat 1
-        /
-        var i
-          int
-        var j
-          int
-      attribute qid 1
-        string-attr VccPrelu.2458:15
-      attribute uniqueId 1
-        string-attr 318
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 2
-    <=
-    int-num 0
-    var i
-      int
-    <
-    int-num 0
-    var j
-      int
-    <=
-    /
-    var i
-      int
-    var j
-      int
-    var i
-      int
-axiom 0
-    forall 2 1 3
-      var i
-        int
-      var j
-        int
-      pat 1
-        /
-        var i
-          int
-        var j
-          int
-      attribute qid 1
-        string-attr VccPrelu.2460:15
-      attribute uniqueId 1
-        string-attr 319
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 2
-    >
-    var i
-      int
-    int-num 0
-    >
-    var j
-      int
-    int-num 0
-    and 2
-    <
-    -
-    var i
-      int
-    var j
-      int
-    *
-    /
-    var i
-      int
-    var j
-      int
-    var j
-      int
-    <=
-    *
-    /
-    var i
-      int
-    var j
-      int
-    var j
-      int
-    var i
-      int
-axiom 0
-    forall 1 1 3
-      var i
-        int
-      pat 1
-        /
-        var i
-          int
-        var i
-          int
-      attribute qid 1
-        string-attr VccPrelu.2461:15
-      attribute uniqueId 1
-        string-attr 320
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    not
-    =
-    var i
-      int
-    int-num 0
-    =
-    /
-    var i
-      int
-    var i
-      int
-    int-num 1
-axiom 0
-    forall 2 2 3
-      var x
-        int
-      var y
-        int
-      pat 1
-        %
-        var x
-          int
-        var y
-          int
-      pat 1
-        /
-        var x
-          int
-        var y
-          int
-      attribute qid 1
-        string-attr VccPrelu.2464:15
-      attribute uniqueId 1
-        string-attr 321
-      attribute bvZ3Native 1
-        string-attr False
-    =
-    %
-    var x
-      int
-    var y
-      int
-    -
-    var x
-      int
-    *
-    /
-    var x
-      int
-    var y
-      int
-    var y
-      int
-axiom 0
-    forall 2 1 3
-      var x
-        int
-      var y
-        int
-      pat 1
-        %
-        var x
-          int
-        var y
-          int
-      attribute qid 1
-        string-attr VccPrelu.2465:15
-      attribute uniqueId 1
-        string-attr 322
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 2
-    <=
-    int-num 0
-    var x
-      int
-    <
-    int-num 0
-    var y
-      int
-    and 2
-    <=
-    int-num 0
-    %
-    var x
-      int
-    var y
-      int
-    <
-    %
-    var x
-      int
-    var y
-      int
-    var y
-      int
-axiom 0
-    forall 2 1 3
-      var x
-        int
-      var y
-        int
-      pat 1
-        %
-        var x
-          int
-        var y
-          int
-      attribute qid 1
-        string-attr VccPrelu.2466:15
-      attribute uniqueId 1
-        string-attr 323
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 2
-    <=
-    int-num 0
-    var x
-      int
-    <
-    var y
-      int
-    int-num 0
-    and 2
-    <=
-    int-num 0
-    %
-    var x
-      int
-    var y
-      int
-    <
-    %
-    var x
-      int
-    var y
-      int
-    -
-    int-num 0
-    var y
-      int
-axiom 0
-    forall 2 1 3
-      var x
-        int
-      var y
-        int
-      pat 1
-        %
-        var x
-          int
-        var y
-          int
-      attribute qid 1
-        string-attr VccPrelu.2467:15
-      attribute uniqueId 1
-        string-attr 324
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 2
-    <=
-    var x
-      int
-    int-num 0
-    <
-    int-num 0
-    var y
-      int
-    and 2
-    <
-    -
-    int-num 0
-    var y
-      int
-    %
-    var x
-      int
-    var y
-      int
-    <=
-    %
-    var x
-      int
-    var y
-      int
-    int-num 0
-axiom 0
-    forall 2 1 3
-      var x
-        int
-      var y
-        int
-      pat 1
-        %
-        var x
-          int
-        var y
-          int
-      attribute qid 1
-        string-attr VccPrelu.2468:15
-      attribute uniqueId 1
-        string-attr 325
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 2
-    <=
-    var x
-      int
-    int-num 0
-    <
-    var y
-      int
-    int-num 0
-    and 2
-    <
-    var y
-      int
-    %
-    var x
-      int
-    var y
-      int
-    <=
-    %
-    var x
-      int
-    var y
-      int
-    int-num 0
-axiom 0
-    forall 3 1 3
-      var t
-        type-con $ctype 0
-      var x
-        int
-      var y
-        int
-      pat 1
-        fun $_and 3
-        var t
-          type-con $ctype 0
-        var x
-          int
-        var y
-          int
-      attribute qid 1
-        string-attr VccPrelu.2477:15
-      attribute uniqueId 1
-        string-attr 326
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 2
-    <=
-    int-num 0
-    var x
-      int
-    fun $in_range_t 2
-    var t
-      type-con $ctype 0
-    var x
-      int
-    and 2
-    <=
-    int-num 0
-    fun $_and 3
-    var t
-      type-con $ctype 0
-    var x
-      int
-    var y
-      int
-    <=
-    fun $_and 3
-    var t
-      type-con $ctype 0
-    var x
-      int
-    var y
-      int
-    var x
-      int
-axiom 0
-    forall 3 1 3
-      var t
-        type-con $ctype 0
-      var x
-        int
-      var y
-        int
-      pat 1
-        fun $_and 3
-        var t
-          type-con $ctype 0
-        var x
-          int
-        var y
-          int
-      attribute qid 1
-        string-attr VccPrelu.2478:15
-      attribute uniqueId 1
-        string-attr 327
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 4
-    <=
-    int-num 0
-    var x
-      int
-    <=
-    int-num 0
-    var y
-      int
-    fun $in_range_t 2
-    var t
-      type-con $ctype 0
-    var x
-      int
-    fun $in_range_t 2
-    var t
-      type-con $ctype 0
-    var y
-      int
-    and 2
-    <=
-    fun $_and 3
-    var t
-      type-con $ctype 0
-    var x
-      int
-    var y
-      int
-    var x
-      int
-    <=
-    fun $_and 3
-    var t
-      type-con $ctype 0
-    var x
-      int
-    var y
-      int
-    var y
-      int
-axiom 0
-    forall 3 1 3
-      var t
-        type-con $ctype 0
-      var x
-        int
-      var y
-        int
-      pat 1
-        fun $_or 3
-        var t
-          type-con $ctype 0
-        var x
-          int
-        var y
-          int
-      attribute qid 1
-        string-attr VccPrelu.2479:15
-      attribute uniqueId 1
-        string-attr 328
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 4
-    <=
-    int-num 0
-    var x
-      int
-    <=
-    int-num 0
-    var y
-      int
-    fun $in_range_t 2
-    var t
-      type-con $ctype 0
-    var x
-      int
-    fun $in_range_t 2
-    var t
-      type-con $ctype 0
-    var y
-      int
-    and 2
-    <=
-    int-num 0
-    fun $_or 3
-    var t
-      type-con $ctype 0
-    var x
-      int
-    var y
-      int
-    <=
-    fun $_or 3
-    var t
-      type-con $ctype 0
-    var x
-      int
-    var y
-      int
-    +
-    var x
-      int
-    var y
-      int
-axiom 0
-    forall 3 1 3
-      var t
-        type-con $ctype 0
-      var x
-        int
-      var y
-        int
-      pat 1
-        fun $_or 3
-        var t
-          type-con $ctype 0
-        var x
-          int
-        var y
-          int
-      attribute qid 1
-        string-attr VccPrelu.2480:15
-      attribute uniqueId 1
-        string-attr 329
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 4
-    <=
-    int-num 0
-    var x
-      int
-    <=
-    int-num 0
-    var y
-      int
-    fun $in_range_t 2
-    var t
-      type-con $ctype 0
-    var x
-      int
-    fun $in_range_t 2
-    var t
-      type-con $ctype 0
-    var y
-      int
-    and 2
-    <=
-    var x
-      int
-    fun $_or 3
-    var t
-      type-con $ctype 0
-    var x
-      int
-    var y
-      int
-    <=
-    var y
-      int
-    fun $_or 3
-    var t
-      type-con $ctype 0
-    var x
-      int
-    var y
-      int
-axiom 0
-    forall 4 1 3
-      var t
-        type-con $ctype 0
-      var x
-        int
-      var y
-        int
-      var z
-        int
-      pat 2
-        fun $_or 3
-        var t
-          type-con $ctype 0
-        var x
-          int
-        var y
-          int
-        fun $_pow2 1
-        var z
-          int
-      attribute qid 1
-        string-attr VccPrelu.2481:15
-      attribute uniqueId 1
-        string-attr 330
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 8
-    <=
-    int-num 0
-    var x
-      int
-    <=
-    int-num 0
-    var y
-      int
-    <=
-    int-num 0
-    var z
-      int
-    <
-    var z
-      int
-    int-num 64
-    <
-    var x
-      int
-    fun $_pow2 1
-    var z
-      int
-    <
-    var y
-      int
-    fun $_pow2 1
-    var z
-      int
-    fun $in_range_t 2
-    var t
-      type-con $ctype 0
-    var x
-      int
-    fun $in_range_t 2
-    var t
-      type-con $ctype 0
-    var y
-      int
-    <
-    fun $_or 3
-    var t
-      type-con $ctype 0
-    var x
-      int
-    var y
-      int
-    fun $_pow2 1
-    var z
-      int
-axiom 0
-    forall 3 1 3
-      var t
-        type-con $ctype 0
-      var x
-        int
-      var y
-        int
-      pat 1
-        fun $_or 3
-        var t
-          type-con $ctype 0
-        var x
-          int
-        var y
-          int
-      attribute qid 1
-        string-attr VccPrelu.2484:15
-      attribute uniqueId 1
-        string-attr 331
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 4
-    <=
-    int-num 0
-    var x
-      int
-    <=
-    var x
-      int
-    fun $max.u1 0
-    <=
-    int-num 0
-    var y
-      int
-    <=
-    var y
-      int
-    fun $max.u1 0
-    and 2
-    <=
-    int-num 0
-    fun $_or 3
-    var t
-      type-con $ctype 0
-    var x
-      int
-    var y
-      int
-    <=
-    fun $_or 3
-    var t
-      type-con $ctype 0
-    var x
-      int
-    var y
-      int
-    fun $max.u1 0
-axiom 0
-    forall 3 1 3
-      var t
-        type-con $ctype 0
-      var x
-        int
-      var y
-        int
-      pat 1
-        fun $_or 3
-        var t
-          type-con $ctype 0
-        var x
-          int
-        var y
-          int
-      attribute qid 1
-        string-attr VccPrelu.2485:15
-      attribute uniqueId 1
-        string-attr 332
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 4
-    <=
-    int-num 0
-    var x
-      int
-    <=
-    var x
-      int
-    fun $max.u2 0
-    <=
-    int-num 0
-    var y
-      int
-    <=
-    var y
-      int
-    fun $max.u2 0
-    and 2
-    <=
-    int-num 0
-    fun $_or 3
-    var t
-      type-con $ctype 0
-    var x
-      int
-    var y
-      int
-    <=
-    fun $_or 3
-    var t
-      type-con $ctype 0
-    var x
-      int
-    var y
-      int
-    fun $max.u2 0
-axiom 0
-    forall 3 1 3
-      var t
-        type-con $ctype 0
-      var x
-        int
-      var y
-        int
-      pat 1
-        fun $_or 3
-        var t
-          type-con $ctype 0
-        var x
-          int
-        var y
-          int
-      attribute qid 1
-        string-attr VccPrelu.2486:15
-      attribute uniqueId 1
-        string-attr 333
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 4
-    <=
-    int-num 0
-    var x
-      int
-    <=
-    var x
-      int
-    fun $max.u4 0
-    <=
-    int-num 0
-    var y
-      int
-    <=
-    var y
-      int
-    fun $max.u4 0
-    and 2
-    <=
-    int-num 0
-    fun $_or 3
-    var t
-      type-con $ctype 0
-    var x
-      int
-    var y
-      int
-    <=
-    fun $_or 3
-    var t
-      type-con $ctype 0
-    var x
-      int
-    var y
-      int
-    fun $max.u4 0
-axiom 0
-    forall 3 1 3
-      var t
-        type-con $ctype 0
-      var x
-        int
-      var y
-        int
-      pat 1
-        fun $_or 3
-        var t
-          type-con $ctype 0
-        var x
-          int
-        var y
-          int
-      attribute qid 1
-        string-attr VccPrelu.2487:15
-      attribute uniqueId 1
-        string-attr 334
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 4
-    <=
-    int-num 0
-    var x
-      int
-    <=
-    var x
-      int
-    fun $max.u8 0
-    <=
-    int-num 0
-    var y
-      int
-    <=
-    var y
-      int
-    fun $max.u8 0
-    and 2
-    <=
-    int-num 0
-    fun $_or 3
-    var t
-      type-con $ctype 0
-    var x
-      int
-    var y
-      int
-    <=
-    fun $_or 3
-    var t
-      type-con $ctype 0
-    var x
-      int
-    var y
-      int
-    fun $max.u8 0
-axiom 0
-    forall 3 1 3
-      var t
-        type-con $ctype 0
-      var x
-        int
-      var y
-        int
-      pat 1
-        fun $_and 3
-        var t
-          type-con $ctype 0
-        var x
-          int
-        var y
-          int
-      attribute qid 1
-        string-attr VccPrelu.2489:15
-      attribute uniqueId 1
-        string-attr 335
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 4
-    <=
-    int-num 0
-    var x
-      int
-    <=
-    var x
-      int
-    fun $max.u1 0
-    <=
-    int-num 0
-    var y
-      int
-    <=
-    var y
-      int
-    fun $max.u1 0
-    and 2
-    <=
-    int-num 0
-    fun $_and 3
-    var t
-      type-con $ctype 0
-    var x
-      int
-    var y
-      int
-    <=
-    fun $_and 3
-    var t
-      type-con $ctype 0
-    var x
-      int
-    var y
-      int
-    fun $max.u1 0
-axiom 0
-    forall 3 1 3
-      var t
-        type-con $ctype 0
-      var x
-        int
-      var y
-        int
-      pat 1
-        fun $_and 3
-        var t
-          type-con $ctype 0
-        var x
-          int
-        var y
-          int
-      attribute qid 1
-        string-attr VccPrelu.2490:15
-      attribute uniqueId 1
-        string-attr 336
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 4
-    <=
-    int-num 0
-    var x
-      int
-    <=
-    var x
-      int
-    fun $max.u2 0
-    <=
-    int-num 0
-    var y
-      int
-    <=
-    var y
-      int
-    fun $max.u2 0
-    and 2
-    <=
-    int-num 0
-    fun $_and 3
-    var t
-      type-con $ctype 0
-    var x
-      int
-    var y
-      int
-    <=
-    fun $_and 3
-    var t
-      type-con $ctype 0
-    var x
-      int
-    var y
-      int
-    fun $max.u2 0
-axiom 0
-    forall 3 1 3
-      var t
-        type-con $ctype 0
-      var x
-        int
-      var y
-        int
-      pat 1
-        fun $_and 3
-        var t
-          type-con $ctype 0
-        var x
-          int
-        var y
-          int
-      attribute qid 1
-        string-attr VccPrelu.2491:15
-      attribute uniqueId 1
-        string-attr 337
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 4
-    <=
-    int-num 0
-    var x
-      int
-    <=
-    var x
-      int
-    fun $max.u4 0
-    <=
-    int-num 0
-    var y
-      int
-    <=
-    var y
-      int
-    fun $max.u4 0
-    and 2
-    <=
-    int-num 0
-    fun $_and 3
-    var t
-      type-con $ctype 0
-    var x
-      int
-    var y
-      int
-    <=
-    fun $_and 3
-    var t
-      type-con $ctype 0
-    var x
-      int
-    var y
-      int
-    fun $max.u4 0
-axiom 0
-    forall 3 1 3
-      var t
-        type-con $ctype 0
-      var x
-        int
-      var y
-        int
-      pat 1
-        fun $_and 3
-        var t
-          type-con $ctype 0
-        var x
-          int
-        var y
-          int
-      attribute qid 1
-        string-attr VccPrelu.2492:15
-      attribute uniqueId 1
-        string-attr 338
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 4
-    <=
-    int-num 0
-    var x
-      int
-    <=
-    var x
-      int
-    fun $max.u8 0
-    <=
-    int-num 0
-    var y
-      int
-    <=
-    var y
-      int
-    fun $max.u8 0
-    and 2
-    <=
-    int-num 0
-    fun $_and 3
-    var t
-      type-con $ctype 0
-    var x
-      int
-    var y
-      int
-    <=
-    fun $_and 3
-    var t
-      type-con $ctype 0
-    var x
-      int
-    var y
-      int
-    fun $max.u8 0
-axiom 0
-    forall 3 1 3
-      var t
-        type-con $ctype 0
-      var x
-        int
-      var y
-        int
-      pat 1
-        fun $_xor 3
-        var t
-          type-con $ctype 0
-        var x
-          int
-        var y
-          int
-      attribute qid 1
-        string-attr VccPrelu.2494:15
-      attribute uniqueId 1
-        string-attr 339
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 4
-    <=
-    int-num 0
-    var x
-      int
-    <=
-    var x
-      int
-    fun $max.u1 0
-    <=
-    int-num 0
-    var y
-      int
-    <=
-    var y
-      int
-    fun $max.u1 0
-    and 2
-    <=
-    int-num 0
-    fun $_xor 3
-    var t
-      type-con $ctype 0
-    var x
-      int
-    var y
-      int
-    <=
-    fun $_xor 3
-    var t
-      type-con $ctype 0
-    var x
-      int
-    var y
-      int
-    fun $max.u1 0
-axiom 0
-    forall 3 1 3
-      var t
-        type-con $ctype 0
-      var x
-        int
-      var y
-        int
-      pat 1
-        fun $_xor 3
-        var t
-          type-con $ctype 0
-        var x
-          int
-        var y
-          int
-      attribute qid 1
-        string-attr VccPrelu.2495:15
-      attribute uniqueId 1
-        string-attr 340
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 4
-    <=
-    int-num 0
-    var x
-      int
-    <=
-    var x
-      int
-    fun $max.u2 0
-    <=
-    int-num 0
-    var y
-      int
-    <=
-    var y
-      int
-    fun $max.u2 0
-    and 2
-    <=
-    int-num 0
-    fun $_xor 3
-    var t
-      type-con $ctype 0
-    var x
-      int
-    var y
-      int
-    <=
-    fun $_xor 3
-    var t
-      type-con $ctype 0
-    var x
-      int
-    var y
-      int
-    fun $max.u2 0
-axiom 0
-    forall 3 1 3
-      var t
-        type-con $ctype 0
-      var x
-        int
-      var y
-        int
-      pat 1
-        fun $_xor 3
-        var t
-          type-con $ctype 0
-        var x
-          int
-        var y
-          int
-      attribute qid 1
-        string-attr VccPrelu.2496:15
-      attribute uniqueId 1
-        string-attr 341
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 4
-    <=
-    int-num 0
-    var x
-      int
-    <=
-    var x
-      int
-    fun $max.u4 0
-    <=
-    int-num 0
-    var y
-      int
-    <=
-    var y
-      int
-    fun $max.u4 0
-    and 2
-    <=
-    int-num 0
-    fun $_xor 3
-    var t
-      type-con $ctype 0
-    var x
-      int
-    var y
-      int
-    <=
-    fun $_xor 3
-    var t
-      type-con $ctype 0
-    var x
-      int
-    var y
-      int
-    fun $max.u4 0
-axiom 0
-    forall 3 1 3
-      var t
-        type-con $ctype 0
-      var x
-        int
-      var y
-        int
-      pat 1
-        fun $_xor 3
-        var t
-          type-con $ctype 0
-        var x
-          int
-        var y
-          int
-      attribute qid 1
-        string-attr VccPrelu.2497:15
-      attribute uniqueId 1
-        string-attr 342
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 4
-    <=
-    int-num 0
-    var x
-      int
-    <=
-    var x
-      int
-    fun $max.u8 0
-    <=
-    int-num 0
-    var y
-      int
-    <=
-    var y
-      int
-    fun $max.u8 0
-    and 2
-    <=
-    int-num 0
-    fun $_xor 3
-    var t
-      type-con $ctype 0
-    var x
-      int
-    var y
-      int
-    <=
-    fun $_xor 3
-    var t
-      type-con $ctype 0
-    var x
-      int
-    var y
-      int
-    fun $max.u8 0
-axiom 0
-    forall 2 1 3
-      var t
-        type-con $ctype 0
-      var x
-        int
-      pat 1
-        fun $_not 2
-        var t
-          type-con $ctype 0
-        var x
-          int
-      attribute qid 1
-        string-attr VccPrelu.2499:15
-      attribute uniqueId 1
-        string-attr 343
-      attribute bvZ3Native 1
-        string-attr False
-    fun $in_range_t 2
-    var t
-      type-con $ctype 0
-    fun $_not 2
-    var t
-      type-con $ctype 0
-    var x
-      int
-axiom 0
-    forall 2 1 3
-      var t
-        type-con $ctype 0
-      var x
-        int
-      pat 1
-        fun $_or 3
-        var t
-          type-con $ctype 0
-        var x
-          int
-        fun $_not 2
-        var t
-          type-con $ctype 0
-        var x
-          int
-      attribute qid 1
-        string-attr VccPrelu.2504:15
-      attribute uniqueId 1
-        string-attr 344
-      attribute bvZ3Native 1
-        string-attr False
-    =
-    fun $_or 3
-    var t
-      type-con $ctype 0
-    var x
-      int
-    fun $_not 2
-    var t
-      type-con $ctype 0
-    var x
-      int
-    fun $_not 2
-    var t
-      type-con $ctype 0
-    int-num 0
-axiom 0
-    forall 2 1 3
-      var t
-        type-con $ctype 0
-      var x
-        int
-      pat 1
-        fun $_and 3
-        var t
-          type-con $ctype 0
-        var x
-          int
-        fun $_not 2
-        var t
-          type-con $ctype 0
-        var x
-          int
-      attribute qid 1
-        string-attr VccPrelu.2505:15
-      attribute uniqueId 1
-        string-attr 345
-      attribute bvZ3Native 1
-        string-attr False
-    =
-    fun $_and 3
-    var t
-      type-con $ctype 0
-    var x
-      int
-    fun $_not 2
-    var t
-      type-con $ctype 0
-    var x
-      int
-    int-num 0
-axiom 0
-    forall 2 1 3
-      var t
-        type-con $ctype 0
-      var x
-        int
-      pat 1
-        fun $_or 3
-        var t
-          type-con $ctype 0
-        var x
-          int
-        int-num 0
-      attribute qid 1
-        string-attr VccPrelu.2506:15
-      attribute uniqueId 1
-        string-attr 346
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    fun $in_range_t 2
-    var t
-      type-con $ctype 0
-    var x
-      int
-    =
-    fun $_or 3
-    var t
-      type-con $ctype 0
-    var x
-      int
-    int-num 0
-    var x
-      int
-axiom 0
-    forall 2 1 3
-      var t
-        type-con $ctype 0
-      var x
-        int
-      pat 1
-        fun $_or 3
-        var t
-          type-con $ctype 0
-        var x
-          int
-        fun $_not 2
-        var t
-          type-con $ctype 0
-        int-num 0
-      attribute qid 1
-        string-attr VccPrelu.2507:15
-      attribute uniqueId 1
-        string-attr 347
-      attribute bvZ3Native 1
-        string-attr False
-    =
-    fun $_or 3
-    var t
-      type-con $ctype 0
-    var x
-      int
-    fun $_not 2
-    var t
-      type-con $ctype 0
-    int-num 0
-    fun $_not 2
-    var t
-      type-con $ctype 0
-    int-num 0
-axiom 0
-    forall 2 1 3
-      var t
-        type-con $ctype 0
-      var x
-        int
-      pat 1
-        fun $_or 3
-        var t
-          type-con $ctype 0
-        var x
-          int
-        var x
-          int
-      attribute qid 1
-        string-attr VccPrelu.2508:15
-      attribute uniqueId 1
-        string-attr 348
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    fun $in_range_t 2
-    var t
-      type-con $ctype 0
-    var x
-      int
-    =
-    fun $_or 3
-    var t
-      type-con $ctype 0
-    var x
-      int
-    var x
-      int
-    var x
-      int
-axiom 0
-    forall 2 1 3
-      var t
-        type-con $ctype 0
-      var x
-        int
-      pat 1
-        fun $_and 3
-        var t
-          type-con $ctype 0
-        var x
-          int
-        int-num 0
-      attribute qid 1
-        string-attr VccPrelu.2509:15
-      attribute uniqueId 1
-        string-attr 349
-      attribute bvZ3Native 1
-        string-attr False
-    =
-    fun $_and 3
-    var t
-      type-con $ctype 0
-    var x
-      int
-    int-num 0
-    int-num 0
-axiom 0
-    forall 2 1 3
-      var t
-        type-con $ctype 0
-      var x
-        int
-      pat 1
-        fun $_and 3
-        var t
-          type-con $ctype 0
-        var x
-          int
-        fun $_not 2
-        var t
-          type-con $ctype 0
-        int-num 0
-      attribute qid 1
-        string-attr VccPrelu.2510:15
-      attribute uniqueId 1
-        string-attr 350
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    fun $in_range_t 2
-    var t
-      type-con $ctype 0
-    var x
-      int
-    =
-    fun $_and 3
-    var t
-      type-con $ctype 0
-    var x
-      int
-    fun $_not 2
-    var t
-      type-con $ctype 0
-    int-num 0
-    var x
-      int
-axiom 0
-    forall 2 1 3
-      var t
-        type-con $ctype 0
-      var x
-        int
-      pat 1
-        fun $_and 3
-        var t
-          type-con $ctype 0
-        var x
-          int
-        var x
-          int
-      attribute qid 1
-        string-attr VccPrelu.2511:15
-      attribute uniqueId 1
-        string-attr 351
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    fun $in_range_t 2
-    var t
-      type-con $ctype 0
-    var x
-      int
-    =
-    fun $_and 3
-    var t
-      type-con $ctype 0
-    var x
-      int
-    var x
-      int
-    var x
-      int
-axiom 0
-    forall 3 1 3
-      var t
-        type-con $ctype 0
-      var x
-        int
-      var y
-        int
-      pat 1
-        fun $_and 3
-        var t
-          type-con $ctype 0
-        fun $_or 3
-        var t
-          type-con $ctype 0
-        var x
-          int
-        var y
-          int
-        var y
-          int
-      attribute qid 1
-        string-attr VccPrelu.2512:15
-      attribute uniqueId 1
-        string-attr 352
-      attribute bvZ3Native 1
-        string-attr False
-    =
-    fun $_and 3
-    var t
-      type-con $ctype 0
-    fun $_or 3
-    var t
-      type-con $ctype 0
-    var x
-      int
-    var y
-      int
-    var y
-      int
-    var y
-      int
-axiom 0
-    forall 3 1 3
-      var t
-        type-con $ctype 0
-      var x
-        int
-      var y
-        int
-      pat 1
-        fun $_and 3
-        var t
-          type-con $ctype 0
-        fun $_or 3
-        var t
-          type-con $ctype 0
-        var x
-          int
-        var y
-          int
-        var x
-          int
-      attribute qid 1
-        string-attr VccPrelu.2513:15
-      attribute uniqueId 1
-        string-attr 353
-      attribute bvZ3Native 1
-        string-attr False
-    =
-    fun $_and 3
-    var t
-      type-con $ctype 0
-    fun $_or 3
-    var t
-      type-con $ctype 0
-    var x
-      int
-    var y
-      int
-    var x
-      int
-    var x
-      int
-axiom 0
-    forall 2 1 3
-      var t
-        type-con $ctype 0
-      var x
-        int
-      pat 1
-        fun $_xor 3
-        var t
-          type-con $ctype 0
-        var x
-          int
-        int-num 0
-      attribute qid 1
-        string-attr VccPrelu.2514:15
-      attribute uniqueId 1
-        string-attr 354
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    fun $in_range_t 2
-    var t
-      type-con $ctype 0
-    var x
-      int
-    =
-    fun $_xor 3
-    var t
-      type-con $ctype 0
-    var x
-      int
-    int-num 0
-    var x
-      int
-axiom 0
-    forall 2 1 3
-      var t
-        type-con $ctype 0
-      var x
-        int
-      pat 1
-        fun $_xor 3
-        var t
-          type-con $ctype 0
-        var x
-          int
-        var x
-          int
-      attribute qid 1
-        string-attr VccPrelu.2515:15
-      attribute uniqueId 1
-        string-attr 355
-      attribute bvZ3Native 1
-        string-attr False
-    =
-    fun $_xor 3
-    var t
-      type-con $ctype 0
-    var x
-      int
-    var x
-      int
-    int-num 0
-axiom 0
-    forall 2 1 3
-      var t
-        type-con $ctype 0
-      var x
-        int
-      pat 1
-        fun $_xor 3
-        var t
-          type-con $ctype 0
-        var x
-          int
-        fun $_not 2
-        var t
-          type-con $ctype 0
-        int-num 0
-      attribute qid 1
-        string-attr VccPrelu.2516:15
-      attribute uniqueId 1
-        string-attr 356
-      attribute bvZ3Native 1
-        string-attr False
-    =
-    fun $_xor 3
-    var t
-      type-con $ctype 0
-    var x
-      int
-    fun $_not 2
-    var t
-      type-con $ctype 0
-    int-num 0
-    fun $_not 2
-    var t
-      type-con $ctype 0
-    var x
-      int
-axiom 0
-    forall 2 1 3
-      var t
-        type-con $ctype 0
-      var x
-        int
-      pat 1
-        fun $_not 2
-        var t
-          type-con $ctype 0
-        fun $_not 2
-        var t
-          type-con $ctype 0
-        var x
-          int
-      attribute qid 1
-        string-attr VccPrelu.2517:15
-      attribute uniqueId 1
-        string-attr 357
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    fun $in_range_t 2
-    var t
-      type-con $ctype 0
-    var x
-      int
-    =
-    fun $_not 2
-    var t
-      type-con $ctype 0
-    fun $_not 2
-    var t
-      type-con $ctype 0
-    var x
-      int
-    var x
-      int
-axiom 0
-    forall 3 1 3
-      var t
-        type-con $ctype 0
-      var x
-        int
-      var y
-        int
-      pat 1
-        fun $_or 3
-        var t
-          type-con $ctype 0
-        var x
-          int
-        var y
-          int
-      attribute qid 1
-        string-attr VccPrelu.2518:15
-      attribute uniqueId 1
-        string-attr 358
-      attribute bvZ3Native 1
-        string-attr False
-    =
-    fun $_or 3
-    var t
-      type-con $ctype 0
-    var x
-      int
-    var y
-      int
-    fun $_or 3
-    var t
-      type-con $ctype 0
-    var y
-      int
-    var x
-      int
-axiom 0
-    forall 3 1 3
-      var t
-        type-con $ctype 0
-      var x
-        int
-      var y
-        int
-      pat 1
-        fun $_xor 3
-        var t
-          type-con $ctype 0
-        var x
-          int
-        var y
-          int
-      attribute qid 1
-        string-attr VccPrelu.2519:15
-      attribute uniqueId 1
-        string-attr 359
-      attribute bvZ3Native 1
-        string-attr False
-    =
-    fun $_xor 3
-    var t
-      type-con $ctype 0
-    var x
-      int
-    var y
-      int
-    fun $_xor 3
-    var t
-      type-con $ctype 0
-    var y
-      int
-    var x
-      int
-axiom 0
-    forall 3 1 3
-      var t
-        type-con $ctype 0
-      var x
-        int
-      var y
-        int
-      pat 1
-        fun $_and 3
-        var t
-          type-con $ctype 0
-        var x
-          int
-        var y
-          int
-      attribute qid 1
-        string-attr VccPrelu.2520:15
-      attribute uniqueId 1
-        string-attr 360
-      attribute bvZ3Native 1
-        string-attr False
-    =
-    fun $_and 3
-    var t
-      type-con $ctype 0
-    var x
-      int
-    var y
-      int
-    fun $_and 3
-    var t
-      type-con $ctype 0
-    var y
-      int
-    var x
-      int
-axiom 0
-    forall 2 1 4
-      var x
-        int
-      var y
-        int
-      pat 1
-        fun $_mul 2
-        var x
-          int
-        var y
-          int
-      attribute qid 1
-        string-attr VccPrelu.2524:28
-      attribute uniqueId 1
-        string-attr 361
-      attribute bvZ3Native 1
-        string-attr False
-      attribute weight 1
-        expr-attr
-          int-num 0
-    =
-    fun $_mul 2
-    var x
-      int
-    var y
-      int
-    *
-    var x
-      int
-    var y
-      int
-axiom 0
-    forall 2 1 3
-      var id
-        int
-      var length
-        int
-      pat 1
-        fun $get_string_literal 2
-        var id
-          int
-        var length
-          int
-      attribute qid 1
-        string-attr VccPrelu.2531:15
-      attribute uniqueId 1
-        string-attr 362
-      attribute bvZ3Native 1
-        string-attr False
-    fun $is 2
-    fun $get_string_literal 2
-    var id
-      int
-    var length
-      int
-    fun ^^u1 0
-axiom 0
-    forall 3 2 3
-      var id
-        int
-      var length
-        int
-      var S
-        type-con $state 0
-      pat 1
-        fun $typed 2
-        var S
-          type-con $state 0
-        fun $get_string_literal 2
-        var id
-          int
-        var length
-          int
-      pat 1
-        fun $is_array 4
-        var S
-          type-con $state 0
-        fun $get_string_literal 2
-        var id
-          int
-        var length
-          int
-        fun ^^u1 0
-        var length
-          int
-      attribute qid 1
-        string-attr VccPrelu.2532:15
-      attribute uniqueId 1
-        string-attr 363
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    fun $good_state 1
-    var S
-      type-con $state 0
-    and 2
-    fun $typed 2
-    var S
-      type-con $state 0
-    fun $get_string_literal 2
-    var id
-      int
-    var length
-      int
-    forall 1 2 3
-      var i
-        int
-      pat 1
-        fun $select.sm 2
-        fun $statusmap 1
-        var S
-          type-con $state 0
-        fun $idx 3
-        fun $get_string_literal 2
-        var id
-          int
-        var length
-          int
-        var i
-          int
-        fun ^^u1 0
-      pat 1
-        fun $select.tm 2
-        fun $typemap 1
-        var S
-          type-con $state 0
-        fun $idx 3
-        fun $get_string_literal 2
-        var id
-          int
-        var length
-          int
-        var i
-          int
-        fun ^^u1 0
-      attribute qid 1
-        string-attr VccPrelu.2043:13
-      attribute uniqueId 1
-        string-attr 236
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 2
-    <=
-    int-num 0
-    var i
-      int
-    <
-    var i
-      int
-    var length
-      int
-    and 3
-    fun $ts_is_array_elt 1
-    fun $select.tm 2
-    fun $typemap 1
-    var S
-      type-con $state 0
-    fun $idx 3
-    fun $get_string_literal 2
-    var id
-      int
-    var length
-      int
-    var i
-      int
-    fun ^^u1 0
-    fun $is 2
-    fun $idx 3
-    fun $get_string_literal 2
-    var id
-      int
-    var length
-      int
-    var i
-      int
-    fun ^^u1 0
-    fun ^^u1 0
-    fun $thread_local 2
-    var S
-      type-con $state 0
-    fun $idx 3
-    fun $get_string_literal 2
-    var id
-      int
-    var length
-      int
-    var i
-      int
-    fun ^^u1 0
-axiom 0
-    forall 2 1 3
-      var no
-        int
-      var t
-        type-con $ctype 0
-      pat 1
-        fun $get_fnptr 2
-        var no
-          int
-        var t
-          type-con $ctype 0
-      attribute qid 1
-        string-attr VccPrelu.2543:21
-      attribute uniqueId 1
-        string-attr 364
-      attribute bvZ3Native 1
-        string-attr False
-    =
-    fun $get_fnptr 2
-    var no
-      int
-    var t
-      type-con $ctype 0
-    fun $ptr 2
-    var t
-      type-con $ctype 0
-    fun $get_fnptr_ref 1
-    var no
-      int
-axiom 0
-    forall 1 0 3
-      var no
-        int
-      attribute qid 1
-        string-attr VccPrelu.2550:15
-      attribute uniqueId 1
-        string-attr 365
-      attribute bvZ3Native 1
-        string-attr False
-    =
-    fun $get_fnptr_inv 1
-    fun $get_fnptr_ref 1
-    var no
-      int
-    var no
-      int
-axiom 0
-    forall 3 2 3
-      var S
-        type-con $state 0
-      var no
-        int
-      var t
-        type-con $ctype 0
-      pat 1
-        fun $select.tm 2
-        fun $typemap 1
-        var S
-          type-con $state 0
-        fun $get_fnptr 2
-        var no
-          int
-        var t
-          type-con $ctype 0
-      pat 1
-        fun $select.sm 2
-        fun $statusmap 1
-        var S
-          type-con $state 0
-        fun $get_fnptr 2
-        var no
-          int
-        var t
-          type-con $ctype 0
-      attribute qid 1
-        string-attr VccPrelu.2553:15
-      attribute uniqueId 1
-        string-attr 366
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 2
-    fun $is_fnptr_type 1
-    var t
-      type-con $ctype 0
-    fun $good_state 1
-    var S
-      type-con $state 0
-    fun $mutable 2
-    var S
-      type-con $state 0
-    fun $get_fnptr 2
-    var no
-      int
-    var t
-      type-con $ctype 0
-axiom 0
-    forall 1 1 3
-      var t
-        type-con $ctype 0
-      pat 1
-        fun $is_math_type 1
-        var t
-          type-con $ctype 0
-      attribute qid 1
-        string-attr VccPrelu.2560:15
-      attribute uniqueId 1
-        string-attr 367
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    fun $is_math_type 1
-    var t
-      type-con $ctype 0
-    fun $is_primitive 1
-    var t
-      type-con $ctype 0
-axiom 0
-    forall 1 1 3
-      var t
-        type-con $ctype 0
-      pat 1
-        fun $is_fnptr_type 1
-        var t
-          type-con $ctype 0
-      attribute qid 1
-        string-attr VccPrelu.2561:15
-      attribute uniqueId 1
-        string-attr 368
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    fun $is_fnptr_type 1
-    var t
-      type-con $ctype 0
-    fun $is_primitive 1
-    var t
-      type-con $ctype 0
-axiom 0
-    forall 2 1 3
-      var S
-        type-con $state 0
-      var c
-        type-con $ptr 0
-      pat 2
-        fun $full_stop 1
-        var S
-          type-con $state 0
-        fun $valid_claim 2
-        var S
-          type-con $state 0
-        var c
-          type-con $ptr 0
-      attribute qid 1
-        string-attr VccPrelu.2572:15
-      attribute uniqueId 1
-        string-attr 369
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 2
-    fun $full_stop 1
-    var S
-      type-con $state 0
-    fun $closed 2
-    var S
-      type-con $state 0
-    var c
-      type-con $ptr 0
-    fun $valid_claim 2
-    var S
-      type-con $state 0
-    var c
-      type-con $ptr 0
-axiom 0
-    forall 2 1 3
-      var S
-        type-con $state 0
-      var c
-        type-con $ptr 0
-      pat 1
-        fun $valid_claim 2
-        var S
-          type-con $state 0
-        var c
-          type-con $ptr 0
-      attribute qid 1
-        string-attr VccPrelu.2575:15
-      attribute uniqueId 1
-        string-attr 370
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    fun $valid_claim 2
-    var S
-      type-con $state 0
-    var c
-      type-con $ptr 0
-    and 2
-    fun $closed 2
-    var S
-      type-con $state 0
-    var c
-      type-con $ptr 0
-    fun $invok_state 1
-    var S
-      type-con $state 0
-axiom 0
-    forall 2 1 3
-      var c1
-        type-con $ptr 0
-      var c2
-        type-con $ptr 0
-      pat 1
-        fun $claims_claim 2
-        var c1
-          type-con $ptr 0
-        var c2
-          type-con $ptr 0
-      attribute qid 1
-        string-attr VccPrelu.2600:15
-      attribute uniqueId 1
-        string-attr 373
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 3
-    fun $is 2
-    var c1
-      type-con $ptr 0
-    fun ^^claim 0
-    fun $is 2
-    var c2
-      type-con $ptr 0
-    fun ^^claim 0
-    forall 1 0 3
-      var S
-        type-con $state 0
-      attribute qid 1
-        string-attr VccPrelu.2602:11
-      attribute uniqueId 1
-        string-attr 372
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    fun $valid_claim 2
-    var S
-      type-con $state 0
-    var c1
-      type-con $ptr 0
-    fun $closed 2
-    var S
-      type-con $state 0
-    var c2
-      type-con $ptr 0
-    fun $claims_claim 2
-    var c1
-      type-con $ptr 0
-    var c2
-      type-con $ptr 0
-axiom 0
-    forall 3 1 3
-      var S
-        type-con $state 0
-      var c1
-        type-con $ptr 0
-      var c2
-        type-con $ptr 0
-      pat 2
-        fun $valid_claim 2
-        var S
-          type-con $state 0
-        var c1
-          type-con $ptr 0
-        fun $claims_claim 2
-        var c1
-          type-con $ptr 0
-        var c2
-          type-con $ptr 0
-      attribute qid 1
-        string-attr VccPrelu.2606:15
-      attribute uniqueId 1
-        string-attr 374
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 2
-    fun $valid_claim 2
-    var S
-      type-con $state 0
-    var c1
-      type-con $ptr 0
-    fun $claims_claim 2
-    var c1
-      type-con $ptr 0
-    var c2
-      type-con $ptr 0
-    fun $valid_claim 2
-    var S
-      type-con $state 0
-    var c2
-      type-con $ptr 0
-axiom 0
-    forall 3 1 3
-      var S
-        type-con $state 0
-      var c
-        type-con $ptr 0
-      var o
-        type-con $ptr 0
-      pat 2
-        fun $closed 2
-        var S
-          type-con $state 0
-        var c
-          type-con $ptr 0
-        fun $claims_obj 2
-        var c
-          type-con $ptr 0
-        var o
-          type-con $ptr 0
-      attribute qid 1
-        string-attr VccPrelu.2609:15
-      attribute uniqueId 1
-        string-attr 375
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    fun $good_state 1
-    var S
-      type-con $state 0
-    implies
-    and 2
-    fun $claims_obj 2
-    var c
-      type-con $ptr 0
-    var o
-      type-con $ptr 0
-    fun $closed 2
-    var S
-      type-con $state 0
-    var c
-      type-con $ptr 0
-    and 3
-    fun $instantiate_ptrset 1
-    fun $owns 2
-    var S
-      type-con $state 0
-    var o
-      type-con $ptr 0
-    fun $closed 2
-    var S
-      type-con $state 0
-    var o
-      type-con $ptr 0
-    >
-    fun $ref_cnt 2
-    var S
-      type-con $state 0
-    var o
-      type-con $ptr 0
-    int-num 0
-axiom 0
-    forall 3 1 3
-      var S
-        type-con $state 0
-      var c
-        type-con $ptr 0
-      var o
-        type-con $ptr 0
-      pat 2
-        fun $valid_claim 2
-        var S
-          type-con $state 0
-        var c
-          type-con $ptr 0
-        fun $claims_obj 2
-        var c
-          type-con $ptr 0
-        var o
-          type-con $ptr 0
-      attribute qid 1
-        string-attr VccPrelu.2614:15
-      attribute uniqueId 1
-        string-attr 376
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 2
-    fun $valid_claim 2
-    var S
-      type-con $state 0
-    var c
-      type-con $ptr 0
-    fun $claims_obj 2
-    var c
-      type-con $ptr 0
-    var o
-      type-con $ptr 0
-    fun $inv2 4
-    var S
-      type-con $state 0
-    var S
-      type-con $state 0
-    var o
-      type-con $ptr 0
-    fun $typ 1
-    var o
-      type-con $ptr 0
-axiom 0
-    forall 3 1 3
-      var S
-        type-con $state 0
-      var c
-        type-con $ptr 0
-      var r
-        int
-      pat 2
-        fun $valid_claim 2
-        var S
-          type-con $state 0
-        var c
-          type-con $ptr 0
-        fun $claims_obj 2
-        var c
-          type-con $ptr 0
-        fun $ptr 2
-        fun ^^claim 0
-        var r
-          int
-      attribute qid 1
-        string-attr VccPrelu.2618:15
-      attribute uniqueId 1
-        string-attr 377
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 2
-    fun $valid_claim 2
-    var S
-      type-con $state 0
-    var c
-      type-con $ptr 0
-    fun $claims_obj 2
-    var c
-      type-con $ptr 0
-    fun $ptr 2
-    fun ^^claim 0
-    var r
-      int
-    fun $valid_claim 2
-    var S
-      type-con $state 0
-    fun $ptr 2
-    fun ^^claim 0
-    var r
-      int
-axiom 0
-    forall 2 1 4
-      var S
-        type-con $state 0
-      var p
-        type-con $ptr 0
-      pat 1
-        fun $not_shared 2
-        var S
-          type-con $state 0
-        var p
-          type-con $ptr 0
-      attribute qid 1
-        string-attr VccPrelu.2623:34
-      attribute uniqueId 1
-        string-attr 378
-      attribute bvZ3Native 1
-        string-attr False
-      attribute weight 1
-        expr-attr
-          int-num 0
-    =
-    fun $not_shared 2
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    and 7
-    fun $closed 2
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    =
-    fun $owner 2
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    fun $me 0
-    fun $is 2
-    var p
-      type-con $ptr 0
-    fun $typ 1
-    var p
-      type-con $ptr 0
-    fun $typed 2
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    not
-    =
-    fun $kind_of 1
-    fun $typ 1
-    var p
-      type-con $ptr 0
-    fun $kind_primitive 0
-    fun $is_non_primitive 1
-    fun $typ 1
-    var p
-      type-con $ptr 0
-    or 2
-    not
-    fun $is_claimable 1
-    fun $typ 1
-    var p
-      type-con $ptr 0
-    =
-    fun $ref_cnt 2
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    int-num 0
-axiom 0
-    forall 2 1 4
-      var s
-        type-con $state 0
-      var p
-        type-con $ptr 0
-      pat 1
-        fun $claimed_closed 2
-        var s
-          type-con $state 0
-        var p
-          type-con $ptr 0
-      attribute qid 1
-        string-attr VccPrelu.2626:38
-      attribute uniqueId 1
-        string-attr 379
-      attribute bvZ3Native 1
-        string-attr False
-      attribute weight 1
-        expr-attr
-          int-num 0
-    =
-    fun $claimed_closed 2
-    var s
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    fun $closed 2
-    var s
-      type-con $state 0
-    var p
-      type-con $ptr 0
-axiom 0
-    forall 2 1 3
-      var S
-        type-con $state 0
-      var p
-        type-con $ptr 0
-      pat 2
-        fun $invok_state 1
-        var S
-          type-con $state 0
-        fun $claimed_closed 2
-        var S
-          type-con $state 0
-        var p
-          type-con $ptr 0
-      attribute qid 1
-        string-attr VccPrelu.2629:15
-      attribute uniqueId 1
-        string-attr 380
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 2
-    fun $invok_state 1
-    var S
-      type-con $state 0
-    fun $claimed_closed 2
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    fun $inv2 4
-    var S
-      type-con $state 0
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    fun $typ 1
-    var p
-      type-con $ptr 0
-axiom 0
-    =
-    fun $no_claim 0
-    fun $ptr 2
-    fun ^^claim 0
-    int-num 0
-axiom 0
-    forall 2 1 4
-      var S
-        type-con $state 0
-      var p
-        type-con $ptr 0
-      pat 1
-        fun $ref_cnt 2
-        var S
-          type-con $state 0
-        var p
-          type-con $ptr 0
-      attribute qid 1
-        string-attr VccPrelu.2685:31
-      attribute uniqueId 1
-        string-attr 388
-      attribute bvZ3Native 1
-        string-attr False
-      attribute weight 1
-        expr-attr
-          int-num 0
-    =
-    fun $ref_cnt 2
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    fun $st_ref_cnt 1
-    fun $select.sm 2
-    fun $statusmap 1
-    var S
-      type-con $state 0
-    var p
-      type-con $ptr 0
-axiom 0
-    fun $is_claimable 1
-    fun ^^claim 0
-axiom 0
-    forall 1 0 3
-      var p
-        type-con $ptrset 0
-      attribute qid 1
-        string-attr VccPrelu.2729:15
-      attribute uniqueId 1
-        string-attr 390
-      attribute bvZ3Native 1
-        string-attr False
-    =
-    fun $int_to_ptrset 1
-    fun $ptrset_to_int 1
-    var p
-      type-con $ptrset 0
-    var p
-      type-con $ptrset 0
-axiom 0
-    forall 1 0 3
-      var p
-        type-con $version 0
-      attribute qid 1
-        string-attr VccPrelu.2733:15
-      attribute uniqueId 1
-        string-attr 391
-      attribute bvZ3Native 1
-        string-attr False
-    =
-    fun $int_to_version 1
-    fun $version_to_int 1
-    var p
-      type-con $version 0
-    var p
-      type-con $version 0
-axiom 0
-    forall 1 0 3
-      var p
-        type-con $vol_version 0
-      attribute qid 1
-        string-attr VccPrelu.2737:15
-      attribute uniqueId 1
-        string-attr 392
-      attribute bvZ3Native 1
-        string-attr False
-    =
-    fun $int_to_vol_version 1
-    fun $vol_version_to_int 1
-    var p
-      type-con $vol_version 0
-    var p
-      type-con $vol_version 0
-axiom 0
-    forall 1 0 3
-      var p
-        type-con $ptr 0
-      attribute qid 1
-        string-attr VccPrelu.2741:15
-      attribute uniqueId 1
-        string-attr 393
-      attribute bvZ3Native 1
-        string-attr False
-    =
-    fun $int_to_ptr 1
-    fun $ptr_to_int 1
-    var p
-      type-con $ptr 0
-    var p
-      type-con $ptr 0
-axiom 0
-    forall 3 1 3
-      var S1
-        type-con $state 0
-      var S2
-        type-con $state 0
-      var W
-        type-con $ptrset 0
-      pat 1
-        fun $updated_only_values 3
-        var S1
-          type-con $state 0
-        var S2
-          type-con $state 0
-        var W
-          type-con $ptrset 0
-      attribute qid 1
-        string-attr VccPrelu.2752:15
-      attribute uniqueId 1
-        string-attr 395
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    forall 1 1 3
-      var p
-        type-con $ptr 0
-      pat 1
-        fun $dont_instantiate 1
-        var p
-          type-con $ptr 0
-      attribute qid 1
-        string-attr VccPrelu.2754:11
-      attribute uniqueId 1
-        string-attr 394
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    or 2
-    fun $is_primitive 1
-    fun $typ 1
-    var p
-      type-con $ptr 0
-    fun $is_non_primitive 1
-    fun $typ 1
-    var p
-      type-con $ptr 0
-    implies
-    and 2
-    fun $typed 2
-    var S1
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    not
-    or 2
-    not
-    =
-    fun $owner 2
-    var S1
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    fun $me 0
-    and 2
-    =
-    fun $kind_of 1
-    fun $typ 1
-    var p
-      type-con $ptr 0
-    fun $kind_primitive 0
-    fun $closed 2
-    var S1
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    or 2
-    =
-    fun $select.mem 2
-    fun $memory 1
-    var S1
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    fun $select.mem 2
-    fun $memory 1
-    var S2
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    fun $set_in 2
-    var p
-      type-con $ptr 0
-    var W
-      type-con $ptrset 0
-    fun $updated_only_values 3
-    var S1
-      type-con $state 0
-    var S2
-      type-con $state 0
-    var W
-      type-con $ptrset 0
-axiom 0
-    forall 3 1 3
-      var S1
-        type-con $state 0
-      var S2
-        type-con $state 0
-      var W
-        type-con $ptrset 0
-      pat 1
-        fun $updated_only_domains 3
-        var S1
-          type-con $state 0
-        var S2
-          type-con $state 0
-        var W
-          type-con $ptrset 0
-      attribute qid 1
-        string-attr VccPrelu.2759:15
-      attribute uniqueId 1
-        string-attr 397
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    forall 1 1 3
-      var p
-        type-con $ptr 0
-      pat 1
-        fun $dont_instantiate 1
-        var p
-          type-con $ptr 0
-      attribute qid 1
-        string-attr VccPrelu.2761:11
-      attribute uniqueId 1
-        string-attr 396
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 2
-    fun $set_in 2
-    var p
-      type-con $ptr 0
-    var W
-      type-con $ptrset 0
-    not
-    =
-    fun $kind_of 1
-    fun $typ 1
-    var p
-      type-con $ptr 0
-    fun $kind_primitive 0
-    or 2
-    =
-    fun $select.mem 2
-    fun $memory 1
-    var S1
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    fun $select.mem 2
-    fun $memory 1
-    var S2
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    fun $domain_updated_at 4
-    var S1
-      type-con $state 0
-    var S2
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    var W
-      type-con $ptrset 0
-    fun $updated_only_domains 3
-    var S1
-      type-con $state 0
-    var S2
-      type-con $state 0
-    var W
-      type-con $ptrset 0
-axiom 0
-    forall 4 1 3
-      var S1
-        type-con $state 0
-      var S2
-        type-con $state 0
-      var p
-        type-con $ptr 0
-      var W
-        type-con $ptrset 0
-      pat 1
-        fun $domain_updated_at 4
-        var S1
-          type-con $state 0
-        var S2
-          type-con $state 0
-        var p
-          type-con $ptr 0
-        var W
-          type-con $ptrset 0
-      attribute qid 1
-        string-attr VccPrelu.2777:29
-      attribute uniqueId 1
-        string-attr 399
-      attribute bvZ3Native 1
-        string-attr False
-    =
-    fun $domain_updated_at 4
-    var S1
-      type-con $state 0
-    var S2
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    var W
-      type-con $ptrset 0
-    and 2
-    forall 1 1 3
-      var q
-        type-con $ptr 0
-      pat 1
-        fun $fetch_from_domain 2
-        fun $read_version 2
-        var S2
-          type-con $state 0
-        var p
-          type-con $ptr 0
-        var q
-          type-con $ptr 0
-      attribute qid 1
-        string-attr VccPrelu.2778:13
-      attribute uniqueId 1
-        string-attr 398
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 2
-    =
-    fun $kind_of 1
-    fun $typ 1
-    var q
-      type-con $ptr 0
-    fun $kind_primitive 0
-    not
-    fun $set_in 2
-    var q
-      type-con $ptr 0
-    var W
-      type-con $ptrset 0
-    =
-    fun $fetch_from_domain 2
-    fun $read_version 2
-    var S1
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    var q
-      type-con $ptr 0
-    fun $fetch_from_domain 2
-    fun $read_version 2
-    var S2
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    var q
-      type-con $ptr 0
-    =
-    fun $domain 2
-    var S1
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    fun $domain 2
-    var S2
-      type-con $state 0
-    var p
-      type-con $ptr 0
-axiom 0
-    =
-    fun #distTp1 0
-    fun $ptr_to 1
-    fun ^^u1 0
-axiom 0
-    fun $type_code_is 2
-    int-num 1
-    fun ^^u4 0
-axiom 0
-    fun $file_name_is 2
-    int-num 1
-    fun #file^Z?3A?5CC?5Cmax.c 0
-var-decl $s 0
-    type-con $state 0
-vc maximum 1
-    implies
-    label pos 0 0
-    true
-    implies
-    and 2
-    <=
-    int-num 0
-    var L#max
-      int
-    <=
-    var L#max
-      int
-    fun $max.u1 0
-    implies
-    and 2
-    <=
-    int-num 0
-    var L#p
-      int
-    <=
-    var L#p
-      int
-    fun $max.u4 0
-    implies
-    and 2
-    <=
-    int-num 0
-    var SL#witness
-      int
-    <=
-    var SL#witness
-      int
-    fun $max.u4 0
-    implies
-    <
-    var P#len
-      int
-    int-num 1099511627776
-    implies
-    <
-    int-num 0
-    var P#len
-      int
-    implies
-    and 6
-    fun $closed 2
-    var $s
-      type-con $state 0
-    fun $ptr 2
-    fun $array 2
-    fun ^^u1 0
-    var P#len
-      int
-    fun $ref 1
-    fun $ptr 2
-    fun ^^u1 0
-    var P#arr
-      int
-    =
-    fun $owner 2
-    var $s
-      type-con $state 0
-    fun $ptr 2
-    fun $array 2
-    fun ^^u1 0
-    var P#len
-      int
-    fun $ref 1
-    fun $ptr 2
-    fun ^^u1 0
-    var P#arr
-      int
-    fun $me 0
-    fun $is 2
-    fun $ptr 2
-    fun $array 2
-    fun ^^u1 0
-    var P#len
-      int
-    fun $ref 1
-    fun $ptr 2
-    fun ^^u1 0
-    var P#arr
-      int
-    fun $array 2
-    fun ^^u1 0
-    var P#len
-      int
-    fun $typed 2
-    var $s
-      type-con $state 0
-    fun $ptr 2
-    fun $array 2
-    fun ^^u1 0
-    var P#len
-      int
-    fun $ref 1
-    fun $ptr 2
-    fun ^^u1 0
-    var P#arr
-      int
-    not
-    =
-    fun $kind_of 1
-    fun $array 2
-    fun ^^u1 0
-    var P#len
-      int
-    fun $kind_primitive 0
-    fun $is_non_primitive 1
-    fun $array 2
-    fun ^^u1 0
-    var P#len
-      int
-    implies
-    true
-    implies
-    fun $function_entry 1
-    var $s
-      type-con $state 0
-    implies
-    and 2
-    fun $good_state_ext 2
-    fun #tok$1^6.1 0
-    var $s
-      type-con $state 0
-    fun $full_stop 1
-    var $s
-      type-con $state 0
-    implies
-    forall 1 1 3
-      var f
-        type-con $pure_function 0
-      pat 1
-        fun $frame_level 1
-        var f
-          type-con $pure_function 0
-      attribute qid 1
-        string-attr VccPrelu.2703:13
-      attribute uniqueId 1
-        string-attr 389
-      attribute bvZ3Native 1
-        string-attr False
-    <
-    fun $frame_level 1
-    var f
-      type-con $pure_function 0
-    fun $current_frame_level 0
-    implies
-    and 2
-    fun $local_value_is 5
-    var $s
-      type-con $state 0
-    fun #tok$1^6.1 0
-    fun #loc.arr 0
-    fun $ptr_to_int 1
-    fun $ptr 2
-    fun ^^u1 0
-    var P#arr
-      int
-    fun $ptr_to 1
-    fun ^^u1 0
-    fun $local_value_is_ptr 5
-    var $s
-      type-con $state 0
-    fun #tok$1^6.1 0
-    fun #loc.arr 0
-    fun $ptr 2
-    fun ^^u1 0
-    var P#arr
-      int
-    fun $ptr_to 1
-    fun ^^u1 0
-    implies
-    fun $local_value_is 5
-    var $s
-      type-con $state 0
-    fun #tok$1^6.1 0
-    fun #loc.len 0
-    var P#len
-      int
-    fun ^^u4 0
-    implies
-    =
-    var #wrTime$1^6.1
-      int
-    fun $current_timestamp 1
-    var $s
-      type-con $state 0
-    implies
-    forall 1 1 3
-      var #p
-        type-con $ptr 0
-      pat 1
-        fun $in_writes_at 2
-        var #wrTime$1^6.1
-          int
-        var #p
-          type-con $ptr 0
-      attribute qid 1
-        string-attr nofile.0:0
-      attribute uniqueId 1
-        string-attr 2
-      attribute bvZ3Native 1
-        string-attr False
-    =
-    fun $in_writes_at 2
-    var #wrTime$1^6.1
-      int
-    var #p
-      type-con $ptr 0
-    false
-    implies
-    and 2
-    <=
-    int-num 0
-    var P#len
-      int
-    <=
-    var P#len
-      int
-    fun $max.u4 0
-    and 2
-    label neg 7 27
-    fun $in_domain_lab 4
-    var $s
-      type-con $state 0
-    fun $ptr 2
-    fun $array 2
-    fun ^^u1 0
-    var P#len
-      int
-    fun $ref 1
-    fun $ptr 2
-    fun ^^u1 0
-    var P#arr
-      int
-    fun $ptr 2
-    fun $array 2
-    fun ^^u1 0
-    var P#len
-      int
-    fun $ref 1
-    fun $ptr 2
-    fun ^^u1 0
-    var P#arr
-      int
-    fun l#public 0
-    implies
-    fun $in_domain_lab 4
-    var $s
-      type-con $state 0
-    fun $ptr 2
-    fun $array 2
-    fun ^^u1 0
-    var P#len
-      int
-    fun $ref 1
-    fun $ptr 2
-    fun ^^u1 0
-    var P#arr
-      int
-    fun $ptr 2
-    fun $array 2
-    fun ^^u1 0
-    var P#len
-      int
-    fun $ref 1
-    fun $ptr 2
-    fun ^^u1 0
-    var P#arr
-      int
-    fun l#public 0
-    and 2
-    label neg 12 14
-    and 2
-    fun $is 2
-    fun $idx 3
-    fun $ptr 2
-    fun ^^u1 0
-    var P#arr
-      int
-    int-num 0
-    fun ^^u1 0
-    fun ^^u1 0
-    fun $typed 2
-    var $s
-      type-con $state 0
-    fun $idx 3
-    fun $ptr 2
-    fun ^^u1 0
-    var P#arr
-      int
-    int-num 0
-    fun ^^u1 0
-    implies
-    and 2
-    fun $is 2
-    fun $idx 3
-    fun $ptr 2
-    fun ^^u1 0
-    var P#arr
-      int
-    int-num 0
-    fun ^^u1 0
-    fun ^^u1 0
-    fun $typed 2
-    var $s
-      type-con $state 0
-    fun $idx 3
-    fun $ptr 2
-    fun ^^u1 0
-    var P#arr
-      int
-    int-num 0
-    fun ^^u1 0
-    and 2
-    label neg 12 14
-    and 2
-    fun $is 2
-    fun $idx 3
-    fun $ptr 2
-    fun ^^u1 0
-    var P#arr
-      int
-    int-num 0
-    fun ^^u1 0
-    fun ^^u1 0
-    fun $thread_local 2
-    var $s
-      type-con $state 0
-    fun $idx 3
-    fun $ptr 2
-    fun ^^u1 0
-    var P#arr
-      int
-    int-num 0
-    fun ^^u1 0
-    implies
-    and 2
-    fun $is 2
-    fun $idx 3
-    fun $ptr 2
-    fun ^^u1 0
-    var P#arr
-      int
-    int-num 0
-    fun ^^u1 0
-    fun ^^u1 0
-    fun $thread_local 2
-    var $s
-      type-con $state 0
-    fun $idx 3
-    fun $ptr 2
-    fun ^^u1 0
-    var P#arr
-      int
-    int-num 0
-    fun ^^u1 0
-    implies
-    =
-    var L#max@0
-      int
-    fun $read_u1 2
-    var $s
-      type-con $state 0
-    fun $idx 3
-    fun $ptr 2
-    fun ^^u1 0
-    var P#arr
-      int
-    int-num 0
-    fun ^^u1 0
-    implies
-    fun $local_value_is 5
-    var $s
-      type-con $state 0
-    fun #tok$1^12.3 0
-    fun #loc.max 0
-    var L#max@0
-      int
-    fun ^^u1 0
-    implies
-    fun $local_value_is 5
-    var $s
-      type-con $state 0
-    fun #tok$1^14.3 0
-    fun #loc.witness 0
-    int-num 0
-    fun ^^u4 0
-    implies
-    fun $local_value_is 5
-    var $s
-      type-con $state 0
-    fun #tok$1^16.8 0
-    fun #loc.p 0
-    int-num 1
-    fun ^^u4 0
-    implies
-    and 4
-    <=
-    int-num 1
-    int-num 1
-    <=
-    int-num 1
-    int-num 1
-    <=
-    int-num 0
-    int-num 0
-    <=
-    int-num 0
-    int-num 0
-    and 2
-    label neg 17 17
-    <=
-    int-num 1
-    var P#len
-      int
-    implies
-    <=
-    int-num 1
-    var P#len
-      int
-    and 2
-    label neg 18 17
-    forall 1 0 3
-      var Q#i$1^18.17#tc1
-        int
-      attribute qid 1
-        string-attr maxc.18:17
-      attribute uniqueId 1
-        string-attr 3
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 2
-    <=
-    int-num 0
-    var Q#i$1^18.17#tc1
-      int
-    <=
-    var Q#i$1^18.17#tc1
-      int
-    fun $max.u4 0
-    implies
-    <
-    var Q#i$1^18.17#tc1
-      int
-    int-num 1
-    <=
-    fun $read_u1 2
-    var $s
-      type-con $state 0
-    fun $idx 3
-    fun $ptr 2
-    fun ^^u1 0
-    var P#arr
-      int
-    var Q#i$1^18.17#tc1
-      int
-    fun ^^u1 0
-    var L#max@0
-      int
-    implies
-    forall 1 0 3
-      var Q#i$1^18.17#tc1
-        int
-      attribute qid 1
-        string-attr maxc.18:17
-      attribute uniqueId 1
-        string-attr 3
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 2
-    <=
-    int-num 0
-    var Q#i$1^18.17#tc1
-      int
-    <=
-    var Q#i$1^18.17#tc1
-      int
-    fun $max.u4 0
-    implies
-    <
-    var Q#i$1^18.17#tc1
-      int
-    int-num 1
-    <=
-    fun $read_u1 2
-    var $s
-      type-con $state 0
-    fun $idx 3
-    fun $ptr 2
-    fun ^^u1 0
-    var P#arr
-      int
-    var Q#i$1^18.17#tc1
-      int
-    fun ^^u1 0
-    var L#max@0
-      int
-    and 2
-    label neg 19 17
-    and 2
-    <
-    int-num 0
-    var P#len
-      int
-    =
-    fun $read_u1 2
-    var $s
-      type-con $state 0
-    fun $idx 3
-    fun $ptr 2
-    fun ^^u1 0
-    var P#arr
-      int
-    int-num 0
-    fun ^^u1 0
-    var L#max@0
-      int
-    implies
-    and 2
-    <
-    int-num 0
-    var P#len
-      int
-    =
-    fun $read_u1 2
-    var $s
-      type-con $state 0
-    fun $idx 3
-    fun $ptr 2
-    fun ^^u1 0
-    var P#arr
-      int
-    int-num 0
-    fun ^^u1 0
-    var L#max@0
-      int
-    implies
-    label pos 0 0
-    true
-    implies
-    and 2
-    <=
-    int-num 0
-    var L#max@1
-      int
-    <=
-    var L#max@1
-      int
-    fun $max.u1 0
-    implies
-    and 2
-    <=
-    int-num 0
-    var SL#witness@0
-      int
-    <=
-    var SL#witness@0
-      int
-    fun $max.u4 0
-    implies
-    and 2
-    <=
-    int-num 0
-    var L#p@0
-      int
-    <=
-    var L#p@0
-      int
-    fun $max.u4 0
-    implies
-    and 2
-    <=
-    int-num 1
-    var L#p@0
-      int
-    <=
-    int-num 0
-    var SL#witness@0
-      int
-    implies
-    <=
-    var L#p@0
-      int
-    var P#len
-      int
-    implies
-    forall 1 0 3
-      var Q#i$1^18.17#tc1
-        int
-      attribute qid 1
-        string-attr maxc.18:17
-      attribute uniqueId 1
-        string-attr 3
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 2
-    <=
-    int-num 0
-    var Q#i$1^18.17#tc1
-      int
-    <=
-    var Q#i$1^18.17#tc1
-      int
-    fun $max.u4 0
-    implies
-    <
-    var Q#i$1^18.17#tc1
-      int
-    var L#p@0
-      int
-    <=
-    fun $read_u1 2
-    var $s
-      type-con $state 0
-    fun $idx 3
-    fun $ptr 2
-    fun ^^u1 0
-    var P#arr
-      int
-    var Q#i$1^18.17#tc1
-      int
-    fun ^^u1 0
-    var L#max@1
-      int
-    implies
-    and 2
-    <
-    var SL#witness@0
-      int
-    var P#len
-      int
-    =
-    fun $read_u1 2
-    var $s
-      type-con $state 0
-    fun $idx 3
-    fun $ptr 2
-    fun ^^u1 0
-    var P#arr
-      int
-    var SL#witness@0
-      int
-    fun ^^u1 0
-    var L#max@1
-      int
-    implies
-    and 2
-    <=
-    int-num 1
-    var L#p@0
-      int
-    <=
-    int-num 0
-    var SL#witness@0
-      int
-    and 2
-    implies
-    label pos 0 0
-    true
-    implies
-    and 2
-    <=
-    int-num 1
-    var L#p@0
-      int
-    <=
-    int-num 0
-    var SL#witness@0
-      int
-    implies
-    not
-    true
-    implies
-    and 2
-    <=
-    int-num 1
-    var L#p@0
-      int
-    <=
-    int-num 0
-    var SL#witness@0
-      int
-    implies
-    label pos 0 0
-    true
-    implies
-    and 2
-    <=
-    int-num 1
-    var L#p@0
-      int
-    <=
-    int-num 0
-    var SL#witness@0
-      int
-    implies
-    and 2
-    fun $good_state_ext 2
-    fun #tok$1^16.3 0
-    var $s
-      type-con $state 0
-    fun $full_stop 1
-    var $s
-      type-con $state 0
-    implies
-    and 2
-    <=
-    int-num 1
-    var L#p@0
-      int
-    <=
-    int-num 0
-    var SL#witness@0
-      int
-    implies
-    label pos 0 0
-    true
-    implies
-    and 2
-    <=
-    int-num 1
-    var L#p@0
-      int
-    <=
-    int-num 0
-    var SL#witness@0
-      int
-    and 2
-    label neg 27 3
-    fun $position_marker 0
-    implies
-    fun $position_marker 0
-    implies
-    and 2
-    <=
-    int-num 1
-    var L#p@0
-      int
-    <=
-    int-num 0
-    var SL#witness@0
-      int
-    implies
-    label pos 0 0
-    true
-    implies
-    and 2
-    <=
-    int-num 1
-    var L#p@0
-      int
-    <=
-    int-num 0
-    var SL#witness@0
-      int
-    implies
-    and 2
-    <=
-    int-num 1
-    var L#p@0
-      int
-    <=
-    int-num 0
-    var SL#witness@0
-      int
-    implies
-    label pos 0 0
-    true
-    implies
-    =
-    var L#max@4
-      int
-    var L#max@1
-      int
-    implies
-    =
-    var L#p@2
-      int
-    var L#p@0
-      int
-    implies
-    =
-    var SL#witness@2
-      int
-    var SL#witness@0
-      int
-    implies
-    =
-    var $result@0
-      int
-    var L#max@1
-      int
-    implies
-    label pos 0 0
-    true
-    and 2
-    label neg 9 14
-    forall 1 0 3
-      var Q#i$1^9.14#tc1
-        int
-      attribute qid 1
-        string-attr maxc.9:14
-      attribute uniqueId 1
-        string-attr 1
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 2
-    <=
-    int-num 0
-    var Q#i$1^9.14#tc1
-      int
-    <=
-    var Q#i$1^9.14#tc1
-      int
-    fun $max.u4 0
-    implies
-    <
-    var Q#i$1^9.14#tc1
-      int
-    var P#len
-      int
-    <=
-    fun $read_u1 2
-    var $s
-      type-con $state 0
-    fun $idx 3
-    fun $ptr 2
-    fun ^^u1 0
-    var P#arr
-      int
-    var Q#i$1^9.14#tc1
-      int
-    fun ^^u1 0
-    var $result@0
-      int
-    implies
-    forall 1 0 3
-      var Q#i$1^9.14#tc1
-        int
-      attribute qid 1
-        string-attr maxc.9:14
-      attribute uniqueId 1
-        string-attr 1
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 2
-    <=
-    int-num 0
-    var Q#i$1^9.14#tc1
-      int
-    <=
-    var Q#i$1^9.14#tc1
-      int
-    fun $max.u4 0
-    implies
-    <
-    var Q#i$1^9.14#tc1
-      int
-    var P#len
-      int
-    <=
-    fun $read_u1 2
-    var $s
-      type-con $state 0
-    fun $idx 3
-    fun $ptr 2
-    fun ^^u1 0
-    var P#arr
-      int
-    var Q#i$1^9.14#tc1
-      int
-    fun ^^u1 0
-    var $result@0
-      int
-    and 2
-    label neg 10 14
-    exists 1 0 3
-      var Q#i$1^10.14#tc1
-        int
-      attribute qid 1
-        string-attr maxc.10:14
-      attribute uniqueId 1
-        string-attr 0
-      attribute bvZ3Native 1
-        string-attr False
-    and 4
-    <=
-    int-num 0
-    var Q#i$1^10.14#tc1
-      int
-    <=
-    var Q#i$1^10.14#tc1
-      int
-    fun $max.u4 0
-    <
-    var Q#i$1^10.14#tc1
-      int
-    var P#len
-      int
-    =
-    fun $read_u1 2
-    var $s
-      type-con $state 0
-    fun $idx 3
-    fun $ptr 2
-    fun ^^u1 0
-    var P#arr
-      int
-    var Q#i$1^10.14#tc1
-      int
-    fun ^^u1 0
-    var $result@0
-      int
-    implies
-    exists 1 0 3
-      var Q#i$1^10.14#tc1
-        int
-      attribute qid 1
-        string-attr maxc.10:14
-      attribute uniqueId 1
-        string-attr 0
-      attribute bvZ3Native 1
-        string-attr False
-    and 4
-    <=
-    int-num 0
-    var Q#i$1^10.14#tc1
-      int
-    <=
-    var Q#i$1^10.14#tc1
-      int
-    fun $max.u4 0
-    <
-    var Q#i$1^10.14#tc1
-      int
-    var P#len
-      int
-    =
-    fun $read_u1 2
-    var $s
-      type-con $state 0
-    fun $idx 3
-    fun $ptr 2
-    fun ^^u1 0
-    var P#arr
-      int
-    var Q#i$1^10.14#tc1
-      int
-    fun ^^u1 0
-    var $result@0
-      int
-    true
-    implies
-    label pos 0 0
-    true
-    implies
-    and 2
-    <=
-    int-num 1
-    var L#p@0
-      int
-    <=
-    int-num 0
-    var SL#witness@0
-      int
-    implies
-    true
-    implies
-    and 2
-    <=
-    int-num 1
-    var L#p@0
-      int
-    <=
-    int-num 0
-    var SL#witness@0
-      int
-    implies
-    label pos 0 0
-    true
-    implies
-    and 2
-    <=
-    int-num 1
-    var L#p@0
-      int
-    <=
-    int-num 0
-    var SL#witness@0
-      int
-    implies
-    and 7
-    forall 1 1 3
-      var p
-        type-con $ptr 0
-      pat 1
-        fun $select.sm 2
-        fun $statusmap 1
-        var $s
-          type-con $state 0
-        var p
-          type-con $ptr 0
-      attribute qid 1
-        string-attr VccPrelu.1970:13
-      attribute uniqueId 1
-        string-attr 220
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    not
-    =
-    fun $kind_of 1
-    fun $typ 1
-    fun $owner 2
-    var $s
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    fun $kind_thread 0
-    not
-    =
-    fun $kind_of 1
-    fun $typ 1
-    fun $owner 2
-    var $s
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    fun $kind_thread 0
-    forall 1 1 3
-      var p
-        type-con $ptr 0
-      pat 1
-        fun $select.mem 2
-        fun $memory 1
-        var $s
-          type-con $state 0
-        var p
-          type-con $ptr 0
-      attribute qid 1
-        string-attr VccPrelu.1972:13
-      attribute uniqueId 1
-        string-attr 221
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    fun $thread_local 2
-    var $s
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    and 2
-    =
-    fun $select.mem 2
-    fun $memory 1
-    var $s
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    fun $select.mem 2
-    fun $memory 1
-    var $s
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    fun $thread_local 2
-    var $s
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    forall 1 1 3
-      var p
-        type-con $ptr 0
-      pat 1
-        fun $select.sm 2
-        fun $statusmap 1
-        var $s
-          type-con $state 0
-        var p
-          type-con $ptr 0
-      attribute qid 1
-        string-attr VccPrelu.1974:13
-      attribute uniqueId 1
-        string-attr 222
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    fun $thread_local 2
-    var $s
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    and 2
-    =
-    fun $select.sm 2
-    fun $statusmap 1
-    var $s
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    fun $select.sm 2
-    fun $statusmap 1
-    var $s
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    fun $thread_local 2
-    var $s
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    forall 1 1 3
-      var p
-        type-con $ptr 0
-      pat 1
-        fun $select.tm 2
-        fun $typemap 1
-        var $s
-          type-con $state 0
-        var p
-          type-con $ptr 0
-      attribute qid 1
-        string-attr VccPrelu.1976:13
-      attribute uniqueId 1
-        string-attr 223
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    fun $thread_local 2
-    var $s
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    and 2
-    =
-    fun $select.tm 2
-    fun $typemap 1
-    var $s
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    fun $select.tm 2
-    fun $typemap 1
-    var $s
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    fun $thread_local 2
-    var $s
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    <=
-    fun $current_timestamp 1
-    var $s
-      type-con $state 0
-    fun $current_timestamp 1
-    var $s
-      type-con $state 0
-    forall 1 1 4
-      var p
-        type-con $ptr 0
-      pat 1
-        fun $timestamp 2
-        var $s
-          type-con $state 0
-        var p
-          type-con $ptr 0
-      attribute qid 1
-        string-attr VccPrelu.1274:13
-      attribute uniqueId 1
-        string-attr 139
-      attribute bvZ3Native 1
-        string-attr False
-      attribute weight 1
-        expr-attr
-          int-num 0
-    <=
-    fun $timestamp 2
-    var $s
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    fun $timestamp 2
-    var $s
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    fun $call_transition 2
-    var $s
-      type-con $state 0
-    var $s
-      type-con $state 0
-    implies
-    and 3
-    <=
-    fun $current_timestamp 1
-    var $s
-      type-con $state 0
-    fun $current_timestamp 1
-    var $s
-      type-con $state 0
-    forall 1 1 4
-      var p
-        type-con $ptr 0
-      pat 1
-        fun $timestamp 2
-        var $s
-          type-con $state 0
-        var p
-          type-con $ptr 0
-      attribute qid 1
-        string-attr VccPrelu.1274:13
-      attribute uniqueId 1
-        string-attr 139
-      attribute bvZ3Native 1
-        string-attr False
-      attribute weight 1
-        expr-attr
-          int-num 0
-    <=
-    fun $timestamp 2
-    var $s
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    fun $timestamp 2
-    var $s
-      type-con $state 0
-    var p
-      type-con $ptr 0
-    fun $call_transition 2
-    var $s
-      type-con $state 0
-    var $s
-      type-con $state 0
-    implies
-    and 2
-    fun $good_state_ext 2
-    fun #tok$1^16.3 0
-    var $s
-      type-con $state 0
-    fun $full_stop 1
-    var $s
-      type-con $state 0
-    implies
-    fun $local_value_is 5
-    var $s
-      type-con $state 0
-    fun #tok$1^16.3 0
-    fun #loc.p 0
-    var L#p@0
-      int
-    fun ^^u4 0
-    implies
-    fun $local_value_is 5
-    var $s
-      type-con $state 0
-    fun #tok$1^16.3 0
-    fun #loc.witness 0
-    var SL#witness@0
-      int
-    fun ^^u4 0
-    implies
-    fun $local_value_is 5
-    var $s
-      type-con $state 0
-    fun #tok$1^16.3 0
-    fun #loc.max 0
-    var L#max@1
-      int
-    fun ^^u1 0
-    implies
-    fun $local_value_is 5
-    var $s
-      type-con $state 0
-    fun #tok$1^16.3 0
-    fun #loc.len 0
-    var P#len
-      int
-    fun ^^u4 0
-    implies
-    and 2
-    fun $local_value_is 5
-    var $s
-      type-con $state 0
-    fun #tok$1^16.3 0
-    fun #loc.arr 0
-    fun $ptr_to_int 1
-    fun $ptr 2
-    fun ^^u1 0
-    var P#arr
-      int
-    fun $ptr_to 1
-    fun ^^u1 0
-    fun $local_value_is_ptr 5
-    var $s
-      type-con $state 0
-    fun #tok$1^16.3 0
-    fun #loc.arr 0
-    fun $ptr 2
-    fun ^^u1 0
-    var P#arr
-      int
-    fun $ptr_to 1
-    fun ^^u1 0
-    implies
-    and 2
-    =
-    fun $typemap 1
-    var $s
-      type-con $state 0
-    fun $typemap 1
-    var $s
-      type-con $state 0
-    =
-    fun $statusmap 1
-    var $s
-      type-con $state 0
-    fun $statusmap 1
-    var $s
-      type-con $state 0
-    implies
-    and 2
-    <=
-    int-num 1
-    var L#p@0
-      int
-    <=
-    int-num 0
-    var SL#witness@0
-      int
-    and 2
-    implies
-    label pos 0 0
-    true
-    implies
-    and 2
-    <=
-    int-num 1
-    var L#p@0
-      int
-    <=
-    int-num 0
-    var SL#witness@0
-      int
-    implies
-    <
-    var L#p@0
-      int
-    var P#len
-      int
-    implies
-    and 2
-    <=
-    int-num 1
-    var L#p@0
-      int
-    <=
-    int-num 0
-    var SL#witness@0
-      int
-    implies
-    label pos 0 0
-    true
-    implies
-    and 2
-    <=
-    int-num 1
-    var L#p@0
-      int
-    <=
-    int-num 0
-    var SL#witness@0
-      int
-    and 2
-    label neg 21 9
-    and 2
-    fun $is 2
-    fun $idx 3
-    fun $ptr 2
-    fun ^^u1 0
-    var P#arr
-      int
-    var L#p@0
-      int
-    fun ^^u1 0
-    fun ^^u1 0
-    fun $typed 2
-    var $s
-      type-con $state 0
-    fun $idx 3
-    fun $ptr 2
-    fun ^^u1 0
-    var P#arr
-      int
-    var L#p@0
-      int
-    fun ^^u1 0
-    implies
-    and 2
-    fun $is 2
-    fun $idx 3
-    fun $ptr 2
-    fun ^^u1 0
-    var P#arr
-      int
-    var L#p@0
-      int
-    fun ^^u1 0
-    fun ^^u1 0
-    fun $typed 2
-    var $s
-      type-con $state 0
-    fun $idx 3
-    fun $ptr 2
-    fun ^^u1 0
-    var P#arr
-      int
-    var L#p@0
-      int
-    fun ^^u1 0
-    and 2
-    label neg 21 9
-    and 2
-    fun $is 2
-    fun $idx 3
-    fun $ptr 2
-    fun ^^u1 0
-    var P#arr
-      int
-    var L#p@0
-      int
-    fun ^^u1 0
-    fun ^^u1 0
-    fun $thread_local 2
-    var $s
-      type-con $state 0
-    fun $idx 3
-    fun $ptr 2
-    fun ^^u1 0
-    var P#arr
-      int
-    var L#p@0
-      int
-    fun ^^u1 0
-    implies
-    and 2
-    fun $is 2
-    fun $idx 3
-    fun $ptr 2
-    fun ^^u1 0
-    var P#arr
-      int
-    var L#p@0
-      int
-    fun ^^u1 0
-    fun ^^u1 0
-    fun $thread_local 2
-    var $s
-      type-con $state 0
-    fun $idx 3
-    fun $ptr 2
-    fun ^^u1 0
-    var P#arr
-      int
-    var L#p@0
-      int
-    fun ^^u1 0
-    implies
-    and 2
-    <=
-    int-num 1
-    var L#p@0
-      int
-    <=
-    int-num 0
-    var SL#witness@0
-      int
-    and 2
-    implies
-    label pos 0 0
-    true
-    implies
-    and 2
-    <=
-    int-num 1
-    var L#p@0
-      int
-    <=
-    int-num 0
-    var SL#witness@0
-      int
-    implies
-    >
-    fun $read_u1 2
-    var $s
-      type-con $state 0
-    fun $idx 3
-    fun $ptr 2
-    fun ^^u1 0
-    var P#arr
-      int
-    var L#p@0
-      int
-    fun ^^u1 0
-    var L#max@1
-      int
-    implies
-    and 2
-    <=
-    int-num 1
-    var L#p@0
-      int
-    <=
-    int-num 0
-    var SL#witness@0
-      int
-    implies
-    label pos 0 0
-    true
-    implies
-    and 2
-    <=
-    int-num 1
-    var L#p@0
-      int
-    <=
-    int-num 0
-    var SL#witness@0
-      int
-    and 2
-    label neg 23 13
-    and 2
-    fun $is 2
-    fun $idx 3
-    fun $ptr 2
-    fun ^^u1 0
-    var P#arr
-      int
-    var L#p@0
-      int
-    fun ^^u1 0
-    fun ^^u1 0
-    fun $typed 2
-    var $s
-      type-con $state 0
-    fun $idx 3
-    fun $ptr 2
-    fun ^^u1 0
-    var P#arr
-      int
-    var L#p@0
-      int
-    fun ^^u1 0
-    implies
-    and 2
-    fun $is 2
-    fun $idx 3
-    fun $ptr 2
-    fun ^^u1 0
-    var P#arr
-      int
-    var L#p@0
-      int
-    fun ^^u1 0
-    fun ^^u1 0
-    fun $typed 2
-    var $s
-      type-con $state 0
-    fun $idx 3
-    fun $ptr 2
-    fun ^^u1 0
-    var P#arr
-      int
-    var L#p@0
-      int
-    fun ^^u1 0
-    and 2
-    label neg 23 13
-    and 2
-    fun $is 2
-    fun $idx 3
-    fun $ptr 2
-    fun ^^u1 0
-    var P#arr
-      int
-    var L#p@0
-      int
-    fun ^^u1 0
-    fun ^^u1 0
-    fun $thread_local 2
-    var $s
-      type-con $state 0
-    fun $idx 3
-    fun $ptr 2
-    fun ^^u1 0
-    var P#arr
-      int
-    var L#p@0
-      int
-    fun ^^u1 0
-    implies
-    and 2
-    fun $is 2
-    fun $idx 3
-    fun $ptr 2
-    fun ^^u1 0
-    var P#arr
-      int
-    var L#p@0
-      int
-    fun ^^u1 0
-    fun ^^u1 0
-    fun $thread_local 2
-    var $s
-      type-con $state 0
-    fun $idx 3
-    fun $ptr 2
-    fun ^^u1 0
-    var P#arr
-      int
-    var L#p@0
-      int
-    fun ^^u1 0
-    implies
-    =
-    var L#max@2
-      int
-    fun $read_u1 2
-    var $s
-      type-con $state 0
-    fun $idx 3
-    fun $ptr 2
-    fun ^^u1 0
-    var P#arr
-      int
-    var L#p@0
-      int
-    fun ^^u1 0
-    implies
-    fun $local_value_is 5
-    var $s
-      type-con $state 0
-    fun #tok$1^23.7 0
-    fun #loc.max 0
-    var L#max@2
-      int
-    fun ^^u1 0
-    implies
-    fun $local_value_is 5
-    var $s
-      type-con $state 0
-    fun #tok$1^24.47 0
-    fun #loc.witness 0
-    var L#p@0
-      int
-    fun ^^u4 0
-    implies
-    and 2
-    <=
-    int-num 1
-    var L#p@0
-      int
-    <=
-    int-num 1
-    var L#p@0
-      int
-    implies
-    label pos 0 0
-    true
-    implies
-    =
-    var L#max@3
-      int
-    var L#max@2
-      int
-    implies
-    =
-    var SL#witness@1
-      int
-    var L#p@0
-      int
-    implies
-    label pos 0 0
-    true
-    implies
-    and 2
-    <=
-    int-num 1
-    var L#p@0
-      int
-    <=
-    int-num 0
-    var SL#witness@1
-      int
-    and 2
-    label neg 16 24
-    and 2
-    <=
-    int-num 0
-    +
-    var L#p@0
-      int
-    int-num 1
-    <=
-    +
-    var L#p@0
-      int
-    int-num 1
-    fun $max.u4 0
-    implies
-    and 2
-    <=
-    int-num 0
-    +
-    var L#p@0
-      int
-    int-num 1
-    <=
-    +
-    var L#p@0
-      int
-    int-num 1
-    fun $max.u4 0
-    implies
-    =
-    var L#p@1
-      int
-    +
-    var L#p@0
-      int
-    int-num 1
-    implies
-    fun $local_value_is 5
-    var $s
-      type-con $state 0
-    fun #tok$1^16.24 0
-    fun #loc.p 0
-    var L#p@1
-      int
-    fun ^^u4 0
-    implies
-    and 2
-    <=
-    int-num 2
-    var L#p@1
-      int
-    <=
-    int-num 0
-    var SL#witness@1
-      int
-    implies
-    label pos 0 0
-    true
-    and 2
-    label neg 17 17
-    <=
-    var L#p@1
-      int
-    var P#len
-      int
-    implies
-    <=
-    var L#p@1
-      int
-    var P#len
-      int
-    and 2
-    label neg 18 17
-    forall 1 0 3
-      var Q#i$1^18.17#tc1
-        int
-      attribute qid 1
-        string-attr maxc.18:17
-      attribute uniqueId 1
-        string-attr 3
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 2
-    <=
-    int-num 0
-    var Q#i$1^18.17#tc1
-      int
-    <=
-    var Q#i$1^18.17#tc1
-      int
-    fun $max.u4 0
-    implies
-    <
-    var Q#i$1^18.17#tc1
-      int
-    var L#p@1
-      int
-    <=
-    fun $read_u1 2
-    var $s
-      type-con $state 0
-    fun $idx 3
-    fun $ptr 2
-    fun ^^u1 0
-    var P#arr
-      int
-    var Q#i$1^18.17#tc1
-      int
-    fun ^^u1 0
-    var L#max@3
-      int
-    implies
-    forall 1 0 3
-      var Q#i$1^18.17#tc1
-        int
-      attribute qid 1
-        string-attr maxc.18:17
-      attribute uniqueId 1
-        string-attr 3
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 2
-    <=
-    int-num 0
-    var Q#i$1^18.17#tc1
-      int
-    <=
-    var Q#i$1^18.17#tc1
-      int
-    fun $max.u4 0
-    implies
-    <
-    var Q#i$1^18.17#tc1
-      int
-    var L#p@1
-      int
-    <=
-    fun $read_u1 2
-    var $s
-      type-con $state 0
-    fun $idx 3
-    fun $ptr 2
-    fun ^^u1 0
-    var P#arr
-      int
-    var Q#i$1^18.17#tc1
-      int
-    fun ^^u1 0
-    var L#max@3
-      int
-    and 2
-    label neg 19 17
-    and 2
-    <
-    var SL#witness@1
-      int
-    var P#len
-      int
-    =
-    fun $read_u1 2
-    var $s
-      type-con $state 0
-    fun $idx 3
-    fun $ptr 2
-    fun ^^u1 0
-    var P#arr
-      int
-    var SL#witness@1
-      int
-    fun ^^u1 0
-    var L#max@3
-      int
-    implies
-    and 2
-    <
-    var SL#witness@1
-      int
-    var P#len
-      int
-    =
-    fun $read_u1 2
-    var $s
-      type-con $state 0
-    fun $idx 3
-    fun $ptr 2
-    fun ^^u1 0
-    var P#arr
-      int
-    var SL#witness@1
-      int
-    fun ^^u1 0
-    var L#max@3
-      int
-    implies
-    false
-    true
-    implies
-    label pos 0 0
-    true
-    implies
-    and 2
-    <=
-    int-num 1
-    var L#p@0
-      int
-    <=
-    int-num 0
-    var SL#witness@0
-      int
-    implies
-    <=
-    fun $read_u1 2
-    var $s
-      type-con $state 0
-    fun $idx 3
-    fun $ptr 2
-    fun ^^u1 0
-    var P#arr
-      int
-    var L#p@0
-      int
-    fun ^^u1 0
-    var L#max@1
-      int
-    implies
-    and 2
-    <=
-    int-num 1
-    var L#p@0
-      int
-    <=
-    int-num 0
-    var SL#witness@0
-      int
-    implies
-    label pos 0 0
-    true
-    implies
-    and 2
-    <=
-    int-num 1
-    var L#p@0
-      int
-    <=
-    int-num 0
-    var SL#witness@0
-      int
-    implies
-    and 2
-    <=
-    int-num 1
-    var L#p@0
-      int
-    <=
-    int-num 0
-    var SL#witness@0
-      int
-    implies
-    label pos 0 0
-    true
-    implies
-    =
-    var L#max@3
-      int
-    var L#max@1
-      int
-    implies
-    =
-    var SL#witness@1
-      int
-    var SL#witness@0
-      int
-    implies
-    label pos 0 0
-    true
-    implies
-    and 2
-    <=
-    int-num 1
-    var L#p@0
-      int
-    <=
-    int-num 0
-    var SL#witness@1
-      int
-    and 2
-    label neg 16 24
-    and 2
-    <=
-    int-num 0
-    +
-    var L#p@0
-      int
-    int-num 1
-    <=
-    +
-    var L#p@0
-      int
-    int-num 1
-    fun $max.u4 0
-    implies
-    and 2
-    <=
-    int-num 0
-    +
-    var L#p@0
-      int
-    int-num 1
-    <=
-    +
-    var L#p@0
-      int
-    int-num 1
-    fun $max.u4 0
-    implies
-    =
-    var L#p@1
-      int
-    +
-    var L#p@0
-      int
-    int-num 1
-    implies
-    fun $local_value_is 5
-    var $s
-      type-con $state 0
-    fun #tok$1^16.24 0
-    fun #loc.p 0
-    var L#p@1
-      int
-    fun ^^u4 0
-    implies
-    and 2
-    <=
-    int-num 2
-    var L#p@1
-      int
-    <=
-    int-num 0
-    var SL#witness@1
-      int
-    implies
-    label pos 0 0
-    true
-    and 2
-    label neg 17 17
-    <=
-    var L#p@1
-      int
-    var P#len
-      int
-    implies
-    <=
-    var L#p@1
-      int
-    var P#len
-      int
-    and 2
-    label neg 18 17
-    forall 1 0 3
-      var Q#i$1^18.17#tc1
-        int
-      attribute qid 1
-        string-attr maxc.18:17
-      attribute uniqueId 1
-        string-attr 3
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 2
-    <=
-    int-num 0
-    var Q#i$1^18.17#tc1
-      int
-    <=
-    var Q#i$1^18.17#tc1
-      int
-    fun $max.u4 0
-    implies
-    <
-    var Q#i$1^18.17#tc1
-      int
-    var L#p@1
-      int
-    <=
-    fun $read_u1 2
-    var $s
-      type-con $state 0
-    fun $idx 3
-    fun $ptr 2
-    fun ^^u1 0
-    var P#arr
-      int
-    var Q#i$1^18.17#tc1
-      int
-    fun ^^u1 0
-    var L#max@3
-      int
-    implies
-    forall 1 0 3
-      var Q#i$1^18.17#tc1
-        int
-      attribute qid 1
-        string-attr maxc.18:17
-      attribute uniqueId 1
-        string-attr 3
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 2
-    <=
-    int-num 0
-    var Q#i$1^18.17#tc1
-      int
-    <=
-    var Q#i$1^18.17#tc1
-      int
-    fun $max.u4 0
-    implies
-    <
-    var Q#i$1^18.17#tc1
-      int
-    var L#p@1
-      int
-    <=
-    fun $read_u1 2
-    var $s
-      type-con $state 0
-    fun $idx 3
-    fun $ptr 2
-    fun ^^u1 0
-    var P#arr
-      int
-    var Q#i$1^18.17#tc1
-      int
-    fun ^^u1 0
-    var L#max@3
-      int
-    and 2
-    label neg 19 17
-    and 2
-    <
-    var SL#witness@1
-      int
-    var P#len
-      int
-    =
-    fun $read_u1 2
-    var $s
-      type-con $state 0
-    fun $idx 3
-    fun $ptr 2
-    fun ^^u1 0
-    var P#arr
-      int
-    var SL#witness@1
-      int
-    fun ^^u1 0
-    var L#max@3
-      int
-    implies
-    and 2
-    <
-    var SL#witness@1
-      int
-    var P#len
-      int
-    =
-    fun $read_u1 2
-    var $s
-      type-con $state 0
-    fun $idx 3
-    fun $ptr 2
-    fun ^^u1 0
-    var P#arr
-      int
-    var SL#witness@1
-      int
-    fun ^^u1 0
-    var L#max@3
-      int
-    implies
-    false
-    true
-    implies
-    label pos 0 0
-    true
-    implies
-    and 2
-    <=
-    int-num 1
-    var L#p@0
-      int
-    <=
-    int-num 0
-    var SL#witness@0
-      int
-    implies
-    >=
-    var L#p@0
-      int
-    var P#len
-      int
-    implies
-    and 2
-    <=
-    int-num 1
-    var L#p@0
-      int
-    <=
-    int-num 0
-    var SL#witness@0
-      int
-    implies
-    label pos 0 0
-    true
-    implies
-    and 2
-    <=
-    int-num 1
-    var L#p@0
-      int
-    <=
-    int-num 0
-    var SL#witness@0
-      int
-    implies
-    and 2
-    <=
-    int-num 1
-    var L#p@0
-      int
-    <=
-    int-num 0
-    var SL#witness@0
-      int
-    implies
-    label pos 0 0
-    true
-    implies
-    and 2
-    <=
-    int-num 1
-    var L#p@0
-      int
-    <=
-    int-num 0
-    var SL#witness@0
-      int
-    and 2
-    label neg 27 3
-    fun $position_marker 0
-    implies
-    fun $position_marker 0
-    implies
-    and 2
-    <=
-    int-num 1
-    var L#p@0
-      int
-    <=
-    int-num 0
-    var SL#witness@0
-      int
-    implies
-    label pos 0 0
-    true
-    implies
-    and 2
-    <=
-    int-num 1
-    var L#p@0
-      int
-    <=
-    int-num 0
-    var SL#witness@0
-      int
-    implies
-    and 2
-    <=
-    int-num 1
-    var L#p@0
-      int
-    <=
-    int-num 0
-    var SL#witness@0
-      int
-    implies
-    label pos 0 0
-    true
-    implies
-    =
-    var L#max@4
-      int
-    var L#max@1
-      int
-    implies
-    =
-    var L#p@2
-      int
-    var L#p@0
-      int
-    implies
-    =
-    var SL#witness@2
-      int
-    var SL#witness@0
-      int
-    implies
-    =
-    var $result@0
-      int
-    var L#max@1
-      int
-    implies
-    label pos 0 0
-    true
-    and 2
-    label neg 9 14
-    forall 1 0 3
-      var Q#i$1^9.14#tc1
-        int
-      attribute qid 1
-        string-attr maxc.9:14
-      attribute uniqueId 1
-        string-attr 1
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 2
-    <=
-    int-num 0
-    var Q#i$1^9.14#tc1
-      int
-    <=
-    var Q#i$1^9.14#tc1
-      int
-    fun $max.u4 0
-    implies
-    <
-    var Q#i$1^9.14#tc1
-      int
-    var P#len
-      int
-    <=
-    fun $read_u1 2
-    var $s
-      type-con $state 0
-    fun $idx 3
-    fun $ptr 2
-    fun ^^u1 0
-    var P#arr
-      int
-    var Q#i$1^9.14#tc1
-      int
-    fun ^^u1 0
-    var $result@0
-      int
-    implies
-    forall 1 0 3
-      var Q#i$1^9.14#tc1
-        int
-      attribute qid 1
-        string-attr maxc.9:14
-      attribute uniqueId 1
-        string-attr 1
-      attribute bvZ3Native 1
-        string-attr False
-    implies
-    and 2
-    <=
-    int-num 0
-    var Q#i$1^9.14#tc1
-      int
-    <=
-    var Q#i$1^9.14#tc1
-      int
-    fun $max.u4 0
-    implies
-    <
-    var Q#i$1^9.14#tc1
-      int
-    var P#len
-      int
-    <=
-    fun $read_u1 2
-    var $s
-      type-con $state 0
-    fun $idx 3
-    fun $ptr 2
-    fun ^^u1 0
-    var P#arr
-      int
-    var Q#i$1^9.14#tc1
-      int
-    fun ^^u1 0
-    var $result@0
-      int
-    and 2
-    label neg 10 14
-    exists 1 0 3
-      var Q#i$1^10.14#tc1
-        int
-      attribute qid 1
-        string-attr maxc.10:14
-      attribute uniqueId 1
-        string-attr 0
-      attribute bvZ3Native 1
-        string-attr False
-    and 4
-    <=
-    int-num 0
-    var Q#i$1^10.14#tc1
-      int
-    <=
-    var Q#i$1^10.14#tc1
-      int
-    fun $max.u4 0
-    <
-    var Q#i$1^10.14#tc1
-      int
-    var P#len
-      int
-    =
-    fun $read_u1 2
-    var $s
-      type-con $state 0
-    fun $idx 3
-    fun $ptr 2
-    fun ^^u1 0
-    var P#arr
-      int
-    var Q#i$1^10.14#tc1
-      int
-    fun ^^u1 0
-    var $result@0
-      int
-    implies
-    exists 1 0 3
-      var Q#i$1^10.14#tc1
-        int
-      attribute qid 1
-        string-attr maxc.10:14
-      attribute uniqueId 1
-        string-attr 0
-      attribute bvZ3Native 1
-        string-attr False
-    and 4
-    <=
-    int-num 0
-    var Q#i$1^10.14#tc1
-      int
-    <=
-    var Q#i$1^10.14#tc1
-      int
-    fun $max.u4 0
-    <
-    var Q#i$1^10.14#tc1
-      int
-    var P#len
-      int
-    =
-    fun $read_u1 2
-    var $s
-      type-con $state 0
-    fun $idx 3
-    fun $ptr 2
-    fun ^^u1 0
-    var P#arr
-      int
-    var Q#i$1^10.14#tc1
-      int
-    fun ^^u1 0
-    var $result@0
-      int
-    true
--- a/src/HOL/Boogie/Examples/VCC_Max.certs	Tue Jul 23 13:14:14 2013 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,6991 +0,0 @@
-08c73deecf4f2eae63cedf503456d56e41e55ef8 6990 0
-WARNING: For problems containing quantifiers, the model finding capabilities of Z3 work better when the formula does not contain nested quantifiers. You can use PULL_NESTED_QUANTIFIERS=true to eliminate nested quantifiers.
-#2 := false
-decl f470 :: Int
-#4955 := f470
-decl f20 :: (-> S16 S10 Int)
-decl f106 :: (-> S60 S6 S10)
-decl f14 :: S6
-#20 := f14
-decl f107 :: (-> S61 Int S60)
-decl f471 :: Int
-#4957 := f471
-decl f120 :: (-> S68 S10 S61)
-decl f53 :: (-> S34 Int S10)
-decl f445 :: Int
-#4655 := f445
-decl f80 :: (-> S49 S6 S34)
-decl f81 :: S49
-#182 := f81
-#4654 := (f80 f81 f14)
-#4656 := (f53 #4654 f445)
-decl f121 :: S68
-#396 := f121
-#4734 := (f120 f121 #4656)
-#4981 := (f107 #4734 f471)
-#4982 := (f106 #4981 f14)
-decl f70 :: (-> S42 S9 S16)
-decl f444 :: S9
-#4649 := f444
-decl f172 :: S42
-#1486 := f172
-#4748 := (f70 f172 f444)
-#4983 := (f20 #4748 #4982)
-#4984 := (= #4983 f470)
-#20769 := (not #4984)
-#184 := 0::Int
-#5434 := -1::Int
-#12502 := (* -1::Int f471)
-decl f443 :: Int
-#4646 := f443
-#12503 := (+ f443 #12502)
-#12504 := (<= #12503 0::Int)
-#20770 := (or #12504 #20769)
-#20771 := (not #20770)
-#243 := (:var 0 Int)
-#4773 := (f107 #4734 #243)
-#21662 := (pattern #4773)
-#12480 := (* -1::Int f470)
-#4774 := (f106 #4773 f14)
-#4775 := (f20 #4748 #4774)
-#12481 := (+ #4775 #12480)
-#12482 := (<= #12481 0::Int)
-decl f472 :: Int
-#4965 := f472
-#12461 := (* -1::Int f472)
-#12469 := (+ #243 #12461)
-#12468 := (>= #12469 0::Int)
-#8086 := 4294967295::Int
-#14701 := (<= #243 4294967295::Int)
-#17965 := (not #14701)
-#5433 := (>= #243 0::Int)
-#6155 := (not #5433)
-#20761 := (or #6155 #17965 #12468 #12482)
-#21679 := (forall (vars (?v0 Int)) (:pat #21662) #20761)
-#21684 := (not #21679)
-#21687 := (or #21684 #20771)
-#21690 := (not #21687)
-decl ?v0!14 :: Int
-#17016 := ?v0!14
-#17023 := (f107 #4734 ?v0!14)
-#17024 := (f106 #17023 f14)
-#17025 := (f20 #4748 #17024)
-#17327 := (* -1::Int #17025)
-#17328 := (+ f470 #17327)
-#17329 := (>= #17328 0::Int)
-#17314 := (* -1::Int ?v0!14)
-#17315 := (+ f472 #17314)
-#17316 := (<= #17315 0::Int)
-#17018 := (<= ?v0!14 4294967295::Int)
-#20735 := (not #17018)
-#17017 := (>= ?v0!14 0::Int)
-#20734 := (not #17017)
-#20750 := (or #20734 #20735 #17316 #17329)
-#20755 := (not #20750)
-#21693 := (or #20755 #21690)
-#21696 := (not #21693)
-#12462 := (+ f443 #12461)
-#12460 := (>= #12462 0::Int)
-#12465 := (not #12460)
-#21699 := (or #12465 #21696)
-#21702 := (not #21699)
-#21705 := (or #12465 #21702)
-#21708 := (not #21705)
-#12435 := (>= f471 0::Int)
-#20811 := (not #12435)
-#1150 := 2::Int
-#12452 := (>= f472 2::Int)
-#20810 := (not #12452)
-decl f1 :: S1
-#3 := f1
-decl f7 :: (-> S5 S6 S1)
-decl f10 :: S6
-#15 := f10
-decl f8 :: (-> S7 Int S5)
-decl f449 :: (-> S178 S3 S7)
-decl f431 :: S3
-#3710 := f431
-decl f450 :: (-> S179 S3 S178)
-decl f427 :: S3
-#3700 := f427
-decl f451 :: (-> S180 S9 S179)
-decl f452 :: S180
-#4695 := f452
-#4696 := (f451 f452 f444)
-#4967 := (f450 #4696 f427)
-#4968 := (f449 #4967 f431)
-#4969 := (f8 #4968 f472)
-#4970 := (f7 #4969 f10)
-#4971 := (= #4970 f1)
-#11608 := (not #4971)
-decl f464 :: Int
-#4790 := f464
-#12524 := (+ f464 #12461)
-#12523 := (= #12524 -1::Int)
-#12527 := (not #12523)
-#13510 := 4294967294::Int
-#13511 := (<= f464 4294967294::Int)
-#16996 := (not #13511)
-#12444 := (>= f464 -1::Int)
-#16993 := (not #12444)
-#21711 := (or #16993 #16996 #12527 #11608 #20810 #20811 #21708)
-#21714 := (not #21711)
-#21717 := (or #16993 #16996 #21714)
-#21720 := (not #21717)
-#8 := 1::Int
-#12423 := (>= f464 1::Int)
-#12561 := (not #12423)
-#4958 := (= f471 f464)
-#11647 := (not #4958)
-decl f469 :: Int
-#4942 := f469
-#4956 := (= f470 f469)
-#11656 := (not #4956)
-decl f433 :: S3
-#3716 := f433
-decl f428 :: S3
-#3701 := f428
-#4949 := (f450 #4696 f428)
-#4950 := (f449 #4949 f433)
-#4951 := (f8 #4950 f464)
-#4952 := (f7 #4951 f10)
-#4953 := (= #4952 f1)
-#11674 := (not #4953)
-decl f435 :: S3
-#3722 := f435
-decl f429 :: S3
-#3704 := f429
-#4944 := (f450 #4696 f429)
-#4945 := (f449 #4944 f435)
-#4946 := (f8 #4945 f469)
-#4947 := (f7 #4946 f14)
-#4948 := (= #4947 f1)
-#11683 := (not #4948)
-#4929 := (f107 #4734 f464)
-#4930 := (f106 #4929 f14)
-#4940 := (f20 #4748 #4930)
-#4943 := (= f469 #4940)
-#11692 := (not #4943)
-decl f37 :: (-> S27 S10 S1)
-decl f46 :: (-> S31 S9 S27)
-decl f124 :: S31
-#418 := f124
-#4743 := (f46 f124 f444)
-#4937 := (f37 #4743 #4930)
-#4938 := (= #4937 f1)
-#16964 := (not #4938)
-decl f85 :: (-> S51 S10 S5)
-decl f90 :: S51
-#218 := f90
-#4931 := (f85 f90 #4930)
-#4932 := (f7 #4931 f14)
-#4933 := (= #4932 f1)
-#16955 := (not #4933)
-#21723 := (or #16955 #16964 #11692 #11683 #11674 #11656 #11647 #12561 #20811 #21720)
-#21726 := (not #21723)
-#21729 := (or #16955 #16964 #21726)
-#21732 := (not #21729)
-decl f217 :: (-> S93 S6 Int)
-decl f218 :: S93
-#1955 := f218
-#3690 := (f217 f218 f14)
-#25413 := (* #3690 f464)
-#4735 := (f107 #4734 0::Int)
-#4736 := (f106 #4735 f14)
-decl f179 :: S16
-#1535 := f179
-#23755 := (f20 f179 #4736)
-#23775 := (f53 #4654 #23755)
-#24163 := (f20 f179 #23775)
-#25416 := (+ #24163 #25413)
-#25423 := (f53 #4654 #25416)
-decl f332 :: S31
-#3028 := f332
-#23974 := (f46 f332 f444)
-#25662 := (f37 #23974 #25423)
-#25663 := (= #25662 f1)
-decl f51 :: S10
-#120 := f51
-decl f48 :: (-> S32 S10 S10)
-decl f49 :: (-> S33 S9 S32)
-decl f50 :: S33
-#117 := f50
-#4661 := (f49 f50 f444)
-#25660 := (f48 #4661 #25423)
-#25661 := (= #25660 f51)
-#25664 := (or #25661 #25663)
-#25665 := (not #25664)
-decl f29 :: S21
-#46 := f29
-decl f26 :: (-> S6 S21)
-decl f27 :: (-> S22 S10 S6)
-decl f28 :: S22
-#43 := f28
-#25635 := (f27 f28 #25423)
-#25636 := (f26 #25635)
-#25637 := (= #25636 f29)
-#25666 := (or #25637 #25665)
-#25667 := (not #25666)
-decl f208 :: (-> S88 S56 S10)
-decl f101 :: (-> S55 S10 S56)
-decl f102 :: (-> S57 S58 S55)
-decl f104 :: (-> S59 S9 S58)
-decl f105 :: S59
-#348 := f105
-#4875 := (f104 f105 f444)
-decl f103 :: S57
-#347 := f103
-#4876 := (f102 f103 #4875)
-#25639 := (f101 #4876 #25423)
-decl f209 :: S88
-#1822 := f209
-#25643 := (f208 f209 #25639)
-#25654 := (f37 #23974 #25643)
-#25655 := (= #25654 f1)
-#25652 := (f48 #4661 #25643)
-#25653 := (= #25652 f51)
-#25656 := (or #25653 #25655)
-#25657 := (not #25656)
-#25649 := (f27 f28 #25643)
-#25650 := (f26 #25649)
-#25651 := (= #25650 f29)
-decl f52 :: S31
-#123 := f52
-#4650 := (f46 f52 f444)
-#25644 := (f37 #4650 #25643)
-#25645 := (= #25644 f1)
-#25646 := (not #25645)
-decl f122 :: (-> S69 S56 S1)
-decl f210 :: S69
-#1829 := f210
-#25640 := (f122 f210 #25639)
-#25641 := (= #25640 f1)
-#25642 := (not #25641)
-#25647 := (or #25642 #25646)
-#25648 := (not #25647)
-#25638 := (not #25637)
-#25658 := (or #25638 #25648 #25651 #25657)
-#25659 := (not #25658)
-#25668 := (or #25659 #25667)
-#25669 := (not #25668)
-decl f47 :: S31
-#113 := f47
-#4667 := (f46 f47 f444)
-#25632 := (f37 #4667 #25423)
-#25633 := (= #25632 f1)
-#4934 := (f37 #4667 #4930)
-#4935 := (= #4934 f1)
-#24364 := (f101 #4876 #4930)
-#25566 := (f122 f210 #24364)
-#25567 := (= #25566 f1)
-#16958 := (not #4935)
-#25568 := (or #16958 #25567)
-#25569 := (not #25568)
-#25700 := [hypothesis]: #25568
-decl f15 :: (-> S12 S11 S1)
-decl f31 :: (-> S23 S10 S11)
-#4657 := (f20 f179 #4656)
-decl f188 :: (-> S78 Int S6)
-decl f189 :: (-> S79 S6 S78)
-decl f190 :: S79
-#1645 := f190
-#4651 := (f189 f190 f14)
-#4652 := (f188 #4651 f443)
-#4653 := (f80 f81 #4652)
-#4658 := (f53 #4653 #4657)
-#22274 := (f20 f179 #4658)
-#23197 := (f53 #4653 #22274)
-decl f32 :: (-> S24 S9 S23)
-decl f33 :: S24
-#63 := f33
-#23595 := (f32 f33 f444)
-#23603 := (f31 #23595 #23197)
-decl f16 :: (-> S13 S10 S12)
-#23868 := (f53 #4653 f445)
-decl f30 :: S13
-#48 := f30
-#23869 := (f16 f30 #23868)
-#23870 := (f15 #23869 #23603)
-#23871 := (= #23870 f1)
-#23594 := (f16 f30 #23197)
-#23604 := (f15 #23594 #23603)
-#23605 := (= #23604 f1)
-decl f96 :: S24
-#274 := f96
-#23354 := (f32 f96 f444)
-#23609 := (f31 #23354 #23197)
-#40 := (:var 0 S10)
-#49 := (f16 f30 #40)
-#23610 := (f15 #49 #23609)
-#23619 := (pattern #23610)
-decl f303 :: (-> S126 S18 S11)
-decl f23 :: (-> S19 S10 S18)
-decl f24 :: (-> S20 S9 S19)
-decl f25 :: S20
-#35 := f25
-#23613 := (f24 f25 f444)
-#23614 := (f23 #23613 #23197)
-decl f304 :: S126
-#2556 := f304
-#23615 := (f303 f304 #23614)
-decl f227 :: S13
-#2007 := f227
-#2815 := (f16 f227 #40)
-#23616 := (f15 #2815 #23615)
-#23617 := (= #23616 f1)
-#23611 := (= #23610 f1)
-#23612 := (not #23611)
-#23477 := (f27 f28 #23197)
-decl f311 :: S5
-#2736 := f311
-#23607 := (f7 f311 #23477)
-#23608 := (= #23607 f1)
-#23618 := (or #23608 #23612 #23617)
-#23620 := (forall (vars (?v3 S10)) (:pat #23619) #23618)
-#23621 := (not #23620)
-#23488 := (f37 #4650 #23197)
-#23489 := (= #23488 f1)
-#23514 := (not #23489)
-#23606 := (not #23605)
-#23622 := (or #23606 #23514 #23621)
-#23623 := (not #23622)
-decl f92 :: (-> S54 S10 S27)
-decl f308 :: (-> S128 S9 S54)
-decl f310 :: S128
-#2723 := f310
-#23570 := (f308 f310 f444)
-#23571 := (f92 #23570 #23197)
-#23572 := (f37 #23571 #23197)
-#23573 := (= #23572 f1)
-decl f312 :: (-> S130 S129 S1)
-decl f460 :: S129
-#4731 := f460
-decl f313 :: (-> S131 S10 S130)
-decl f314 :: (-> S132 S10 S131)
-decl f315 :: (-> S133 S9 S132)
-decl f319 :: S133
-#2845 := f319
-#4728 := (f315 f319 f444)
-#23559 := (f314 #4728 #23197)
-#23560 := (f313 #23559 #23197)
-#23568 := (f312 #23560 f460)
-#23569 := (= #23568 f1)
-#23574 := (iff #23569 #23573)
-#2829 := (:var 0 S129)
-#28 := (:var 1 S10)
-#37 := (:var 2 S10)
-#24 := (:var 3 S9)
-#2846 := (f315 f319 #24)
-#2847 := (f314 #2846 #37)
-#2848 := (f313 #2847 #28)
-#2849 := (f312 #2848 #2829)
-#2850 := (pattern #2849)
-#2773 := (f308 f310 #24)
-#2852 := (f92 #2773 #37)
-#2853 := (f37 #2852 #28)
-#2854 := (= #2853 f1)
-#2851 := (= #2849 f1)
-#2855 := (iff #2851 #2854)
-#2856 := (forall (vars (?v0 S9) (?v1 S10) (?v2 S10) (?v3 S129)) (:pat #2850) #2855)
-#16349 := (~ #2856 #2856)
-#16347 := (~ #2855 #2855)
-#16348 := [refl]: #16347
-#16350 := [nnf-pos #16348]: #16349
-#9423 := [asserted]: #2856
-#16351 := [mp~ #9423 #16350]: #2856
-#23583 := (not #2856)
-#23585 := (or #23583 #23574)
-#23586 := [quant-inst #4649 #23197 #23197 #4731]: #23585
-#23729 := [unit-resolution #23586 #16351]: #23574
-#4729 := (f314 #4728 #4658)
-#4730 := (f313 #4729 #4658)
-#4732 := (f312 #4730 f460)
-#4733 := (= #4732 f1)
-#23649 := (f26 #23477)
-#23650 := (= #23649 f29)
-#23738 := (not #23650)
-decl f420 :: S21
-#3596 := f420
-#4622 := (= f29 f420)
-#4623 := (not #4622)
-#23739 := (iff #4623 #23738)
-#23736 := (iff #4622 #23650)
-#23734 := (iff #23650 #4622)
-#23712 := (= f420 f29)
-#23732 := (iff #23712 #4622)
-#23733 := [commutativity]: #23732
-#23713 := (iff #23650 #23712)
-#23723 := (= #23649 f420)
-#4670 := (f26 #4652)
-#23371 := (= #4670 f420)
-decl f221 :: S5
-#1974 := f221
-#23299 := (f7 f221 #4652)
-#23300 := (= #23299 f1)
-#23372 := (iff #23300 #23371)
-#333 := (:var 0 S6)
-#2348 := (f7 f221 #333)
-#3582 := (pattern #2348)
-#3591 := (f26 #333)
-#3597 := (= #3591 f420)
-#2350 := (= #2348 f1)
-#3598 := (iff #2350 #3597)
-#3599 := (forall (vars (?v0 S6)) (:pat #3582) #3598)
-#16784 := (~ #3599 #3599)
-#16782 := (~ #3598 #3598)
-#16783 := [refl]: #16782
-#16785 := [nnf-pos #16783]: #16784
-#9912 := [asserted]: #3599
-#16786 := [mp~ #9912 #16785]: #3599
-#23381 := (not #3599)
-#23382 := (or #23381 #23372)
-#23383 := [quant-inst #4652]: #23382
-#23580 := [unit-resolution #23383 #16786]: #23372
-#23384 := (not #23372)
-#23582 := (or #23384 #23371)
-#472 := (:var 1 S6)
-#1646 := (f189 f190 #472)
-#1647 := (f188 #1646 #243)
-#1969 := (pattern #1647)
-#1975 := (f7 f221 #1647)
-#1976 := (= #1975 f1)
-#1977 := (forall (vars (?v0 S6) (?v1 Int)) (:pat #1969) #1976)
-#15883 := (~ #1977 #1977)
-#15881 := (~ #1976 #1976)
-#15882 := [refl]: #15881
-#15884 := [nnf-pos #15882]: #15883
-#8601 := [asserted]: #1977
-#15885 := [mp~ #8601 #15884]: #1977
-#23306 := (not #1977)
-#23307 := (or #23306 #23300)
-#23308 := [quant-inst #20 #4646]: #23307
-#24771 := [unit-resolution #23308 #15885]: #23300
-#23388 := (not #23300)
-#23389 := (or #23384 #23388 #23371)
-#23390 := [def-axiom]: #23389
-#23698 := [unit-resolution #23390 #24771]: #23582
-#23699 := [unit-resolution #23698 #23580]: #23371
-#23721 := (= #23649 #4670)
-#23719 := (= #23477 #4652)
-#23212 := (f27 f28 #4658)
-#23213 := (= #23212 #4652)
-#1719 := (f80 f81 #472)
-#3411 := (f53 #1719 #243)
-#21607 := (pattern #3411)
-#3415 := (f27 f28 #3411)
-#3416 := (= #3415 #472)
-#21614 := (forall (vars (?v0 S6) (?v1 Int)) (:pat #21607) #3416)
-#3417 := (forall (vars (?v0 S6) (?v1 Int)) #3416)
-#21617 := (iff #3417 #21614)
-#21615 := (iff #3416 #3416)
-#21616 := [refl]: #21615
-#21618 := [quant-intro #21616]: #21617
-#16699 := (~ #3417 #3417)
-#16697 := (~ #3416 #3416)
-#16698 := [refl]: #16697
-#16700 := [nnf-pos #16698]: #16699
-#9831 := [asserted]: #3417
-#16701 := [mp~ #9831 #16700]: #3417
-#21619 := [mp #16701 #21618]: #21614
-#23239 := (not #21614)
-#23278 := (or #23239 #23213)
-#23279 := [quant-inst #4652 #4657]: #23278
-#23700 := [unit-resolution #23279 #21619]: #23213
-#23717 := (= #23477 #23212)
-#23715 := (= #23197 #4658)
-#23210 := (= #4658 #23197)
-#4664 := (f85 f90 #4658)
-#4665 := (f7 #4664 #4652)
-#4666 := (= #4665 f1)
-decl f45 :: S5
-#109 := f45
-#4673 := (f7 f45 #4652)
-#4674 := (= #4673 f1)
-#4671 := (= #4670 f29)
-#4672 := (not #4671)
-#4668 := (f37 #4667 #4658)
-#4669 := (= #4668 f1)
-#4662 := (f48 #4661 #4658)
-#4663 := (= #4662 f51)
-#4659 := (f37 #4650 #4658)
-#4660 := (= #4659 f1)
-#13118 := (and #4660 #4663 #4666 #4669 #4672 #4674)
-decl f468 :: Int
-#4819 := f468
-#4826 := (= #4775 f468)
-#12352 := (* -1::Int f443)
-#12735 := (+ #243 #12352)
-#12734 := (>= #12735 0::Int)
-#12736 := (not #12734)
-decl f135 :: Int
-#565 := f135
-#5615 := (* -1::Int f135)
-#5616 := (+ #243 #5615)
-#5617 := (<= #5616 0::Int)
-#12777 := (and #5433 #5617 #12736 #4826)
-#12782 := (exists (vars (?v0 Int)) #12777)
-#12746 := (* -1::Int f468)
-#12747 := (+ #4775 #12746)
-#12748 := (<= #12747 0::Int)
-#5624 := (and #5433 #5617)
-#6637 := (not #5624)
-#12757 := (or #6637 #12734 #12748)
-#12762 := (forall (vars (?v0 Int)) #12757)
-#12765 := (not #12762)
-#12785 := (or #12765 #12782)
-#12788 := (and #12762 #12785)
-decl f462 :: Int
-#4782 := f462
-#4820 := (= f468 f462)
-#11320 := (not #4820)
-decl f463 :: Int
-#4786 := f463
-decl f467 :: Int
-#4817 := f467
-#4818 := (= f467 f463)
-#11329 := (not #4818)
-decl f466 :: Int
-#4815 := f466
-#4816 := (= f466 f464)
-#11338 := (not #4816)
-decl f465 :: Int
-#4813 := f465
-#4814 := (= f465 f462)
-#11347 := (not #4814)
-#12426 := (>= f463 0::Int)
-#12428 := (and #12423 #12426)
-#12431 := (not #12428)
-decl f341 :: S1
-#3124 := f341
-#3125 := (= f341 f1)
-#11381 := (not #3125)
-#12815 := (or #11381 #12431 #11347 #11338 #11329 #11320 #12788)
-#12820 := (and #3125 #12815)
-#12447 := (* -1::Int f464)
-#12705 := (+ f443 #12447)
-#12706 := (<= #12705 0::Int)
-#12707 := (not #12706)
-#12845 := (or #12431 #12707 #12820)
-#12505 := (not #12504)
-#12508 := (and #12505 #4984)
-#12491 := (or #6637 #12468 #12482)
-#12496 := (forall (vars (?v0 Int)) #12491)
-#12499 := (not #12496)
-#12511 := (or #12499 #12508)
-#12514 := (and #12496 #12511)
-#12517 := (or #12465 #12514)
-#12520 := (and #12460 #12517)
-#12454 := (and #12452 #12435)
-#12457 := (not #12454)
-#12448 := (+ f135 #12447)
-#12446 := (>= #12448 1::Int)
-#12530 := (and #12444 #12446)
-#12533 := (not #12530)
-#12548 := (or #12533 #12527 #11608 #12457 #12520)
-#12556 := (and #12444 #12446 #12548)
-#12437 := (and #12423 #12435)
-#12440 := (not #12437)
-#5013 := (= f471 f463)
-#11758 := (not #5013)
-#5012 := (= f470 f462)
-#11767 := (not #5012)
-#12614 := (* -1::Int #4940)
-#12615 := (+ f462 #12614)
-#12613 := (>= #12615 0::Int)
-#12612 := (not #12613)
-#12667 := (or #12431 #12612 #11767 #11758 #12440 #12556)
-#4939 := (and #4933 #4938)
-#11701 := (not #4939)
-#12588 := (or #11701 #11692 #11683 #11674 #12561 #11656 #11647 #12440 #12556)
-#12596 := (and #4933 #4938 #12588)
-#4936 := (and #4933 #4935)
-#11713 := (not #4936)
-#12601 := (or #11713 #12596)
-#12607 := (and #4933 #4935 #12601)
-#12637 := (or #12431 #12613 #12607)
-#12672 := (and #12637 #12667)
-#12681 := (or #11701 #12431 #12672)
-#12689 := (and #4933 #4938 #12681)
-#12694 := (or #11713 #12689)
-#12700 := (and #4933 #4935 #12694)
-#12729 := (or #12431 #12706 #12700)
-#12850 := (and #12729 #12845)
-decl f12 :: (-> S8 S6 S6)
-decl f13 :: S8
-#19 := f13
-#21 := (f12 f13 f14)
-decl f453 :: (-> S181 S3 S51)
-decl f438 :: S3
-#3731 := f438
-decl f454 :: (-> S182 S3 S181)
-decl f430 :: S3
-#3707 := f430
-decl f455 :: (-> S183 S9 S182)
-decl f456 :: S183
-#4703 := f456
-#4704 := (f455 f456 f444)
-#4919 := (f454 #4704 f430)
-#4920 := (f453 #4919 f438)
-#4921 := (f85 #4920 #4656)
-#4922 := (f7 #4921 #21)
-#4923 := (= #4922 f1)
-decl f55 :: S16
-#139 := f55
-#4699 := (f20 f55 #4656)
-#4898 := (f450 #4696 f430)
-#4915 := (f449 #4898 f438)
-#4916 := (f8 #4915 #4699)
-#4917 := (f7 #4916 #21)
-#4918 := (= #4917 f1)
-#4924 := (and #4918 #4923)
-#11934 := (not #4924)
-decl f437 :: S3
-#3728 := f437
-#4911 := (f449 #4898 f437)
-#4912 := (f8 #4911 f443)
-#4913 := (f7 #4912 f10)
-#4914 := (= #4913 f1)
-#11943 := (not #4914)
-#4907 := (f449 #4898 f435)
-#4908 := (f8 #4907 f462)
-#4909 := (f7 #4908 f14)
-#4910 := (= #4909 f1)
-#11952 := (not #4910)
-#4903 := (f449 #4898 f433)
-#4904 := (f8 #4903 f463)
-#4905 := (f7 #4904 f10)
-#4906 := (= #4905 f1)
-#11961 := (not #4906)
-#4899 := (f449 #4898 f431)
-#4900 := (f8 #4899 f464)
-#4901 := (f7 #4900 f10)
-#4902 := (= #4901 f1)
-#11970 := (not #4902)
-decl f82 :: (-> S50 S9 S1)
-decl f98 :: S50
-#325 := f98
-#4685 := (f82 f98 f444)
-#4686 := (= #4685 f1)
-decl f322 :: (-> S136 S3 S50)
-decl f323 :: S136
-#2881 := f323
-#4809 := (f322 f323 f430)
-#4810 := (f82 #4809 f444)
-#4811 := (= #4810 f1)
-#4812 := (and #4811 #4686)
-#11409 := (not #4812)
-decl f278 :: (-> S115 S9 S50)
-decl f279 :: S115
-#2314 := f279
-#4889 := (f278 f279 f444)
-#4890 := (f82 #4889 f444)
-#4891 := (= #4890 f1)
-#12874 := (not #4891)
-#4803 := (f107 #4734 f463)
-#4804 := (f106 #4803 f14)
-#4805 := (f20 #4748 #4804)
-#4806 := (= #4805 f462)
-#12877 := (* -1::Int f463)
-#12878 := (+ f443 #12877)
-#12879 := (<= #12878 0::Int)
-#12880 := (not #12879)
-#12883 := (and #12880 #4806)
-#12886 := (not #12883)
-#12901 := (* -1::Int f462)
-#12902 := (+ #4775 #12901)
-#12903 := (<= #12902 0::Int)
-#12890 := (+ #243 #12447)
-#12889 := (>= #12890 0::Int)
-#12912 := (or #6637 #12889 #12903)
-#12917 := (forall (vars (?v0 Int)) #12912)
-#12920 := (not #12917)
-#12923 := (>= #12705 0::Int)
-#12926 := (not #12923)
-#12932 := (>= #12448 0::Int)
-#12929 := (>= f464 0::Int)
-#12935 := (and #12929 #12932)
-#12938 := (not #12935)
-#12942 := (+ f135 #12877)
-#12941 := (>= #12942 0::Int)
-#12945 := (and #12426 #12941)
-#12948 := (not #12945)
-decl f137 :: Int
-#585 := f137
-#12955 := (+ f137 #12901)
-#12954 := (>= #12955 0::Int)
-#12951 := (>= f462 0::Int)
-#12958 := (and #12951 #12954)
-#12961 := (not #12958)
-decl f461 :: Int
-#4747 := f461
-#4749 := (f20 #4748 #4736)
-#4780 := (= #4749 f461)
-#12418 := (<= f443 0::Int)
-#12419 := (not #12418)
-#12964 := (and #12419 #4780)
-#12967 := (not #12964)
-#13033 := (or #12967 #12961 #12948 #12938 #12431 #12926 #12920 #12886 #12874 #11409 #11970 #11961 #11952 #11943 #11934 #12850)
-#13041 := (and #12419 #4780 #13033)
-#12398 := (* -1::Int #4775)
-#12399 := (+ f461 #12398)
-#12397 := (>= #12399 0::Int)
-#12385 := (>= #243 1::Int)
-#12407 := (or #6637 #12385 #12397)
-#12412 := (forall (vars (?v0 Int)) #12407)
-#12415 := (not #12412)
-#13046 := (or #12415 #13041)
-#13049 := (and #12412 #13046)
-#12379 := (>= f443 1::Int)
-#12382 := (not #12379)
-#13052 := (or #12382 #13049)
-#13055 := (and #12379 #13052)
-decl f432 :: S3
-#3713 := f432
-#4761 := (f450 #4696 f432)
-#4762 := (f449 #4761 f431)
-#4763 := (f8 #4762 1::Int)
-#4764 := (f7 #4763 f10)
-#4765 := (= #4764 f1)
-#12150 := (not #4765)
-decl f434 :: S3
-#3719 := f434
-#4756 := (f450 #4696 f434)
-#4757 := (f449 #4756 f433)
-#4758 := (f8 #4757 0::Int)
-#4759 := (f7 #4758 f10)
-#4760 := (= #4759 f1)
-#12159 := (not #4760)
-decl f436 :: S3
-#3725 := f436
-#4751 := (f450 #4696 f436)
-#4752 := (f449 #4751 f435)
-#4753 := (f8 #4752 f461)
-#4754 := (f7 #4753 f14)
-#4755 := (= #4754 f1)
-#12168 := (not #4755)
-#4750 := (= f461 #4749)
-#12177 := (not #4750)
-#4744 := (f37 #4743 #4736)
-#4745 := (= #4744 f1)
-#4737 := (f85 f90 #4736)
-#4738 := (f7 #4737 f14)
-#4739 := (= #4738 f1)
-#4746 := (and #4739 #4745)
-#12186 := (not #4746)
-#13076 := (or #12186 #12177 #12168 #12159 #12150 #13055)
-#13084 := (and #4739 #4745 #13076)
-#4740 := (f37 #4667 #4736)
-#4741 := (= #4740 f1)
-#4742 := (and #4739 #4741)
-#12198 := (not #4742)
-#13089 := (or #12198 #13084)
-#13095 := (and #4739 #4741 #13089)
-#12210 := (not #4733)
-#13100 := (or #12210 #13095)
-#13103 := (and #4733 #13100)
-#12353 := (+ f135 #12352)
-#12351 := (>= #12353 0::Int)
-#12349 := (>= f443 0::Int)
-#12356 := (and #12349 #12351)
-#12359 := (not #12356)
-decl f458 :: (-> S184 Int S27)
-decl f457 :: Int
-#4715 := f457
-decl f459 :: S184
-#4718 := f459
-#4719 := (f458 f459 f457)
-#4720 := (f37 #4719 #40)
-#4721 := (pattern #4720)
-#4722 := (= #4720 f1)
-#11242 := (not #4722)
-#11245 := (forall (vars (?v0 S10)) (:pat #4721) #11242)
-#12231 := (not #11245)
-decl f266 :: (-> S108 S9 Int)
-decl f267 :: S108
-#2247 := f267
-#4716 := (f266 f267 f444)
-#4717 := (= f457 #4716)
-#12240 := (not #4717)
-decl f439 :: S3
-#3734 := f439
-#4697 := (f450 #4696 f439)
-#4711 := (f449 #4697 f437)
-#4712 := (f8 #4711 f443)
-#4713 := (f7 #4712 f10)
-#4714 := (= #4713 f1)
-#12249 := (not #4714)
-#4705 := (f454 #4704 f439)
-#4706 := (f453 #4705 f438)
-#4707 := (f85 #4706 #4656)
-#4708 := (f7 #4707 #21)
-#4709 := (= #4708 f1)
-#4698 := (f449 #4697 f438)
-#4700 := (f8 #4698 #4699)
-#4701 := (f7 #4700 #21)
-#4702 := (= #4701 f1)
-#4710 := (and #4702 #4709)
-#12258 := (not #4710)
-decl f446 :: (-> S177 S176 Int)
-#4689 := (:var 0 S176)
-decl f447 :: S177
-#4688 := f447
-#4690 := (f446 f447 #4689)
-#4691 := (pattern #4690)
-decl f448 :: Int
-#4692 := f448
-#13108 := (* -1::Int f448)
-#13109 := (+ #4690 #13108)
-#13107 := (>= #13109 0::Int)
-#13106 := (not #13107)
-#13112 := (forall (vars (?v0 S176)) (:pat #4691) #13106)
-#13115 := (not #13112)
-#4682 := (f322 f323 f439)
-#4683 := (f82 #4682 f444)
-#4684 := (= #4683 f1)
-#4687 := (and #4684 #4686)
-#12276 := (not #4687)
-decl f324 :: S50
-#2900 := f324
-#4680 := (f82 f324 f444)
-#4681 := (= #4680 f1)
-#12285 := (not #4681)
-#13121 := (not #13118)
-#1300 := 1099511627776::Int
-#13131 := (>= f443 1099511627776::Int)
-decl f442 :: Int
-#4642 := f442
-#13146 := (* -1::Int f442)
-#13147 := (+ f135 #13146)
-#13145 := (>= #13147 0::Int)
-#13143 := (>= f442 0::Int)
-#13150 := (and #13143 #13145)
-#13153 := (not #13150)
-decl f441 :: Int
-#4638 := f441
-#13160 := (* -1::Int f441)
-#13161 := (+ f135 #13160)
-#13159 := (>= #13161 0::Int)
-#13157 := (>= f441 0::Int)
-#13164 := (and #13157 #13159)
-#13167 := (not #13164)
-decl f440 :: Int
-#4634 := f440
-#13174 := (* -1::Int f440)
-#13175 := (+ f137 #13174)
-#13173 := (>= #13175 0::Int)
-#13171 := (>= f440 0::Int)
-#13178 := (and #13171 #13173)
-#13181 := (not #13178)
-#13226 := (or #13181 #13167 #13153 #13131 #12418 #13121 #12285 #12276 #13115 #12258 #12249 #12240 #12231 #12359 #13103)
-#13231 := (not #13226)
-#1 := true
-#4821 := (< #243 f443)
-#4827 := (and #4821 #4826)
-#567 := (<= #243 f135)
-#4828 := (and #567 #4827)
-#409 := (<= 0::Int #243)
-#4829 := (and #409 #4828)
-#4830 := (exists (vars (?v0 Int)) #4829)
-#4831 := (and #4830 true)
-#4822 := (<= #4775 f468)
-#4823 := (implies #4821 #4822)
-#568 := (and #409 #567)
-#4824 := (implies #568 #4823)
-#4825 := (forall (vars (?v0 Int)) #4824)
-#4832 := (implies #4825 #4831)
-#4833 := (and #4825 #4832)
-#4834 := (implies #4820 #4833)
-#4835 := (implies #4818 #4834)
-#4836 := (implies #4816 #4835)
-#4837 := (implies #4814 #4836)
-#4787 := (<= 0::Int f463)
-#4794 := (<= 1::Int f464)
-#4795 := (and #4794 #4787)
-#4838 := (implies #4795 #4837)
-#4839 := (implies #4795 #4838)
-#4840 := (implies #4795 #4839)
-#4841 := (implies #3125 #4840)
-#4842 := (and #3125 #4841)
-#4843 := (implies #4795 #4842)
-#4844 := (implies #4795 #4843)
-#5032 := (implies #4795 #4844)
-#5033 := (implies #4795 #5032)
-#5031 := (<= f443 f464)
-#5034 := (implies #5031 #5033)
-#5035 := (implies #4795 #5034)
-#4980 := (< f471 f443)
-#4985 := (and #4980 #4984)
-#4986 := (and #4985 true)
-#4976 := (<= #4775 f470)
-#4975 := (< #243 f472)
-#4977 := (implies #4975 #4976)
-#4978 := (implies #568 #4977)
-#4979 := (forall (vars (?v0 Int)) #4978)
-#4987 := (implies #4979 #4986)
-#4988 := (and #4979 #4987)
-#4974 := (<= f472 f443)
-#4989 := (implies #4974 #4988)
-#4990 := (and #4974 #4989)
-#4959 := (<= 0::Int f471)
-#4972 := (<= 2::Int f472)
-#4973 := (and #4972 #4959)
-#4991 := (implies #4973 #4990)
-#4992 := (implies #4971 #4991)
-#4961 := (+ f464 1::Int)
-#4966 := (= f472 #4961)
-#4993 := (implies #4966 #4992)
-#4963 := (<= #4961 f135)
-#4962 := (<= 0::Int #4961)
-#4964 := (and #4962 #4963)
-#4994 := (implies #4964 #4993)
-#4995 := (and #4964 #4994)
-#4960 := (and #4794 #4959)
-#4996 := (implies #4960 #4995)
-#5014 := (implies #5013 #4996)
-#5015 := (implies #5012 #5014)
-#5016 := (implies #4795 #5015)
-#5017 := (implies #4795 #5016)
-#5018 := (implies #4795 #5017)
-#5011 := (<= #4940 f462)
-#5019 := (implies #5011 #5018)
-#5020 := (implies #4795 #5019)
-#4997 := (implies #4958 #4996)
-#4998 := (implies #4956 #4997)
-#4954 := (and #4794 #4794)
-#4999 := (implies #4954 #4998)
-#5000 := (implies #4953 #4999)
-#5001 := (implies #4948 #5000)
-#5002 := (implies #4943 #5001)
-#5003 := (implies #4939 #5002)
-#5004 := (and #4939 #5003)
-#5005 := (implies #4936 #5004)
-#5006 := (and #4936 #5005)
-#5007 := (implies #4795 #5006)
-#5008 := (implies #4795 #5007)
-#4941 := (< f462 #4940)
-#5009 := (implies #4941 #5008)
-#5010 := (implies #4795 #5009)
-#5021 := (and #5010 #5020)
-#5022 := (implies #4795 #5021)
-#5023 := (implies #4939 #5022)
-#5024 := (and #4939 #5023)
-#5025 := (implies #4936 #5024)
-#5026 := (and #4936 #5025)
-#5027 := (implies #4795 #5026)
-#5028 := (implies #4795 #5027)
-#4928 := (< f464 f443)
-#5029 := (implies #4928 #5028)
-#5030 := (implies #4795 #5029)
-#5036 := (and #5030 #5035)
-#5037 := (implies #4795 #5036)
-decl f77 :: (-> S48 S9 S47)
-decl f78 :: S48
-#174 := f78
-#4850 := (f77 f78 f444)
-#4926 := (= #4850 #4850)
-#4925 := (= #4875 #4875)
-#4927 := (and #4925 #4926)
-#5038 := (implies #4927 #5037)
-#5039 := (implies #4924 #5038)
-#5040 := (implies #4914 #5039)
-#5041 := (implies #4910 #5040)
-#5042 := (implies #4906 #5041)
-#5043 := (implies #4902 #5042)
-#5044 := (implies #4812 #5043)
-decl f265 :: S42
-#2244 := f265
-#4884 := (f70 f265 f444)
-#4885 := (f20 #4884 #40)
-#4886 := (pattern #4885)
-#4887 := (<= #4885 #4885)
-#4888 := (forall (vars (?v0 S10)) (:pat #4886) #4887)
-#4892 := (and #4888 #4891)
-#4883 := (<= #4716 #4716)
-#4893 := (and #4883 #4892)
-#5045 := (implies #4893 #5044)
-#4877 := (f101 #4876 #40)
-#4878 := (pattern #4877)
-#4865 := (f37 #4743 #40)
-#4866 := (= #4865 f1)
-#4879 := (= #4877 #4877)
-#4880 := (and #4879 #4866)
-#4881 := (implies #4866 #4880)
-#4882 := (forall (vars (?v0 S10)) (:pat #4878) #4881)
-#4894 := (and #4882 #4893)
-decl f74 :: (-> S45 S10 S44)
-decl f75 :: (-> S46 S47 S45)
-decl f76 :: S46
-#173 := f76
-#4851 := (f75 f76 #4850)
-#4852 := (f74 #4851 #40)
-#4853 := (pattern #4852)
-#4871 := (= #4852 #4852)
-#4872 := (and #4871 #4866)
-#4873 := (implies #4866 #4872)
-#4874 := (forall (vars (?v0 S10)) (:pat #4853) #4873)
-#4895 := (and #4874 #4894)
-decl f39 :: (-> S28 S29 S16)
-decl f41 :: (-> S30 S9 S29)
-decl f42 :: S30
-#84 := f42
-#4861 := (f41 f42 f444)
-decl f40 :: S28
-#83 := f40
-#4862 := (f39 f40 #4861)
-#4863 := (f20 #4862 #40)
-#4864 := (pattern #4863)
-#4867 := (= #4863 #4863)
-#4868 := (and #4867 #4866)
-#4869 := (implies #4866 #4868)
-#4870 := (forall (vars (?v0 S10)) (:pat #4864) #4869)
-#4896 := (and #4870 #4895)
-decl f419 :: S21
-#3592 := f419
-#4854 := (f48 #4661 #40)
-#4855 := (f27 f28 #4854)
-#4856 := (f26 #4855)
-#4857 := (= #4856 f419)
-#4858 := (not #4857)
-#4859 := (implies #4858 #4858)
-#4860 := (forall (vars (?v0 S10)) (:pat #4853) #4859)
-#4897 := (and #4860 #4896)
-#5046 := (implies #4897 #5045)
-#5047 := (implies #4795 #5046)
-#5048 := (implies #4795 #5047)
-#5049 := (implies #4795 #5048)
-#4845 := (implies #4812 #4844)
-#4846 := (implies #4795 #4845)
-#4847 := (implies #4795 #4846)
-#4808 := (not true)
-#4848 := (implies #4808 #4847)
-#4849 := (implies #4795 #4848)
-#5050 := (and #4849 #5049)
-#5051 := (implies #4795 #5050)
-#4802 := (< f463 f443)
-#4807 := (and #4802 #4806)
-#5052 := (implies #4807 #5051)
-#4798 := (<= #4775 f462)
-#4797 := (< #243 f464)
-#4799 := (implies #4797 #4798)
-#4800 := (implies #568 #4799)
-#4801 := (forall (vars (?v0 Int)) #4800)
-#5053 := (implies #4801 #5052)
-#4796 := (<= f464 f443)
-#5054 := (implies #4796 #5053)
-#5055 := (implies #4795 #5054)
-#4792 := (<= f464 f135)
-#4791 := (<= 0::Int f464)
-#4793 := (and #4791 #4792)
-#5056 := (implies #4793 #5055)
-#4788 := (<= f463 f135)
-#4789 := (and #4787 #4788)
-#5057 := (implies #4789 #5056)
-#4784 := (<= f462 f137)
-#4783 := (<= 0::Int f462)
-#4785 := (and #4783 #4784)
-#5058 := (implies #4785 #5057)
-#4648 := (< 0::Int f443)
-#4781 := (and #4648 #4780)
-#5059 := (implies #4781 #5058)
-#5060 := (and #4781 #5059)
-#4776 := (<= #4775 f461)
-#4772 := (< #243 1::Int)
-#4777 := (implies #4772 #4776)
-#4778 := (implies #568 #4777)
-#4779 := (forall (vars (?v0 Int)) #4778)
-#5061 := (implies #4779 #5060)
-#5062 := (and #4779 #5061)
-#4771 := (<= 1::Int f443)
-#5063 := (implies #4771 #5062)
-#5064 := (and #4771 #5063)
-#4767 := (<= 0::Int 0::Int)
-#4768 := (and #4767 #4767)
-#4766 := (<= 1::Int 1::Int)
-#4769 := (and #4766 #4768)
-#4770 := (and #4766 #4769)
-#5065 := (implies #4770 #5064)
-#5066 := (implies #4765 #5065)
-#5067 := (implies #4760 #5066)
-#5068 := (implies #4755 #5067)
-#5069 := (implies #4750 #5068)
-#5070 := (implies #4746 #5069)
-#5071 := (and #4746 #5070)
-#5072 := (implies #4742 #5071)
-#5073 := (and #4742 #5072)
-#5074 := (implies #4733 #5073)
-#5075 := (and #4733 #5074)
-#4726 := (<= f443 f135)
-#4725 := (<= 0::Int f443)
-#4727 := (and #4725 #4726)
-#5076 := (implies #4727 #5075)
-#4723 := (iff #4722 false)
-#4724 := (forall (vars (?v0 S10)) (:pat #4721) #4723)
-#5077 := (implies #4724 #5076)
-#5078 := (implies #4717 #5077)
-#5079 := (implies #4714 #5078)
-#5080 := (implies #4710 #5079)
-#4693 := (< #4690 f448)
-#4694 := (forall (vars (?v0 S176)) (:pat #4691) #4693)
-#5081 := (implies #4694 #5080)
-#5082 := (implies #4687 #5081)
-#5083 := (implies #4681 #5082)
-#4675 := (and #4672 #4674)
-#4676 := (and #4669 #4675)
-#4677 := (and #4666 #4676)
-#4678 := (and #4663 #4677)
-#4679 := (and #4660 #4678)
-#5084 := (implies #4679 #5083)
-#5085 := (implies #4648 #5084)
-#4647 := (< f443 1099511627776::Int)
-#5086 := (implies #4647 #5085)
-#4644 := (<= f442 f135)
-#4643 := (<= 0::Int f442)
-#4645 := (and #4643 #4644)
-#5087 := (implies #4645 #5086)
-#4640 := (<= f441 f135)
-#4639 := (<= 0::Int f441)
-#4641 := (and #4639 #4640)
-#5088 := (implies #4641 #5087)
-#4636 := (<= f440 f137)
-#4635 := (<= 0::Int f440)
-#4637 := (and #4635 #4636)
-#5089 := (implies #4637 #5088)
-#5090 := (not #5089)
-#13234 := (iff #5090 #13231)
-#11288 := (not #4821)
-#11289 := (or #11288 #4822)
-#6630 := (not #568)
-#11295 := (or #6630 #11289)
-#11300 := (forall (vars (?v0 Int)) #11295)
-#11308 := (not #11300)
-#11309 := (or #11308 #4830)
-#11314 := (and #11300 #11309)
-#11321 := (or #11320 #11314)
-#11330 := (or #11329 #11321)
-#11339 := (or #11338 #11330)
-#11348 := (or #11347 #11339)
-#11356 := (not #4795)
-#11357 := (or #11356 #11348)
-#11365 := (or #11356 #11357)
-#11373 := (or #11356 #11365)
-#11382 := (or #11381 #11373)
-#11387 := (and #3125 #11382)
-#11393 := (or #11356 #11387)
-#11401 := (or #11356 #11393)
-#11883 := (or #11356 #11401)
-#11891 := (or #11356 #11883)
-#11899 := (not #5031)
-#11900 := (or #11899 #11891)
-#11908 := (or #11356 #11900)
-#11555 := (not #4975)
-#11556 := (or #11555 #4976)
-#11562 := (or #6630 #11556)
-#11567 := (forall (vars (?v0 Int)) #11562)
-#11575 := (not #11567)
-#11576 := (or #11575 #4985)
-#11581 := (and #11567 #11576)
-#11587 := (not #4974)
-#11588 := (or #11587 #11581)
-#11593 := (and #4974 #11588)
-#11599 := (not #4973)
-#11600 := (or #11599 #11593)
-#11609 := (or #11608 #11600)
-#11540 := (+ 1::Int f464)
-#11552 := (= f472 #11540)
-#11617 := (not #11552)
-#11618 := (or #11617 #11609)
-#11546 := (<= #11540 f135)
-#11543 := (<= 0::Int #11540)
-#11549 := (and #11543 #11546)
-#11626 := (not #11549)
-#11627 := (or #11626 #11618)
-#11632 := (and #11549 #11627)
-#11638 := (not #4960)
-#11639 := (or #11638 #11632)
-#11759 := (or #11758 #11639)
-#11768 := (or #11767 #11759)
-#11776 := (or #11356 #11768)
-#11784 := (or #11356 #11776)
-#11792 := (or #11356 #11784)
-#11800 := (not #5011)
-#11801 := (or #11800 #11792)
-#11809 := (or #11356 #11801)
-#11648 := (or #11647 #11639)
-#11657 := (or #11656 #11648)
-#11665 := (not #4794)
-#11666 := (or #11665 #11657)
-#11675 := (or #11674 #11666)
-#11684 := (or #11683 #11675)
-#11693 := (or #11692 #11684)
-#11702 := (or #11701 #11693)
-#11707 := (and #4939 #11702)
-#11714 := (or #11713 #11707)
-#11719 := (and #4936 #11714)
-#11725 := (or #11356 #11719)
-#11733 := (or #11356 #11725)
-#11741 := (not #4941)
-#11742 := (or #11741 #11733)
-#11750 := (or #11356 #11742)
-#11814 := (and #11750 #11809)
-#11820 := (or #11356 #11814)
-#11828 := (or #11701 #11820)
-#11833 := (and #4939 #11828)
-#11839 := (or #11713 #11833)
-#11844 := (and #4936 #11839)
-#11850 := (or #11356 #11844)
-#11858 := (or #11356 #11850)
-#11866 := (not #4928)
-#11867 := (or #11866 #11858)
-#11875 := (or #11356 #11867)
-#11913 := (and #11875 #11908)
-#11919 := (or #11356 #11913)
-#11935 := (or #11934 #11919)
-#11944 := (or #11943 #11935)
-#11953 := (or #11952 #11944)
-#11962 := (or #11961 #11953)
-#11971 := (or #11970 #11962)
-#11979 := (or #11409 #11971)
-#11987 := (not #4893)
-#11988 := (or #11987 #11979)
-#11996 := (or #11987 #11988)
-#12004 := (or #11356 #11996)
-#12012 := (or #11356 #12004)
-#12020 := (or #11356 #12012)
-#12035 := (or #11356 #12020)
-#12043 := (not #4807)
-#12044 := (or #12043 #12035)
-#11271 := (not #4797)
-#11272 := (or #11271 #4798)
-#11278 := (or #6630 #11272)
-#11283 := (forall (vars (?v0 Int)) #11278)
-#12052 := (not #11283)
-#12053 := (or #12052 #12044)
-#12061 := (not #4796)
-#12062 := (or #12061 #12053)
-#12070 := (or #11356 #12062)
-#12078 := (not #4793)
-#12079 := (or #12078 #12070)
-#12087 := (not #4789)
-#12088 := (or #12087 #12079)
-#12096 := (not #4785)
-#12097 := (or #12096 #12088)
-#12105 := (not #4781)
-#12106 := (or #12105 #12097)
-#12111 := (and #4781 #12106)
-#11256 := (not #4772)
-#11257 := (or #11256 #4776)
-#11263 := (or #6630 #11257)
-#11268 := (forall (vars (?v0 Int)) #11263)
-#12117 := (not #11268)
-#12118 := (or #12117 #12111)
-#12123 := (and #11268 #12118)
-#12129 := (not #4771)
-#12130 := (or #12129 #12123)
-#12135 := (and #4771 #12130)
-#11250 := (and #4766 #4767)
-#11253 := (and #4766 #11250)
-#12141 := (not #11253)
-#12142 := (or #12141 #12135)
-#12151 := (or #12150 #12142)
-#12160 := (or #12159 #12151)
-#12169 := (or #12168 #12160)
-#12178 := (or #12177 #12169)
-#12187 := (or #12186 #12178)
-#12192 := (and #4746 #12187)
-#12199 := (or #12198 #12192)
-#12204 := (and #4742 #12199)
-#12211 := (or #12210 #12204)
-#12216 := (and #4733 #12211)
-#12222 := (not #4727)
-#12223 := (or #12222 #12216)
-#12232 := (or #12231 #12223)
-#12241 := (or #12240 #12232)
-#12250 := (or #12249 #12241)
-#12259 := (or #12258 #12250)
-#12267 := (not #4694)
-#12268 := (or #12267 #12259)
-#12277 := (or #12276 #12268)
-#12286 := (or #12285 #12277)
-#12294 := (not #4679)
-#12295 := (or #12294 #12286)
-#12303 := (not #4648)
-#12304 := (or #12303 #12295)
-#12312 := (not #4647)
-#12313 := (or #12312 #12304)
-#12321 := (not #4645)
-#12322 := (or #12321 #12313)
-#12330 := (not #4641)
-#12331 := (or #12330 #12322)
-#12339 := (not #4637)
-#12340 := (or #12339 #12331)
-#12345 := (not #12340)
-#13232 := (iff #12345 #13231)
-#13229 := (iff #12340 #13226)
-#13184 := (or #12359 #13103)
-#13187 := (or #12231 #13184)
-#13190 := (or #12240 #13187)
-#13193 := (or #12249 #13190)
-#13196 := (or #12258 #13193)
-#13199 := (or #13115 #13196)
-#13202 := (or #12276 #13199)
-#13205 := (or #12285 #13202)
-#13208 := (or #13121 #13205)
-#13211 := (or #12418 #13208)
-#13214 := (or #13131 #13211)
-#13217 := (or #13153 #13214)
-#13220 := (or #13167 #13217)
-#13223 := (or #13181 #13220)
-#13227 := (iff #13223 #13226)
-#13228 := [rewrite]: #13227
-#13224 := (iff #12340 #13223)
-#13221 := (iff #12331 #13220)
-#13218 := (iff #12322 #13217)
-#13215 := (iff #12313 #13214)
-#13212 := (iff #12304 #13211)
-#13209 := (iff #12295 #13208)
-#13206 := (iff #12286 #13205)
-#13203 := (iff #12277 #13202)
-#13200 := (iff #12268 #13199)
-#13197 := (iff #12259 #13196)
-#13194 := (iff #12250 #13193)
-#13191 := (iff #12241 #13190)
-#13188 := (iff #12232 #13187)
-#13185 := (iff #12223 #13184)
-#13104 := (iff #12216 #13103)
-#13101 := (iff #12211 #13100)
-#13098 := (iff #12204 #13095)
-#13092 := (and #4742 #13089)
-#13096 := (iff #13092 #13095)
-#13097 := [rewrite]: #13096
-#13093 := (iff #12204 #13092)
-#13090 := (iff #12199 #13089)
-#13087 := (iff #12192 #13084)
-#13081 := (and #4746 #13076)
-#13085 := (iff #13081 #13084)
-#13086 := [rewrite]: #13085
-#13082 := (iff #12192 #13081)
-#13079 := (iff #12187 #13076)
-#13058 := (or false #13055)
-#13061 := (or #12150 #13058)
-#13064 := (or #12159 #13061)
-#13067 := (or #12168 #13064)
-#13070 := (or #12177 #13067)
-#13073 := (or #12186 #13070)
-#13077 := (iff #13073 #13076)
-#13078 := [rewrite]: #13077
-#13074 := (iff #12187 #13073)
-#13071 := (iff #12178 #13070)
-#13068 := (iff #12169 #13067)
-#13065 := (iff #12160 #13064)
-#13062 := (iff #12151 #13061)
-#13059 := (iff #12142 #13058)
-#13056 := (iff #12135 #13055)
-#13053 := (iff #12130 #13052)
-#13050 := (iff #12123 #13049)
-#13047 := (iff #12118 #13046)
-#13044 := (iff #12111 #13041)
-#13038 := (and #12964 #13033)
-#13042 := (iff #13038 #13041)
-#13043 := [rewrite]: #13042
-#13039 := (iff #12111 #13038)
-#13036 := (iff #12106 #13033)
-#12970 := (or #12431 #12850)
-#12973 := (or #11934 #12970)
-#12976 := (or #11943 #12973)
-#12979 := (or #11952 #12976)
-#12982 := (or #11961 #12979)
-#12985 := (or #11970 #12982)
-#12988 := (or #11409 #12985)
-#12991 := (or #12874 #12988)
-#12994 := (or #12874 #12991)
-#12997 := (or #12431 #12994)
-#13000 := (or #12431 #12997)
-#13003 := (or #12431 #13000)
-#13006 := (or #12431 #13003)
-#13009 := (or #12886 #13006)
-#13012 := (or #12920 #13009)
-#13015 := (or #12926 #13012)
-#13018 := (or #12431 #13015)
-#13021 := (or #12938 #13018)
-#13024 := (or #12948 #13021)
-#13027 := (or #12961 #13024)
-#13030 := (or #12967 #13027)
-#13034 := (iff #13030 #13033)
-#13035 := [rewrite]: #13034
-#13031 := (iff #12106 #13030)
-#13028 := (iff #12097 #13027)
-#13025 := (iff #12088 #13024)
-#13022 := (iff #12079 #13021)
-#13019 := (iff #12070 #13018)
-#13016 := (iff #12062 #13015)
-#13013 := (iff #12053 #13012)
-#13010 := (iff #12044 #13009)
-#13007 := (iff #12035 #13006)
-#13004 := (iff #12020 #13003)
-#13001 := (iff #12012 #13000)
-#12998 := (iff #12004 #12997)
-#12995 := (iff #11996 #12994)
-#12992 := (iff #11988 #12991)
-#12989 := (iff #11979 #12988)
-#12986 := (iff #11971 #12985)
-#12983 := (iff #11962 #12982)
-#12980 := (iff #11953 #12979)
-#12977 := (iff #11944 #12976)
-#12974 := (iff #11935 #12973)
-#12971 := (iff #11919 #12970)
-#12851 := (iff #11913 #12850)
-#12848 := (iff #11908 #12845)
-#12827 := (or #12431 #12820)
-#12830 := (or #12431 #12827)
-#12833 := (or #12431 #12830)
-#12836 := (or #12431 #12833)
-#12839 := (or #12707 #12836)
-#12842 := (or #12431 #12839)
-#12846 := (iff #12842 #12845)
-#12847 := [rewrite]: #12846
-#12843 := (iff #11908 #12842)
-#12840 := (iff #11900 #12839)
-#12837 := (iff #11891 #12836)
-#12834 := (iff #11883 #12833)
-#12831 := (iff #11401 #12830)
-#12828 := (iff #11393 #12827)
-#12821 := (iff #11387 #12820)
-#12818 := (iff #11382 #12815)
-#12791 := (or #11320 #12788)
-#12794 := (or #11329 #12791)
-#12797 := (or #11338 #12794)
-#12800 := (or #11347 #12797)
-#12803 := (or #12431 #12800)
-#12806 := (or #12431 #12803)
-#12809 := (or #12431 #12806)
-#12812 := (or #11381 #12809)
-#12816 := (iff #12812 #12815)
-#12817 := [rewrite]: #12816
-#12813 := (iff #11382 #12812)
-#12810 := (iff #11373 #12809)
-#12807 := (iff #11365 #12806)
-#12804 := (iff #11357 #12803)
-#12801 := (iff #11348 #12800)
-#12798 := (iff #11339 #12797)
-#12795 := (iff #11330 #12794)
-#12792 := (iff #11321 #12791)
-#12789 := (iff #11314 #12788)
-#12786 := (iff #11309 #12785)
-#12783 := (iff #4830 #12782)
-#12780 := (iff #4829 #12777)
-#12768 := (and #12736 #4826)
-#12771 := (and #5617 #12768)
-#12774 := (and #5433 #12771)
-#12778 := (iff #12774 #12777)
-#12779 := [rewrite]: #12778
-#12775 := (iff #4829 #12774)
-#12772 := (iff #4828 #12771)
-#12769 := (iff #4827 #12768)
-#12737 := (iff #4821 #12736)
-#12738 := [rewrite]: #12737
-#12770 := [monotonicity #12738]: #12769
-#5618 := (iff #567 #5617)
-#5619 := [rewrite]: #5618
-#12773 := [monotonicity #5619 #12770]: #12772
-#5431 := (iff #409 #5433)
-#5432 := [rewrite]: #5431
-#12776 := [monotonicity #5432 #12773]: #12775
-#12781 := [trans #12776 #12779]: #12780
-#12784 := [quant-intro #12781]: #12783
-#12766 := (iff #11308 #12765)
-#12763 := (iff #11300 #12762)
-#12760 := (iff #11295 #12757)
-#12751 := (or #12734 #12748)
-#12754 := (or #6637 #12751)
-#12758 := (iff #12754 #12757)
-#12759 := [rewrite]: #12758
-#12755 := (iff #11295 #12754)
-#12752 := (iff #11289 #12751)
-#12749 := (iff #4822 #12748)
-#12750 := [rewrite]: #12749
-#12744 := (iff #11288 #12734)
-#12739 := (not #12736)
-#12742 := (iff #12739 #12734)
-#12743 := [rewrite]: #12742
-#12740 := (iff #11288 #12739)
-#12741 := [monotonicity #12738]: #12740
-#12745 := [trans #12741 #12743]: #12744
-#12753 := [monotonicity #12745 #12750]: #12752
-#6638 := (iff #6630 #6637)
-#5625 := (iff #568 #5624)
-#5626 := [monotonicity #5432 #5619]: #5625
-#6639 := [monotonicity #5626]: #6638
-#12756 := [monotonicity #6639 #12753]: #12755
-#12761 := [trans #12756 #12759]: #12760
-#12764 := [quant-intro #12761]: #12763
-#12767 := [monotonicity #12764]: #12766
-#12787 := [monotonicity #12767 #12784]: #12786
-#12790 := [monotonicity #12764 #12787]: #12789
-#12793 := [monotonicity #12790]: #12792
-#12796 := [monotonicity #12793]: #12795
-#12799 := [monotonicity #12796]: #12798
-#12802 := [monotonicity #12799]: #12801
-#12432 := (iff #11356 #12431)
-#12429 := (iff #4795 #12428)
-#12425 := (iff #4787 #12426)
-#12427 := [rewrite]: #12425
-#12422 := (iff #4794 #12423)
-#12424 := [rewrite]: #12422
-#12430 := [monotonicity #12424 #12427]: #12429
-#12433 := [monotonicity #12430]: #12432
-#12805 := [monotonicity #12433 #12802]: #12804
-#12808 := [monotonicity #12433 #12805]: #12807
-#12811 := [monotonicity #12433 #12808]: #12810
-#12814 := [monotonicity #12811]: #12813
-#12819 := [trans #12814 #12817]: #12818
-#12822 := [monotonicity #12819]: #12821
-#12829 := [monotonicity #12433 #12822]: #12828
-#12832 := [monotonicity #12433 #12829]: #12831
-#12835 := [monotonicity #12433 #12832]: #12834
-#12838 := [monotonicity #12433 #12835]: #12837
-#12825 := (iff #11899 #12707)
-#12823 := (iff #5031 #12706)
-#12824 := [rewrite]: #12823
-#12826 := [monotonicity #12824]: #12825
-#12841 := [monotonicity #12826 #12838]: #12840
-#12844 := [monotonicity #12433 #12841]: #12843
-#12849 := [trans #12844 #12847]: #12848
-#12732 := (iff #11875 #12729)
-#12717 := (or #12431 #12700)
-#12720 := (or #12431 #12717)
-#12723 := (or #12706 #12720)
-#12726 := (or #12431 #12723)
-#12730 := (iff #12726 #12729)
-#12731 := [rewrite]: #12730
-#12727 := (iff #11875 #12726)
-#12724 := (iff #11867 #12723)
-#12721 := (iff #11858 #12720)
-#12718 := (iff #11850 #12717)
-#12703 := (iff #11844 #12700)
-#12697 := (and #4936 #12694)
-#12701 := (iff #12697 #12700)
-#12702 := [rewrite]: #12701
-#12698 := (iff #11844 #12697)
-#12695 := (iff #11839 #12694)
-#12692 := (iff #11833 #12689)
-#12686 := (and #4939 #12681)
-#12690 := (iff #12686 #12689)
-#12691 := [rewrite]: #12690
-#12687 := (iff #11833 #12686)
-#12684 := (iff #11828 #12681)
-#12675 := (or #12431 #12672)
-#12678 := (or #11701 #12675)
-#12682 := (iff #12678 #12681)
-#12683 := [rewrite]: #12682
-#12679 := (iff #11828 #12678)
-#12676 := (iff #11820 #12675)
-#12673 := (iff #11814 #12672)
-#12670 := (iff #11809 #12667)
-#12564 := (or #12440 #12556)
-#12646 := (or #11758 #12564)
-#12649 := (or #11767 #12646)
-#12652 := (or #12431 #12649)
-#12655 := (or #12431 #12652)
-#12658 := (or #12431 #12655)
-#12661 := (or #12612 #12658)
-#12664 := (or #12431 #12661)
-#12668 := (iff #12664 #12667)
-#12669 := [rewrite]: #12668
-#12665 := (iff #11809 #12664)
-#12662 := (iff #11801 #12661)
-#12659 := (iff #11792 #12658)
-#12656 := (iff #11784 #12655)
-#12653 := (iff #11776 #12652)
-#12650 := (iff #11768 #12649)
-#12647 := (iff #11759 #12646)
-#12565 := (iff #11639 #12564)
-#12559 := (iff #11632 #12556)
-#12553 := (and #12530 #12548)
-#12557 := (iff #12553 #12556)
-#12558 := [rewrite]: #12557
-#12554 := (iff #11632 #12553)
-#12551 := (iff #11627 #12548)
-#12536 := (or #12457 #12520)
-#12539 := (or #11608 #12536)
-#12542 := (or #12527 #12539)
-#12545 := (or #12533 #12542)
-#12549 := (iff #12545 #12548)
-#12550 := [rewrite]: #12549
-#12546 := (iff #11627 #12545)
-#12543 := (iff #11618 #12542)
-#12540 := (iff #11609 #12539)
-#12537 := (iff #11600 #12536)
-#12521 := (iff #11593 #12520)
-#12518 := (iff #11588 #12517)
-#12515 := (iff #11581 #12514)
-#12512 := (iff #11576 #12511)
-#12509 := (iff #4985 #12508)
-#12506 := (iff #4980 #12505)
-#12507 := [rewrite]: #12506
-#12510 := [monotonicity #12507]: #12509
-#12500 := (iff #11575 #12499)
-#12497 := (iff #11567 #12496)
-#12494 := (iff #11562 #12491)
-#12485 := (or #12468 #12482)
-#12488 := (or #6637 #12485)
-#12492 := (iff #12488 #12491)
-#12493 := [rewrite]: #12492
-#12489 := (iff #11562 #12488)
-#12486 := (iff #11556 #12485)
-#12483 := (iff #4976 #12482)
-#12484 := [rewrite]: #12483
-#12478 := (iff #11555 #12468)
-#12470 := (not #12468)
-#12473 := (not #12470)
-#12476 := (iff #12473 #12468)
-#12477 := [rewrite]: #12476
-#12474 := (iff #11555 #12473)
-#12471 := (iff #4975 #12470)
-#12472 := [rewrite]: #12471
-#12475 := [monotonicity #12472]: #12474
-#12479 := [trans #12475 #12477]: #12478
-#12487 := [monotonicity #12479 #12484]: #12486
-#12490 := [monotonicity #6639 #12487]: #12489
-#12495 := [trans #12490 #12493]: #12494
-#12498 := [quant-intro #12495]: #12497
-#12501 := [monotonicity #12498]: #12500
-#12513 := [monotonicity #12501 #12510]: #12512
-#12516 := [monotonicity #12498 #12513]: #12515
-#12466 := (iff #11587 #12465)
-#12463 := (iff #4974 #12460)
-#12464 := [rewrite]: #12463
-#12467 := [monotonicity #12464]: #12466
-#12519 := [monotonicity #12467 #12516]: #12518
-#12522 := [monotonicity #12464 #12519]: #12521
-#12458 := (iff #11599 #12457)
-#12455 := (iff #4973 #12454)
-#12434 := (iff #4959 #12435)
-#12436 := [rewrite]: #12434
-#12451 := (iff #4972 #12452)
-#12453 := [rewrite]: #12451
-#12456 := [monotonicity #12453 #12436]: #12455
-#12459 := [monotonicity #12456]: #12458
-#12538 := [monotonicity #12459 #12522]: #12537
-#12541 := [monotonicity #12538]: #12540
-#12528 := (iff #11617 #12527)
-#12525 := (iff #11552 #12523)
-#12526 := [rewrite]: #12525
-#12529 := [monotonicity #12526]: #12528
-#12544 := [monotonicity #12529 #12541]: #12543
-#12534 := (iff #11626 #12533)
-#12531 := (iff #11549 #12530)
-#12449 := (iff #11546 #12446)
-#12450 := [rewrite]: #12449
-#12443 := (iff #11543 #12444)
-#12445 := [rewrite]: #12443
-#12532 := [monotonicity #12445 #12450]: #12531
-#12535 := [monotonicity #12532]: #12534
-#12547 := [monotonicity #12535 #12544]: #12546
-#12552 := [trans #12547 #12550]: #12551
-#12555 := [monotonicity #12532 #12552]: #12554
-#12560 := [trans #12555 #12558]: #12559
-#12441 := (iff #11638 #12440)
-#12438 := (iff #4960 #12437)
-#12439 := [monotonicity #12424 #12436]: #12438
-#12442 := [monotonicity #12439]: #12441
-#12566 := [monotonicity #12442 #12560]: #12565
-#12648 := [monotonicity #12566]: #12647
-#12651 := [monotonicity #12648]: #12650
-#12654 := [monotonicity #12433 #12651]: #12653
-#12657 := [monotonicity #12433 #12654]: #12656
-#12660 := [monotonicity #12433 #12657]: #12659
-#12644 := (iff #11800 #12612)
-#12642 := (iff #5011 #12613)
-#12643 := [rewrite]: #12642
-#12645 := [monotonicity #12643]: #12644
-#12663 := [monotonicity #12645 #12660]: #12662
-#12666 := [monotonicity #12433 #12663]: #12665
-#12671 := [trans #12666 #12669]: #12670
-#12640 := (iff #11750 #12637)
-#12625 := (or #12431 #12607)
-#12628 := (or #12431 #12625)
-#12631 := (or #12613 #12628)
-#12634 := (or #12431 #12631)
-#12638 := (iff #12634 #12637)
-#12639 := [rewrite]: #12638
-#12635 := (iff #11750 #12634)
-#12632 := (iff #11742 #12631)
-#12629 := (iff #11733 #12628)
-#12626 := (iff #11725 #12625)
-#12610 := (iff #11719 #12607)
-#12604 := (and #4936 #12601)
-#12608 := (iff #12604 #12607)
-#12609 := [rewrite]: #12608
-#12605 := (iff #11719 #12604)
-#12602 := (iff #11714 #12601)
-#12599 := (iff #11707 #12596)
-#12593 := (and #4939 #12588)
-#12597 := (iff #12593 #12596)
-#12598 := [rewrite]: #12597
-#12594 := (iff #11707 #12593)
-#12591 := (iff #11702 #12588)
-#12567 := (or #11647 #12564)
-#12570 := (or #11656 #12567)
-#12573 := (or #12561 #12570)
-#12576 := (or #11674 #12573)
-#12579 := (or #11683 #12576)
-#12582 := (or #11692 #12579)
-#12585 := (or #11701 #12582)
-#12589 := (iff #12585 #12588)
-#12590 := [rewrite]: #12589
-#12586 := (iff #11702 #12585)
-#12583 := (iff #11693 #12582)
-#12580 := (iff #11684 #12579)
-#12577 := (iff #11675 #12576)
-#12574 := (iff #11666 #12573)
-#12571 := (iff #11657 #12570)
-#12568 := (iff #11648 #12567)
-#12569 := [monotonicity #12566]: #12568
-#12572 := [monotonicity #12569]: #12571
-#12562 := (iff #11665 #12561)
-#12563 := [monotonicity #12424]: #12562
-#12575 := [monotonicity #12563 #12572]: #12574
-#12578 := [monotonicity #12575]: #12577
-#12581 := [monotonicity #12578]: #12580
-#12584 := [monotonicity #12581]: #12583
-#12587 := [monotonicity #12584]: #12586
-#12592 := [trans #12587 #12590]: #12591
-#12595 := [monotonicity #12592]: #12594
-#12600 := [trans #12595 #12598]: #12599
-#12603 := [monotonicity #12600]: #12602
-#12606 := [monotonicity #12603]: #12605
-#12611 := [trans #12606 #12609]: #12610
-#12627 := [monotonicity #12433 #12611]: #12626
-#12630 := [monotonicity #12433 #12627]: #12629
-#12623 := (iff #11741 #12613)
-#12618 := (not #12612)
-#12621 := (iff #12618 #12613)
-#12622 := [rewrite]: #12621
-#12619 := (iff #11741 #12618)
-#12616 := (iff #4941 #12612)
-#12617 := [rewrite]: #12616
-#12620 := [monotonicity #12617]: #12619
-#12624 := [trans #12620 #12622]: #12623
-#12633 := [monotonicity #12624 #12630]: #12632
-#12636 := [monotonicity #12433 #12633]: #12635
-#12641 := [trans #12636 #12639]: #12640
-#12674 := [monotonicity #12641 #12671]: #12673
-#12677 := [monotonicity #12433 #12674]: #12676
-#12680 := [monotonicity #12677]: #12679
-#12685 := [trans #12680 #12683]: #12684
-#12688 := [monotonicity #12685]: #12687
-#12693 := [trans #12688 #12691]: #12692
-#12696 := [monotonicity #12693]: #12695
-#12699 := [monotonicity #12696]: #12698
-#12704 := [trans #12699 #12702]: #12703
-#12719 := [monotonicity #12433 #12704]: #12718
-#12722 := [monotonicity #12433 #12719]: #12721
-#12715 := (iff #11866 #12706)
-#12710 := (not #12707)
-#12713 := (iff #12710 #12706)
-#12714 := [rewrite]: #12713
-#12711 := (iff #11866 #12710)
-#12708 := (iff #4928 #12707)
-#12709 := [rewrite]: #12708
-#12712 := [monotonicity #12709]: #12711
-#12716 := [trans #12712 #12714]: #12715
-#12725 := [monotonicity #12716 #12722]: #12724
-#12728 := [monotonicity #12433 #12725]: #12727
-#12733 := [trans #12728 #12731]: #12732
-#12852 := [monotonicity #12733 #12849]: #12851
-#12972 := [monotonicity #12433 #12852]: #12971
-#12975 := [monotonicity #12972]: #12974
-#12978 := [monotonicity #12975]: #12977
-#12981 := [monotonicity #12978]: #12980
-#12984 := [monotonicity #12981]: #12983
-#12987 := [monotonicity #12984]: #12986
-#12990 := [monotonicity #12987]: #12989
-#12875 := (iff #11987 #12874)
-#12872 := (iff #4893 #4891)
-#12864 := (and true #4891)
-#12867 := (and true #12864)
-#12870 := (iff #12867 #4891)
-#12871 := [rewrite]: #12870
-#12868 := (iff #4893 #12867)
-#12865 := (iff #4892 #12864)
-#12860 := (iff #4888 true)
-#12855 := (forall (vars (?v0 S10)) (:pat #4886) true)
-#12858 := (iff #12855 true)
-#12859 := [elim-unused]: #12858
-#12856 := (iff #4888 #12855)
-#12853 := (iff #4887 true)
-#12854 := [rewrite]: #12853
-#12857 := [quant-intro #12854]: #12856
-#12861 := [trans #12857 #12859]: #12860
-#12866 := [monotonicity #12861]: #12865
-#12862 := (iff #4883 true)
-#12863 := [rewrite]: #12862
-#12869 := [monotonicity #12863 #12866]: #12868
-#12873 := [trans #12869 #12871]: #12872
-#12876 := [monotonicity #12873]: #12875
-#12993 := [monotonicity #12876 #12990]: #12992
-#12996 := [monotonicity #12876 #12993]: #12995
-#12999 := [monotonicity #12433 #12996]: #12998
-#13002 := [monotonicity #12433 #12999]: #13001
-#13005 := [monotonicity #12433 #13002]: #13004
-#13008 := [monotonicity #12433 #13005]: #13007
-#12887 := (iff #12043 #12886)
-#12884 := (iff #4807 #12883)
-#12881 := (iff #4802 #12880)
-#12882 := [rewrite]: #12881
-#12885 := [monotonicity #12882]: #12884
-#12888 := [monotonicity #12885]: #12887
-#13011 := [monotonicity #12888 #13008]: #13010
-#12921 := (iff #12052 #12920)
-#12918 := (iff #11283 #12917)
-#12915 := (iff #11278 #12912)
-#12906 := (or #12889 #12903)
-#12909 := (or #6637 #12906)
-#12913 := (iff #12909 #12912)
-#12914 := [rewrite]: #12913
-#12910 := (iff #11278 #12909)
-#12907 := (iff #11272 #12906)
-#12904 := (iff #4798 #12903)
-#12905 := [rewrite]: #12904
-#12899 := (iff #11271 #12889)
-#12891 := (not #12889)
-#12894 := (not #12891)
-#12897 := (iff #12894 #12889)
-#12898 := [rewrite]: #12897
-#12895 := (iff #11271 #12894)
-#12892 := (iff #4797 #12891)
-#12893 := [rewrite]: #12892
-#12896 := [monotonicity #12893]: #12895
-#12900 := [trans #12896 #12898]: #12899
-#12908 := [monotonicity #12900 #12905]: #12907
-#12911 := [monotonicity #6639 #12908]: #12910
-#12916 := [trans #12911 #12914]: #12915
-#12919 := [quant-intro #12916]: #12918
-#12922 := [monotonicity #12919]: #12921
-#13014 := [monotonicity #12922 #13011]: #13013
-#12927 := (iff #12061 #12926)
-#12924 := (iff #4796 #12923)
-#12925 := [rewrite]: #12924
-#12928 := [monotonicity #12925]: #12927
-#13017 := [monotonicity #12928 #13014]: #13016
-#13020 := [monotonicity #12433 #13017]: #13019
-#12939 := (iff #12078 #12938)
-#12936 := (iff #4793 #12935)
-#12933 := (iff #4792 #12932)
-#12934 := [rewrite]: #12933
-#12930 := (iff #4791 #12929)
-#12931 := [rewrite]: #12930
-#12937 := [monotonicity #12931 #12934]: #12936
-#12940 := [monotonicity #12937]: #12939
-#13023 := [monotonicity #12940 #13020]: #13022
-#12949 := (iff #12087 #12948)
-#12946 := (iff #4789 #12945)
-#12943 := (iff #4788 #12941)
-#12944 := [rewrite]: #12943
-#12947 := [monotonicity #12427 #12944]: #12946
-#12950 := [monotonicity #12947]: #12949
-#13026 := [monotonicity #12950 #13023]: #13025
-#12962 := (iff #12096 #12961)
-#12959 := (iff #4785 #12958)
-#12956 := (iff #4784 #12954)
-#12957 := [rewrite]: #12956
-#12952 := (iff #4783 #12951)
-#12953 := [rewrite]: #12952
-#12960 := [monotonicity #12953 #12957]: #12959
-#12963 := [monotonicity #12960]: #12962
-#13029 := [monotonicity #12963 #13026]: #13028
-#12968 := (iff #12105 #12967)
-#12965 := (iff #4781 #12964)
-#12420 := (iff #4648 #12419)
-#12421 := [rewrite]: #12420
-#12966 := [monotonicity #12421]: #12965
-#12969 := [monotonicity #12966]: #12968
-#13032 := [monotonicity #12969 #13029]: #13031
-#13037 := [trans #13032 #13035]: #13036
-#13040 := [monotonicity #12966 #13037]: #13039
-#13045 := [trans #13040 #13043]: #13044
-#12416 := (iff #12117 #12415)
-#12413 := (iff #11268 #12412)
-#12410 := (iff #11263 #12407)
-#12401 := (or #12385 #12397)
-#12404 := (or #6637 #12401)
-#12408 := (iff #12404 #12407)
-#12409 := [rewrite]: #12408
-#12405 := (iff #11263 #12404)
-#12402 := (iff #11257 #12401)
-#12396 := (iff #4776 #12397)
-#12400 := [rewrite]: #12396
-#12394 := (iff #11256 #12385)
-#12386 := (not #12385)
-#12389 := (not #12386)
-#12392 := (iff #12389 #12385)
-#12393 := [rewrite]: #12392
-#12390 := (iff #11256 #12389)
-#12387 := (iff #4772 #12386)
-#12388 := [rewrite]: #12387
-#12391 := [monotonicity #12388]: #12390
-#12395 := [trans #12391 #12393]: #12394
-#12403 := [monotonicity #12395 #12400]: #12402
-#12406 := [monotonicity #6639 #12403]: #12405
-#12411 := [trans #12406 #12409]: #12410
-#12414 := [quant-intro #12411]: #12413
-#12417 := [monotonicity #12414]: #12416
-#13048 := [monotonicity #12417 #13045]: #13047
-#13051 := [monotonicity #12414 #13048]: #13050
-#12383 := (iff #12129 #12382)
-#12380 := (iff #4771 #12379)
-#12381 := [rewrite]: #12380
-#12384 := [monotonicity #12381]: #12383
-#13054 := [monotonicity #12384 #13051]: #13053
-#13057 := [monotonicity #12381 #13054]: #13056
-#12377 := (iff #12141 false)
-#11286 := (iff #4808 false)
-#11287 := [rewrite]: #11286
-#12375 := (iff #12141 #4808)
-#12373 := (iff #11253 true)
-#11531 := (and true true)
-#12368 := (and true #11531)
-#12371 := (iff #12368 true)
-#12372 := [rewrite]: #12371
-#12369 := (iff #11253 #12368)
-#12366 := (iff #11250 #11531)
-#12364 := (iff #4767 true)
-#12365 := [rewrite]: #12364
-#12362 := (iff #4766 true)
-#12363 := [rewrite]: #12362
-#12367 := [monotonicity #12363 #12365]: #12366
-#12370 := [monotonicity #12363 #12367]: #12369
-#12374 := [trans #12370 #12372]: #12373
-#12376 := [monotonicity #12374]: #12375
-#12378 := [trans #12376 #11287]: #12377
-#13060 := [monotonicity #12378 #13057]: #13059
-#13063 := [monotonicity #13060]: #13062
-#13066 := [monotonicity #13063]: #13065
-#13069 := [monotonicity #13066]: #13068
-#13072 := [monotonicity #13069]: #13071
-#13075 := [monotonicity #13072]: #13074
-#13080 := [trans #13075 #13078]: #13079
-#13083 := [monotonicity #13080]: #13082
-#13088 := [trans #13083 #13086]: #13087
-#13091 := [monotonicity #13088]: #13090
-#13094 := [monotonicity #13091]: #13093
-#13099 := [trans #13094 #13097]: #13098
-#13102 := [monotonicity #13099]: #13101
-#13105 := [monotonicity #13102]: #13104
-#12360 := (iff #12222 #12359)
-#12357 := (iff #4727 #12356)
-#12354 := (iff #4726 #12351)
-#12355 := [rewrite]: #12354
-#12348 := (iff #4725 #12349)
-#12350 := [rewrite]: #12348
-#12358 := [monotonicity #12350 #12355]: #12357
-#12361 := [monotonicity #12358]: #12360
-#13186 := [monotonicity #12361 #13105]: #13185
-#13189 := [monotonicity #13186]: #13188
-#13192 := [monotonicity #13189]: #13191
-#13195 := [monotonicity #13192]: #13194
-#13198 := [monotonicity #13195]: #13197
-#13116 := (iff #12267 #13115)
-#13113 := (iff #4694 #13112)
-#13110 := (iff #4693 #13106)
-#13111 := [rewrite]: #13110
-#13114 := [quant-intro #13111]: #13113
-#13117 := [monotonicity #13114]: #13116
-#13201 := [monotonicity #13117 #13198]: #13200
-#13204 := [monotonicity #13201]: #13203
-#13207 := [monotonicity #13204]: #13206
-#13122 := (iff #12294 #13121)
-#13119 := (iff #4679 #13118)
-#13120 := [rewrite]: #13119
-#13123 := [monotonicity #13120]: #13122
-#13210 := [monotonicity #13123 #13207]: #13209
-#13129 := (iff #12303 #12418)
-#13124 := (not #12419)
-#13127 := (iff #13124 #12418)
-#13128 := [rewrite]: #13127
-#13125 := (iff #12303 #13124)
-#13126 := [monotonicity #12421]: #13125
-#13130 := [trans #13126 #13128]: #13129
-#13213 := [monotonicity #13130 #13210]: #13212
-#13140 := (iff #12312 #13131)
-#13132 := (not #13131)
-#13135 := (not #13132)
-#13138 := (iff #13135 #13131)
-#13139 := [rewrite]: #13138
-#13136 := (iff #12312 #13135)
-#13133 := (iff #4647 #13132)
-#13134 := [rewrite]: #13133
-#13137 := [monotonicity #13134]: #13136
-#13141 := [trans #13137 #13139]: #13140
-#13216 := [monotonicity #13141 #13213]: #13215
-#13154 := (iff #12321 #13153)
-#13151 := (iff #4645 #13150)
-#13148 := (iff #4644 #13145)
-#13149 := [rewrite]: #13148
-#13142 := (iff #4643 #13143)
-#13144 := [rewrite]: #13142
-#13152 := [monotonicity #13144 #13149]: #13151
-#13155 := [monotonicity #13152]: #13154
-#13219 := [monotonicity #13155 #13216]: #13218
-#13168 := (iff #12330 #13167)
-#13165 := (iff #4641 #13164)
-#13162 := (iff #4640 #13159)
-#13163 := [rewrite]: #13162
-#13156 := (iff #4639 #13157)
-#13158 := [rewrite]: #13156
-#13166 := [monotonicity #13158 #13163]: #13165
-#13169 := [monotonicity #13166]: #13168
-#13222 := [monotonicity #13169 #13219]: #13221
-#13182 := (iff #12339 #13181)
-#13179 := (iff #4637 #13178)
-#13176 := (iff #4636 #13173)
-#13177 := [rewrite]: #13176
-#13170 := (iff #4635 #13171)
-#13172 := [rewrite]: #13170
-#13180 := [monotonicity #13172 #13177]: #13179
-#13183 := [monotonicity #13180]: #13182
-#13225 := [monotonicity #13183 #13222]: #13224
-#13230 := [trans #13225 #13228]: #13229
-#13233 := [monotonicity #13230]: #13232
-#12346 := (iff #5090 #12345)
-#12343 := (iff #5089 #12340)
-#12336 := (implies #4637 #12331)
-#12341 := (iff #12336 #12340)
-#12342 := [rewrite]: #12341
-#12337 := (iff #5089 #12336)
-#12334 := (iff #5088 #12331)
-#12327 := (implies #4641 #12322)
-#12332 := (iff #12327 #12331)
-#12333 := [rewrite]: #12332
-#12328 := (iff #5088 #12327)
-#12325 := (iff #5087 #12322)
-#12318 := (implies #4645 #12313)
-#12323 := (iff #12318 #12322)
-#12324 := [rewrite]: #12323
-#12319 := (iff #5087 #12318)
-#12316 := (iff #5086 #12313)
-#12309 := (implies #4647 #12304)
-#12314 := (iff #12309 #12313)
-#12315 := [rewrite]: #12314
-#12310 := (iff #5086 #12309)
-#12307 := (iff #5085 #12304)
-#12300 := (implies #4648 #12295)
-#12305 := (iff #12300 #12304)
-#12306 := [rewrite]: #12305
-#12301 := (iff #5085 #12300)
-#12298 := (iff #5084 #12295)
-#12291 := (implies #4679 #12286)
-#12296 := (iff #12291 #12295)
-#12297 := [rewrite]: #12296
-#12292 := (iff #5084 #12291)
-#12289 := (iff #5083 #12286)
-#12282 := (implies #4681 #12277)
-#12287 := (iff #12282 #12286)
-#12288 := [rewrite]: #12287
-#12283 := (iff #5083 #12282)
-#12280 := (iff #5082 #12277)
-#12273 := (implies #4687 #12268)
-#12278 := (iff #12273 #12277)
-#12279 := [rewrite]: #12278
-#12274 := (iff #5082 #12273)
-#12271 := (iff #5081 #12268)
-#12264 := (implies #4694 #12259)
-#12269 := (iff #12264 #12268)
-#12270 := [rewrite]: #12269
-#12265 := (iff #5081 #12264)
-#12262 := (iff #5080 #12259)
-#12255 := (implies #4710 #12250)
-#12260 := (iff #12255 #12259)
-#12261 := [rewrite]: #12260
-#12256 := (iff #5080 #12255)
-#12253 := (iff #5079 #12250)
-#12246 := (implies #4714 #12241)
-#12251 := (iff #12246 #12250)
-#12252 := [rewrite]: #12251
-#12247 := (iff #5079 #12246)
-#12244 := (iff #5078 #12241)
-#12237 := (implies #4717 #12232)
-#12242 := (iff #12237 #12241)
-#12243 := [rewrite]: #12242
-#12238 := (iff #5078 #12237)
-#12235 := (iff #5077 #12232)
-#12228 := (implies #11245 #12223)
-#12233 := (iff #12228 #12232)
-#12234 := [rewrite]: #12233
-#12229 := (iff #5077 #12228)
-#12226 := (iff #5076 #12223)
-#12219 := (implies #4727 #12216)
-#12224 := (iff #12219 #12223)
-#12225 := [rewrite]: #12224
-#12220 := (iff #5076 #12219)
-#12217 := (iff #5075 #12216)
-#12214 := (iff #5074 #12211)
-#12207 := (implies #4733 #12204)
-#12212 := (iff #12207 #12211)
-#12213 := [rewrite]: #12212
-#12208 := (iff #5074 #12207)
-#12205 := (iff #5073 #12204)
-#12202 := (iff #5072 #12199)
-#12195 := (implies #4742 #12192)
-#12200 := (iff #12195 #12199)
-#12201 := [rewrite]: #12200
-#12196 := (iff #5072 #12195)
-#12193 := (iff #5071 #12192)
-#12190 := (iff #5070 #12187)
-#12183 := (implies #4746 #12178)
-#12188 := (iff #12183 #12187)
-#12189 := [rewrite]: #12188
-#12184 := (iff #5070 #12183)
-#12181 := (iff #5069 #12178)
-#12174 := (implies #4750 #12169)
-#12179 := (iff #12174 #12178)
-#12180 := [rewrite]: #12179
-#12175 := (iff #5069 #12174)
-#12172 := (iff #5068 #12169)
-#12165 := (implies #4755 #12160)
-#12170 := (iff #12165 #12169)
-#12171 := [rewrite]: #12170
-#12166 := (iff #5068 #12165)
-#12163 := (iff #5067 #12160)
-#12156 := (implies #4760 #12151)
-#12161 := (iff #12156 #12160)
-#12162 := [rewrite]: #12161
-#12157 := (iff #5067 #12156)
-#12154 := (iff #5066 #12151)
-#12147 := (implies #4765 #12142)
-#12152 := (iff #12147 #12151)
-#12153 := [rewrite]: #12152
-#12148 := (iff #5066 #12147)
-#12145 := (iff #5065 #12142)
-#12138 := (implies #11253 #12135)
-#12143 := (iff #12138 #12142)
-#12144 := [rewrite]: #12143
-#12139 := (iff #5065 #12138)
-#12136 := (iff #5064 #12135)
-#12133 := (iff #5063 #12130)
-#12126 := (implies #4771 #12123)
-#12131 := (iff #12126 #12130)
-#12132 := [rewrite]: #12131
-#12127 := (iff #5063 #12126)
-#12124 := (iff #5062 #12123)
-#12121 := (iff #5061 #12118)
-#12114 := (implies #11268 #12111)
-#12119 := (iff #12114 #12118)
-#12120 := [rewrite]: #12119
-#12115 := (iff #5061 #12114)
-#12112 := (iff #5060 #12111)
-#12109 := (iff #5059 #12106)
-#12102 := (implies #4781 #12097)
-#12107 := (iff #12102 #12106)
-#12108 := [rewrite]: #12107
-#12103 := (iff #5059 #12102)
-#12100 := (iff #5058 #12097)
-#12093 := (implies #4785 #12088)
-#12098 := (iff #12093 #12097)
-#12099 := [rewrite]: #12098
-#12094 := (iff #5058 #12093)
-#12091 := (iff #5057 #12088)
-#12084 := (implies #4789 #12079)
-#12089 := (iff #12084 #12088)
-#12090 := [rewrite]: #12089
-#12085 := (iff #5057 #12084)
-#12082 := (iff #5056 #12079)
-#12075 := (implies #4793 #12070)
-#12080 := (iff #12075 #12079)
-#12081 := [rewrite]: #12080
-#12076 := (iff #5056 #12075)
-#12073 := (iff #5055 #12070)
-#12067 := (implies #4795 #12062)
-#12071 := (iff #12067 #12070)
-#12072 := [rewrite]: #12071
-#12068 := (iff #5055 #12067)
-#12065 := (iff #5054 #12062)
-#12058 := (implies #4796 #12053)
-#12063 := (iff #12058 #12062)
-#12064 := [rewrite]: #12063
-#12059 := (iff #5054 #12058)
-#12056 := (iff #5053 #12053)
-#12049 := (implies #11283 #12044)
-#12054 := (iff #12049 #12053)
-#12055 := [rewrite]: #12054
-#12050 := (iff #5053 #12049)
-#12047 := (iff #5052 #12044)
-#12040 := (implies #4807 #12035)
-#12045 := (iff #12040 #12044)
-#12046 := [rewrite]: #12045
-#12041 := (iff #5052 #12040)
-#12038 := (iff #5051 #12035)
-#12032 := (implies #4795 #12020)
-#12036 := (iff #12032 #12035)
-#12037 := [rewrite]: #12036
-#12033 := (iff #5051 #12032)
-#12030 := (iff #5050 #12020)
-#12025 := (and true #12020)
-#12028 := (iff #12025 #12020)
-#12029 := [rewrite]: #12028
-#12026 := (iff #5050 #12025)
-#12023 := (iff #5049 #12020)
-#12017 := (implies #4795 #12012)
-#12021 := (iff #12017 #12020)
-#12022 := [rewrite]: #12021
-#12018 := (iff #5049 #12017)
-#12015 := (iff #5048 #12012)
-#12009 := (implies #4795 #12004)
-#12013 := (iff #12009 #12012)
-#12014 := [rewrite]: #12013
-#12010 := (iff #5048 #12009)
-#12007 := (iff #5047 #12004)
-#12001 := (implies #4795 #11996)
-#12005 := (iff #12001 #12004)
-#12006 := [rewrite]: #12005
-#12002 := (iff #5047 #12001)
-#11999 := (iff #5046 #11996)
-#11993 := (implies #4893 #11988)
-#11997 := (iff #11993 #11996)
-#11998 := [rewrite]: #11997
-#11994 := (iff #5046 #11993)
-#11991 := (iff #5045 #11988)
-#11984 := (implies #4893 #11979)
-#11989 := (iff #11984 #11988)
-#11990 := [rewrite]: #11989
-#11985 := (iff #5045 #11984)
-#11982 := (iff #5044 #11979)
-#11976 := (implies #4812 #11971)
-#11980 := (iff #11976 #11979)
-#11981 := [rewrite]: #11980
-#11977 := (iff #5044 #11976)
-#11974 := (iff #5043 #11971)
-#11967 := (implies #4902 #11962)
-#11972 := (iff #11967 #11971)
-#11973 := [rewrite]: #11972
-#11968 := (iff #5043 #11967)
-#11965 := (iff #5042 #11962)
-#11958 := (implies #4906 #11953)
-#11963 := (iff #11958 #11962)
-#11964 := [rewrite]: #11963
-#11959 := (iff #5042 #11958)
-#11956 := (iff #5041 #11953)
-#11949 := (implies #4910 #11944)
-#11954 := (iff #11949 #11953)
-#11955 := [rewrite]: #11954
-#11950 := (iff #5041 #11949)
-#11947 := (iff #5040 #11944)
-#11940 := (implies #4914 #11935)
-#11945 := (iff #11940 #11944)
-#11946 := [rewrite]: #11945
-#11941 := (iff #5040 #11940)
-#11938 := (iff #5039 #11935)
-#11931 := (implies #4924 #11919)
-#11936 := (iff #11931 #11935)
-#11937 := [rewrite]: #11936
-#11932 := (iff #5039 #11931)
-#11929 := (iff #5038 #11919)
-#11924 := (implies true #11919)
-#11927 := (iff #11924 #11919)
-#11928 := [rewrite]: #11927
-#11925 := (iff #5038 #11924)
-#11922 := (iff #5037 #11919)
-#11916 := (implies #4795 #11913)
-#11920 := (iff #11916 #11919)
-#11921 := [rewrite]: #11920
-#11917 := (iff #5037 #11916)
-#11914 := (iff #5036 #11913)
-#11911 := (iff #5035 #11908)
-#11905 := (implies #4795 #11900)
-#11909 := (iff #11905 #11908)
-#11910 := [rewrite]: #11909
-#11906 := (iff #5035 #11905)
-#11903 := (iff #5034 #11900)
-#11896 := (implies #5031 #11891)
-#11901 := (iff #11896 #11900)
-#11902 := [rewrite]: #11901
-#11897 := (iff #5034 #11896)
-#11894 := (iff #5033 #11891)
-#11888 := (implies #4795 #11883)
-#11892 := (iff #11888 #11891)
-#11893 := [rewrite]: #11892
-#11889 := (iff #5033 #11888)
-#11886 := (iff #5032 #11883)
-#11880 := (implies #4795 #11401)
-#11884 := (iff #11880 #11883)
-#11885 := [rewrite]: #11884
-#11881 := (iff #5032 #11880)
-#11404 := (iff #4844 #11401)
-#11398 := (implies #4795 #11393)
-#11402 := (iff #11398 #11401)
-#11403 := [rewrite]: #11402
-#11399 := (iff #4844 #11398)
-#11396 := (iff #4843 #11393)
-#11390 := (implies #4795 #11387)
-#11394 := (iff #11390 #11393)
-#11395 := [rewrite]: #11394
-#11391 := (iff #4843 #11390)
-#11388 := (iff #4842 #11387)
-#11385 := (iff #4841 #11382)
-#11378 := (implies #3125 #11373)
-#11383 := (iff #11378 #11382)
-#11384 := [rewrite]: #11383
-#11379 := (iff #4841 #11378)
-#11376 := (iff #4840 #11373)
-#11370 := (implies #4795 #11365)
-#11374 := (iff #11370 #11373)
-#11375 := [rewrite]: #11374
-#11371 := (iff #4840 #11370)
-#11368 := (iff #4839 #11365)
-#11362 := (implies #4795 #11357)
-#11366 := (iff #11362 #11365)
-#11367 := [rewrite]: #11366
-#11363 := (iff #4839 #11362)
-#11360 := (iff #4838 #11357)
-#11353 := (implies #4795 #11348)
-#11358 := (iff #11353 #11357)
-#11359 := [rewrite]: #11358
-#11354 := (iff #4838 #11353)
-#11351 := (iff #4837 #11348)
-#11344 := (implies #4814 #11339)
-#11349 := (iff #11344 #11348)
-#11350 := [rewrite]: #11349
-#11345 := (iff #4837 #11344)
-#11342 := (iff #4836 #11339)
-#11335 := (implies #4816 #11330)
-#11340 := (iff #11335 #11339)
-#11341 := [rewrite]: #11340
-#11336 := (iff #4836 #11335)
-#11333 := (iff #4835 #11330)
-#11326 := (implies #4818 #11321)
-#11331 := (iff #11326 #11330)
-#11332 := [rewrite]: #11331
-#11327 := (iff #4835 #11326)
-#11324 := (iff #4834 #11321)
-#11317 := (implies #4820 #11314)
-#11322 := (iff #11317 #11321)
-#11323 := [rewrite]: #11322
-#11318 := (iff #4834 #11317)
-#11315 := (iff #4833 #11314)
-#11312 := (iff #4832 #11309)
-#11305 := (implies #11300 #4830)
-#11310 := (iff #11305 #11309)
-#11311 := [rewrite]: #11310
-#11306 := (iff #4832 #11305)
-#11303 := (iff #4831 #4830)
-#11304 := [rewrite]: #11303
-#11301 := (iff #4825 #11300)
-#11298 := (iff #4824 #11295)
-#11292 := (implies #568 #11289)
-#11296 := (iff #11292 #11295)
-#11297 := [rewrite]: #11296
-#11293 := (iff #4824 #11292)
-#11290 := (iff #4823 #11289)
-#11291 := [rewrite]: #11290
-#11294 := [monotonicity #11291]: #11293
-#11299 := [trans #11294 #11297]: #11298
-#11302 := [quant-intro #11299]: #11301
-#11307 := [monotonicity #11302 #11304]: #11306
-#11313 := [trans #11307 #11311]: #11312
-#11316 := [monotonicity #11302 #11313]: #11315
-#11319 := [monotonicity #11316]: #11318
-#11325 := [trans #11319 #11323]: #11324
-#11328 := [monotonicity #11325]: #11327
-#11334 := [trans #11328 #11332]: #11333
-#11337 := [monotonicity #11334]: #11336
-#11343 := [trans #11337 #11341]: #11342
-#11346 := [monotonicity #11343]: #11345
-#11352 := [trans #11346 #11350]: #11351
-#11355 := [monotonicity #11352]: #11354
-#11361 := [trans #11355 #11359]: #11360
-#11364 := [monotonicity #11361]: #11363
-#11369 := [trans #11364 #11367]: #11368
-#11372 := [monotonicity #11369]: #11371
-#11377 := [trans #11372 #11375]: #11376
-#11380 := [monotonicity #11377]: #11379
-#11386 := [trans #11380 #11384]: #11385
-#11389 := [monotonicity #11386]: #11388
-#11392 := [monotonicity #11389]: #11391
-#11397 := [trans #11392 #11395]: #11396
-#11400 := [monotonicity #11397]: #11399
-#11405 := [trans #11400 #11403]: #11404
-#11882 := [monotonicity #11405]: #11881
-#11887 := [trans #11882 #11885]: #11886
-#11890 := [monotonicity #11887]: #11889
-#11895 := [trans #11890 #11893]: #11894
-#11898 := [monotonicity #11895]: #11897
-#11904 := [trans #11898 #11902]: #11903
-#11907 := [monotonicity #11904]: #11906
-#11912 := [trans #11907 #11910]: #11911
-#11878 := (iff #5030 #11875)
-#11872 := (implies #4795 #11867)
-#11876 := (iff #11872 #11875)
-#11877 := [rewrite]: #11876
-#11873 := (iff #5030 #11872)
-#11870 := (iff #5029 #11867)
-#11863 := (implies #4928 #11858)
-#11868 := (iff #11863 #11867)
-#11869 := [rewrite]: #11868
-#11864 := (iff #5029 #11863)
-#11861 := (iff #5028 #11858)
-#11855 := (implies #4795 #11850)
-#11859 := (iff #11855 #11858)
-#11860 := [rewrite]: #11859
-#11856 := (iff #5028 #11855)
-#11853 := (iff #5027 #11850)
-#11847 := (implies #4795 #11844)
-#11851 := (iff #11847 #11850)
-#11852 := [rewrite]: #11851
-#11848 := (iff #5027 #11847)
-#11845 := (iff #5026 #11844)
-#11842 := (iff #5025 #11839)
-#11836 := (implies #4936 #11833)
-#11840 := (iff #11836 #11839)
-#11841 := [rewrite]: #11840
-#11837 := (iff #5025 #11836)
-#11834 := (iff #5024 #11833)
-#11831 := (iff #5023 #11828)
-#11825 := (implies #4939 #11820)
-#11829 := (iff #11825 #11828)
-#11830 := [rewrite]: #11829
-#11826 := (iff #5023 #11825)
-#11823 := (iff #5022 #11820)
-#11817 := (implies #4795 #11814)
-#11821 := (iff #11817 #11820)
-#11822 := [rewrite]: #11821
-#11818 := (iff #5022 #11817)
-#11815 := (iff #5021 #11814)
-#11812 := (iff #5020 #11809)
-#11806 := (implies #4795 #11801)
-#11810 := (iff #11806 #11809)
-#11811 := [rewrite]: #11810
-#11807 := (iff #5020 #11806)
-#11804 := (iff #5019 #11801)
-#11797 := (implies #5011 #11792)
-#11802 := (iff #11797 #11801)
-#11803 := [rewrite]: #11802
-#11798 := (iff #5019 #11797)
-#11795 := (iff #5018 #11792)
-#11789 := (implies #4795 #11784)
-#11793 := (iff #11789 #11792)
-#11794 := [rewrite]: #11793
-#11790 := (iff #5018 #11789)
-#11787 := (iff #5017 #11784)
-#11781 := (implies #4795 #11776)
-#11785 := (iff #11781 #11784)
-#11786 := [rewrite]: #11785
-#11782 := (iff #5017 #11781)
-#11779 := (iff #5016 #11776)
-#11773 := (implies #4795 #11768)
-#11777 := (iff #11773 #11776)
-#11778 := [rewrite]: #11777
-#11774 := (iff #5016 #11773)
-#11771 := (iff #5015 #11768)
-#11764 := (implies #5012 #11759)
-#11769 := (iff #11764 #11768)
-#11770 := [rewrite]: #11769
-#11765 := (iff #5015 #11764)
-#11762 := (iff #5014 #11759)
-#11755 := (implies #5013 #11639)
-#11760 := (iff #11755 #11759)
-#11761 := [rewrite]: #11760
-#11756 := (iff #5014 #11755)
-#11642 := (iff #4996 #11639)
-#11635 := (implies #4960 #11632)
-#11640 := (iff #11635 #11639)
-#11641 := [rewrite]: #11640
-#11636 := (iff #4996 #11635)
-#11633 := (iff #4995 #11632)
-#11630 := (iff #4994 #11627)
-#11623 := (implies #11549 #11618)
-#11628 := (iff #11623 #11627)
-#11629 := [rewrite]: #11628
-#11624 := (iff #4994 #11623)
-#11621 := (iff #4993 #11618)
-#11614 := (implies #11552 #11609)
-#11619 := (iff #11614 #11618)
-#11620 := [rewrite]: #11619
-#11615 := (iff #4993 #11614)
-#11612 := (iff #4992 #11609)
-#11605 := (implies #4971 #11600)
-#11610 := (iff #11605 #11609)
-#11611 := [rewrite]: #11610
-#11606 := (iff #4992 #11605)
-#11603 := (iff #4991 #11600)
-#11596 := (implies #4973 #11593)
-#11601 := (iff #11596 #11600)
-#11602 := [rewrite]: #11601
-#11597 := (iff #4991 #11596)
-#11594 := (iff #4990 #11593)
-#11591 := (iff #4989 #11588)
-#11584 := (implies #4974 #11581)
-#11589 := (iff #11584 #11588)
-#11590 := [rewrite]: #11589
-#11585 := (iff #4989 #11584)
-#11582 := (iff #4988 #11581)
-#11579 := (iff #4987 #11576)
-#11572 := (implies #11567 #4985)
-#11577 := (iff #11572 #11576)
-#11578 := [rewrite]: #11577
-#11573 := (iff #4987 #11572)
-#11570 := (iff #4986 #4985)
-#11571 := [rewrite]: #11570
-#11568 := (iff #4979 #11567)
-#11565 := (iff #4978 #11562)
-#11559 := (implies #568 #11556)
-#11563 := (iff #11559 #11562)
-#11564 := [rewrite]: #11563
-#11560 := (iff #4978 #11559)
-#11557 := (iff #4977 #11556)
-#11558 := [rewrite]: #11557
-#11561 := [monotonicity #11558]: #11560
-#11566 := [trans #11561 #11564]: #11565
-#11569 := [quant-intro #11566]: #11568
-#11574 := [monotonicity #11569 #11571]: #11573
-#11580 := [trans #11574 #11578]: #11579
-#11583 := [monotonicity #11569 #11580]: #11582
-#11586 := [monotonicity #11583]: #11585
-#11592 := [trans #11586 #11590]: #11591
-#11595 := [monotonicity #11592]: #11594
-#11598 := [monotonicity #11595]: #11597
-#11604 := [trans #11598 #11602]: #11603
-#11607 := [monotonicity #11604]: #11606
-#11613 := [trans #11607 #11611]: #11612
-#11553 := (iff #4966 #11552)
-#11541 := (= #4961 #11540)
-#11542 := [rewrite]: #11541
-#11554 := [monotonicity #11542]: #11553
-#11616 := [monotonicity #11554 #11613]: #11615
-#11622 := [trans #11616 #11620]: #11621
-#11550 := (iff #4964 #11549)
-#11547 := (iff #4963 #11546)
-#11548 := [monotonicity #11542]: #11547
-#11544 := (iff #4962 #11543)
-#11545 := [monotonicity #11542]: #11544
-#11551 := [monotonicity #11545 #11548]: #11550
-#11625 := [monotonicity #11551 #11622]: #11624
-#11631 := [trans #11625 #11629]: #11630
-#11634 := [monotonicity #11551 #11631]: #11633
-#11637 := [monotonicity #11634]: #11636
-#11643 := [trans #11637 #11641]: #11642
-#11757 := [monotonicity #11643]: #11756
-#11763 := [trans #11757 #11761]: #11762
-#11766 := [monotonicity #11763]: #11765
-#11772 := [trans #11766 #11770]: #11771
-#11775 := [monotonicity #11772]: #11774
-#11780 := [trans #11775 #11778]: #11779
-#11783 := [monotonicity #11780]: #11782
-#11788 := [trans #11783 #11786]: #11787
-#11791 := [monotonicity #11788]: #11790
-#11796 := [trans #11791 #11794]: #11795
-#11799 := [monotonicity #11796]: #11798
-#11805 := [trans #11799 #11803]: #11804
-#11808 := [monotonicity #11805]: #11807
-#11813 := [trans #11808 #11811]: #11812
-#11753 := (iff #5010 #11750)
-#11747 := (implies #4795 #11742)
-#11751 := (iff #11747 #11750)
-#11752 := [rewrite]: #11751
-#11748 := (iff #5010 #11747)
-#11745 := (iff #5009 #11742)
-#11738 := (implies #4941 #11733)
-#11743 := (iff #11738 #11742)
-#11744 := [rewrite]: #11743
-#11739 := (iff #5009 #11738)
-#11736 := (iff #5008 #11733)
-#11730 := (implies #4795 #11725)
-#11734 := (iff #11730 #11733)
-#11735 := [rewrite]: #11734
-#11731 := (iff #5008 #11730)
-#11728 := (iff #5007 #11725)
-#11722 := (implies #4795 #11719)
-#11726 := (iff #11722 #11725)
-#11727 := [rewrite]: #11726
-#11723 := (iff #5007 #11722)
-#11720 := (iff #5006 #11719)
-#11717 := (iff #5005 #11714)
-#11710 := (implies #4936 #11707)
-#11715 := (iff #11710 #11714)
-#11716 := [rewrite]: #11715
-#11711 := (iff #5005 #11710)
-#11708 := (iff #5004 #11707)
-#11705 := (iff #5003 #11702)
-#11698 := (implies #4939 #11693)
-#11703 := (iff #11698 #11702)
-#11704 := [rewrite]: #11703
-#11699 := (iff #5003 #11698)
-#11696 := (iff #5002 #11693)
-#11689 := (implies #4943 #11684)
-#11694 := (iff #11689 #11693)
-#11695 := [rewrite]: #11694
-#11690 := (iff #5002 #11689)
-#11687 := (iff #5001 #11684)
-#11680 := (implies #4948 #11675)
-#11685 := (iff #11680 #11684)
-#11686 := [rewrite]: #11685
-#11681 := (iff #5001 #11680)
-#11678 := (iff #5000 #11675)
-#11671 := (implies #4953 #11666)
-#11676 := (iff #11671 #11675)
-#11677 := [rewrite]: #11676
-#11672 := (iff #5000 #11671)
-#11669 := (iff #4999 #11666)
-#11662 := (implies #4794 #11657)
-#11667 := (iff #11662 #11666)
-#11668 := [rewrite]: #11667
-#11663 := (iff #4999 #11662)
-#11660 := (iff #4998 #11657)
-#11653 := (implies #4956 #11648)
-#11658 := (iff #11653 #11657)
-#11659 := [rewrite]: #11658
-#11654 := (iff #4998 #11653)
-#11651 := (iff #4997 #11648)
-#11644 := (implies #4958 #11639)
-#11649 := (iff #11644 #11648)
-#11650 := [rewrite]: #11649
-#11645 := (iff #4997 #11644)
-#11646 := [monotonicity #11643]: #11645
-#11652 := [trans #11646 #11650]: #11651
-#11655 := [monotonicity #11652]: #11654
-#11661 := [trans #11655 #11659]: #11660
-#11538 := (iff #4954 #4794)
-#11539 := [rewrite]: #11538
-#11664 := [monotonicity #11539 #11661]: #11663
-#11670 := [trans #11664 #11668]: #11669
-#11673 := [monotonicity #11670]: #11672
-#11679 := [trans #11673 #11677]: #11678
-#11682 := [monotonicity #11679]: #11681
-#11688 := [trans #11682 #11686]: #11687
-#11691 := [monotonicity #11688]: #11690
-#11697 := [trans #11691 #11695]: #11696
-#11700 := [monotonicity #11697]: #11699
-#11706 := [trans #11700 #11704]: #11705
-#11709 := [monotonicity #11706]: #11708
-#11712 := [monotonicity #11709]: #11711
-#11718 := [trans #11712 #11716]: #11717
-#11721 := [monotonicity #11718]: #11720
-#11724 := [monotonicity #11721]: #11723
-#11729 := [trans #11724 #11727]: #11728
-#11732 := [monotonicity #11729]: #11731
-#11737 := [trans #11732 #11735]: #11736
-#11740 := [monotonicity #11737]: #11739
-#11746 := [trans #11740 #11744]: #11745
-#11749 := [monotonicity #11746]: #11748
-#11754 := [trans #11749 #11752]: #11753
-#11816 := [monotonicity #11754 #11813]: #11815
-#11819 := [monotonicity #11816]: #11818
-#11824 := [trans #11819 #11822]: #11823
-#11827 := [monotonicity #11824]: #11826
-#11832 := [trans #11827 #11830]: #11831
-#11835 := [monotonicity #11832]: #11834
-#11838 := [monotonicity #11835]: #11837
-#11843 := [trans #11838 #11841]: #11842
-#11846 := [monotonicity #11843]: #11845
-#11849 := [monotonicity #11846]: #11848
-#11854 := [trans #11849 #11852]: #11853
-#11857 := [monotonicity #11854]: #11856
-#11862 := [trans #11857 #11860]: #11861
-#11865 := [monotonicity #11862]: #11864
-#11871 := [trans #11865 #11869]: #11870
-#11874 := [monotonicity #11871]: #11873
-#11879 := [trans #11874 #11877]: #11878
-#11915 := [monotonicity #11879 #11912]: #11914
-#11918 := [monotonicity #11915]: #11917
-#11923 := [trans #11918 #11921]: #11922
-#11536 := (iff #4927 true)
-#11534 := (iff #11531 true)
-#11535 := [rewrite]: #11534
-#11532 := (iff #4927 #11531)
-#11529 := (iff #4926 true)
-#11530 := [rewrite]: #11529
-#11527 := (iff #4925 true)
-#11528 := [rewrite]: #11527
-#11533 := [monotonicity #11528 #11530]: #11532
-#11537 := [trans #11533 #11535]: #11536
-#11926 := [monotonicity #11537 #11923]: #11925
-#11930 := [trans #11926 #11928]: #11929
-#11933 := [monotonicity #11930]: #11932
-#11939 := [trans #11933 #11937]: #11938
-#11942 := [monotonicity #11939]: #11941
-#11948 := [trans #11942 #11946]: #11947
-#11951 := [monotonicity #11948]: #11950
-#11957 := [trans #11951 #11955]: #11956
-#11960 := [monotonicity #11957]: #11959
-#11966 := [trans #11960 #11964]: #11965
-#11969 := [monotonicity #11966]: #11968
-#11975 := [trans #11969 #11973]: #11974
-#11978 := [monotonicity #11975]: #11977
-#11983 := [trans #11978 #11981]: #11982
-#11986 := [monotonicity #11983]: #11985
-#11992 := [trans #11986 #11990]: #11991
-#11525 := (iff #4897 #4893)
-#11508 := (and true #4893)
-#11511 := (iff #11508 #4893)
-#11512 := [rewrite]: #11511
-#11523 := (iff #4897 #11508)
-#11521 := (iff #4896 #4893)
-#11519 := (iff #4896 #11508)
-#11517 := (iff #4895 #4893)
-#11515 := (iff #4895 #11508)
-#11513 := (iff #4894 #4893)
-#11509 := (iff #4894 #11508)
-#11506 := (iff #4882 true)
-#11501 := (forall (vars (?v0 S10)) (:pat #4878) true)
-#11504 := (iff #11501 true)
-#11505 := [elim-unused]: #11504
-#11502 := (iff #4882 #11501)
-#11499 := (iff #4881 true)
-#11463 := (implies #4866 #4866)
-#11466 := (iff #11463 true)
-#11467 := [rewrite]: #11466
-#11497 := (iff #4881 #11463)
-#11495 := (iff #4880 #4866)
-#11456 := (and true #4866)
-#11459 := (iff #11456 #4866)
-#11460 := [rewrite]: #11459
-#11493 := (iff #4880 #11456)
-#11491 := (iff #4879 true)
-#11492 := [rewrite]: #11491
-#11494 := [monotonicity #11492]: #11493
-#11496 := [trans #11494 #11460]: #11495
-#11498 := [monotonicity #11496]: #11497
-#11500 := [trans #11498 #11467]: #11499
-#11503 := [quant-intro #11500]: #11502
-#11507 := [trans #11503 #11505]: #11506
-#11510 := [monotonicity #11507]: #11509
-#11514 := [trans #11510 #11512]: #11513
-#11489 := (iff #4874 true)
-#11447 := (forall (vars (?v0 S10)) (:pat #4853) true)
-#11450 := (iff #11447 true)
-#11451 := [elim-unused]: #11450
-#11487 := (iff #4874 #11447)
-#11485 := (iff #4873 true)
-#11483 := (iff #4873 #11463)
-#11481 := (iff #4872 #4866)
-#11479 := (iff #4872 #11456)
-#11477 := (iff #4871 true)
-#11478 := [rewrite]: #11477
-#11480 := [monotonicity #11478]: #11479
-#11482 := [trans #11480 #11460]: #11481
-#11484 := [monotonicity #11482]: #11483
-#11486 := [trans #11484 #11467]: #11485
-#11488 := [quant-intro #11486]: #11487
-#11490 := [trans #11488 #11451]: #11489
-#11516 := [monotonicity #11490 #11514]: #11515
-#11518 := [trans #11516 #11512]: #11517
-#11475 := (iff #4870 true)
-#11470 := (forall (vars (?v0 S10)) (:pat #4864) true)
-#11473 := (iff #11470 true)
-#11474 := [elim-unused]: #11473
-#11471 := (iff #4870 #11470)
-#11468 := (iff #4869 true)
-#11464 := (iff #4869 #11463)
-#11461 := (iff #4868 #4866)
-#11457 := (iff #4868 #11456)
-#11454 := (iff #4867 true)
-#11455 := [rewrite]: #11454
-#11458 := [monotonicity #11455]: #11457
-#11462 := [trans #11458 #11460]: #11461
-#11465 := [monotonicity #11462]: #11464
-#11469 := [trans #11465 #11467]: #11468
-#11472 := [quant-intro #11469]: #11471
-#11476 := [trans #11472 #11474]: #11475
-#11520 := [monotonicity #11476 #11518]: #11519
-#11522 := [trans #11520 #11512]: #11521
-#11452 := (iff #4860 true)
-#11448 := (iff #4860 #11447)
-#11445 := (iff #4859 true)
-#11446 := [rewrite]: #11445
-#11449 := [quant-intro #11446]: #11448
-#11453 := [trans #11449 #11451]: #11452
-#11524 := [monotonicity #11453 #11522]: #11523
-#11526 := [trans #11524 #11512]: #11525
-#11995 := [monotonicity #11526 #11992]: #11994
-#12000 := [trans #11995 #11998]: #11999
-#12003 := [monotonicity #12000]: #12002
-#12008 := [trans #12003 #12006]: #12007
-#12011 := [monotonicity #12008]: #12010
-#12016 := [trans #12011 #12014]: #12015
-#12019 := [monotonicity #12016]: #12018
-#12024 := [trans #12019 #12022]: #12023
-#11443 := (iff #4849 true)
-#11438 := (implies #4795 true)
-#11441 := (iff #11438 true)
-#11442 := [rewrite]: #11441
-#11439 := (iff #4849 #11438)
-#11436 := (iff #4848 true)
-#11410 := (or #11409 #11401)
-#11418 := (or #11356 #11410)
-#11426 := (or #11356 #11418)
-#11431 := (implies false #11426)
-#11434 := (iff #11431 true)
-#11435 := [rewrite]: #11434
-#11432 := (iff #4848 #11431)
-#11429 := (iff #4847 #11426)
-#11423 := (implies #4795 #11418)
-#11427 := (iff #11423 #11426)
-#11428 := [rewrite]: #11427
-#11424 := (iff #4847 #11423)
-#11421 := (iff #4846 #11418)
-#11415 := (implies #4795 #11410)
-#11419 := (iff #11415 #11418)
-#11420 := [rewrite]: #11419
-#11416 := (iff #4846 #11415)
-#11413 := (iff #4845 #11410)
-#11406 := (implies #4812 #11401)
-#11411 := (iff #11406 #11410)
-#11412 := [rewrite]: #11411
-#11407 := (iff #4845 #11406)
-#11408 := [monotonicity #11405]: #11407
-#11414 := [trans #11408 #11412]: #11413
-#11417 := [monotonicity #11414]: #11416
-#11422 := [trans #11417 #11420]: #11421
-#11425 := [monotonicity #11422]: #11424
-#11430 := [trans #11425 #11428]: #11429
-#11433 := [monotonicity #11287 #11430]: #11432
-#11437 := [trans #11433 #11435]: #11436
-#11440 := [monotonicity #11437]: #11439
-#11444 := [trans #11440 #11442]: #11443
-#12027 := [monotonicity #11444 #12024]: #12026
-#12031 := [trans #12027 #12029]: #12030
-#12034 := [monotonicity #12031]: #12033
-#12039 := [trans #12034 #12037]: #12038
-#12042 := [monotonicity #12039]: #12041
-#12048 := [trans #12042 #12046]: #12047
-#11284 := (iff #4801 #11283)
-#11281 := (iff #4800 #11278)
-#11275 := (implies #568 #11272)
-#11279 := (iff #11275 #11278)
-#11280 := [rewrite]: #11279
-#11276 := (iff #4800 #11275)
-#11273 := (iff #4799 #11272)
-#11274 := [rewrite]: #11273
-#11277 := [monotonicity #11274]: #11276
-#11282 := [trans #11277 #11280]: #11281
-#11285 := [quant-intro #11282]: #11284
-#12051 := [monotonicity #11285 #12048]: #12050
-#12057 := [trans #12051 #12055]: #12056
-#12060 := [monotonicity #12057]: #12059
-#12066 := [trans #12060 #12064]: #12065
-#12069 := [monotonicity #12066]: #12068
-#12074 := [trans #12069 #12072]: #12073
-#12077 := [monotonicity #12074]: #12076
-#12083 := [trans #12077 #12081]: #12082
-#12086 := [monotonicity #12083]: #12085
-#12092 := [trans #12086 #12090]: #12091
-#12095 := [monotonicity #12092]: #12094
-#12101 := [trans #12095 #12099]: #12100
-#12104 := [monotonicity #12101]: #12103
-#12110 := [trans #12104 #12108]: #12109
-#12113 := [monotonicity #12110]: #12112
-#11269 := (iff #4779 #11268)
-#11266 := (iff #4778 #11263)
-#11260 := (implies #568 #11257)
-#11264 := (iff #11260 #11263)
-#11265 := [rewrite]: #11264
-#11261 := (iff #4778 #11260)
-#11258 := (iff #4777 #11257)
-#11259 := [rewrite]: #11258
-#11262 := [monotonicity #11259]: #11261
-#11267 := [trans #11262 #11265]: #11266
-#11270 := [quant-intro #11267]: #11269
-#12116 := [monotonicity #11270 #12113]: #12115
-#12122 := [trans #12116 #12120]: #12121
-#12125 := [monotonicity #11270 #12122]: #12124
-#12128 := [monotonicity #12125]: #12127
-#12134 := [trans #12128 #12132]: #12133
-#12137 := [monotonicity #12134]: #12136
-#11254 := (iff #4770 #11253)
-#11251 := (iff #4769 #11250)
-#11248 := (iff #4768 #4767)
-#11249 := [rewrite]: #11248
-#11252 := [monotonicity #11249]: #11251
-#11255 := [monotonicity #11252]: #11254
-#12140 := [monotonicity #11255 #12137]: #12139
-#12146 := [trans #12140 #12144]: #12145
-#12149 := [monotonicity #12146]: #12148
-#12155 := [trans #12149 #12153]: #12154
-#12158 := [monotonicity #12155]: #12157
-#12164 := [trans #12158 #12162]: #12163
-#12167 := [monotonicity #12164]: #12166
-#12173 := [trans #12167 #12171]: #12172
-#12176 := [monotonicity #12173]: #12175
-#12182 := [trans #12176 #12180]: #12181
-#12185 := [monotonicity #12182]: #12184
-#12191 := [trans #12185 #12189]: #12190
-#12194 := [monotonicity #12191]: #12193
-#12197 := [monotonicity #12194]: #12196
-#12203 := [trans #12197 #12201]: #12202
-#12206 := [monotonicity #12203]: #12205
-#12209 := [monotonicity #12206]: #12208
-#12215 := [trans #12209 #12213]: #12214
-#12218 := [monotonicity #12215]: #12217
-#12221 := [monotonicity #12218]: #12220
-#12227 := [trans #12221 #12225]: #12226
-#11246 := (iff #4724 #11245)
-#11243 := (iff #4723 #11242)
-#11244 := [rewrite]: #11243
-#11247 := [quant-intro #11244]: #11246
-#12230 := [monotonicity #11247 #12227]: #12229
-#12236 := [trans #12230 #12234]: #12235
-#12239 := [monotonicity #12236]: #12238
-#12245 := [trans #12239 #12243]: #12244
-#12248 := [monotonicity #12245]: #12247
-#12254 := [trans #12248 #12252]: #12253
-#12257 := [monotonicity #12254]: #12256
-#12263 := [trans #12257 #12261]: #12262
-#12266 := [monotonicity #12263]: #12265
-#12272 := [trans #12266 #12270]: #12271
-#12275 := [monotonicity #12272]: #12274
-#12281 := [trans #12275 #12279]: #12280
-#12284 := [monotonicity #12281]: #12283
-#12290 := [trans #12284 #12288]: #12289
-#12293 := [monotonicity #12290]: #12292
-#12299 := [trans #12293 #12297]: #12298
-#12302 := [monotonicity #12299]: #12301
-#12308 := [trans #12302 #12306]: #12307
-#12311 := [monotonicity #12308]: #12310
-#12317 := [trans #12311 #12315]: #12316
-#12320 := [monotonicity #12317]: #12319
-#12326 := [trans #12320 #12324]: #12325
-#12329 := [monotonicity #12326]: #12328
-#12335 := [trans #12329 #12333]: #12334
-#12338 := [monotonicity #12335]: #12337
-#12344 := [trans #12338 #12342]: #12343
-#12347 := [monotonicity #12344]: #12346
-#13235 := [trans #12347 #13233]: #13234
-#11241 := [asserted]: #5090
-#13236 := [mp #11241 #13235]: #13231
-#13248 := [not-or-elim #13236]: #13118
-#13251 := [and-elim #13248]: #4666
-#300 := (f85 f90 #28)
-#3371 := (f7 #300 #333)
-#3372 := (pattern #3371)
-#1536 := (f20 f179 #28)
-#375 := (f80 f81 #333)
-#3374 := (f53 #375 #1536)
-#3375 := (= #28 #3374)
-#3373 := (= #3371 f1)
-#9811 := (not #3373)
-#9812 := (or #9811 #3375)
-#9815 := (forall (vars (?v0 S10) (?v1 S6)) (:pat #3372) #9812)
-#16674 := (~ #9815 #9815)
-#16672 := (~ #9812 #9812)
-#16673 := [refl]: #16672
-#16675 := [nnf-pos #16673]: #16674
-#3376 := (implies #3373 #3375)
-#3377 := (forall (vars (?v0 S10) (?v1 S6)) (:pat #3372) #3376)
-#9816 := (iff #3377 #9815)
-#9813 := (iff #3376 #9812)
-#9814 := [rewrite]: #9813
-#9817 := [quant-intro #9814]: #9816
-#9810 := [asserted]: #3377
-#9820 := [mp #9810 #9817]: #9815
-#16676 := [mp~ #9820 #16675]: #9815
-#23196 := (not #4666)
-#23214 := (not #9815)
-#23215 := (or #23214 #23196 #23210)
-#23211 := (or #23196 #23210)
-#23216 := (or #23214 #23211)
-#23218 := (iff #23216 #23215)
-#23219 := [rewrite]: #23218
-#23217 := [quant-inst #4658 #4652]: #23216
-#23220 := [mp #23217 #23219]: #23215
-#24763 := [unit-resolution #23220 #16676 #13251]: #23210
-#23716 := [symm #24763]: #23715
-#23718 := [monotonicity #23716]: #23717
-#23720 := [trans #23718 #23700]: #23719
-#23722 := [monotonicity #23720]: #23721
-#23724 := [trans #23722 #23699]: #23723
-#23714 := [monotonicity #23724]: #23713
-#23735 := [trans #23714 #23733]: #23734
-#23737 := [symm #23735]: #23736
-#23740 := [monotonicity #23737]: #23739
-#4626 := (= f420 f419)
-#4627 := (not #4626)
-#4624 := (= f29 f419)
-#4625 := (not #4624)
-decl f421 :: S21
-#3600 := f421
-#4620 := (= f421 f419)
-#4621 := (not #4620)
-#4618 := (= f421 f420)
-#4619 := (not #4618)
-#4616 := (= f421 f29)
-#4617 := (not #4616)
-#11229 := (and #4617 #4619 #4621 #4623 #4625 #4627)
-#4628 := (and #4627 true)
-#4629 := (and #4625 #4628)
-#4630 := (and #4623 #4629)
-#4631 := (and #4621 #4630)
-#4632 := (and #4619 #4631)
-#4633 := (and #4617 #4632)
-#11232 := (iff #4633 #11229)
-#11214 := (and #4625 #4627)
-#11217 := (and #4623 #11214)
-#11220 := (and #4621 #11217)
-#11223 := (and #4619 #11220)
-#11226 := (and #4617 #11223)
-#11230 := (iff #11226 #11229)
-#11231 := [rewrite]: #11230
-#11227 := (iff #4633 #11226)
-#11224 := (iff #4632 #11223)
-#11221 := (iff #4631 #11220)
-#11218 := (iff #4630 #11217)
-#11215 := (iff #4629 #11214)
-#11212 := (iff #4628 #4627)
-#11213 := [rewrite]: #11212
-#11216 := [monotonicity #11213]: #11215
-#11219 := [monotonicity #11216]: #11218
-#11222 := [monotonicity #11219]: #11221
-#11225 := [monotonicity #11222]: #11224
-#11228 := [monotonicity #11225]: #11227
-#11233 := [trans #11228 #11231]: #11232
-#11211 := [asserted]: #4633
-#11234 := [mp #11211 #11233]: #11229
-#11238 := [and-elim #11234]: #4623
-#23741 := [mp #11238 #23740]: #23738
-#23581 := (not #23573)
-#23584 := (not #23569)
-#23766 := (iff #12210 #23584)
-#23764 := (iff #4733 #23569)
-#23751 := (iff #23569 #4733)
-#23749 := (= #23568 #4732)
-#23744 := (= #23560 #4730)
-#23731 := (= #23559 #4729)
-#23743 := [monotonicity #23716]: #23731
-#23745 := [monotonicity #23743 #23716]: #23744
-#23750 := [monotonicity #23745]: #23749
-#23763 := [monotonicity #23750]: #23751
-#23765 := [symm #23763]: #23764
-#23767 := [monotonicity #23765]: #23766
-#23730 := [hypothesis]: #12210
-#23768 := [mp #23730 #23767]: #23584
-#23587 := (not #23574)
-#23588 := (or #23587 #23569 #23581)
-#23589 := [def-axiom]: #23588
-#23769 := [unit-resolution #23589 #23768 #23729]: #23581
-#23646 := (f37 #4667 #23197)
-#23647 := (= #23646 f1)
-#13252 := [and-elim #13248]: #4669
-#23770 := (= #23646 #4668)
-#23771 := [monotonicity #23716]: #23770
-#23772 := [trans #23771 #13252]: #23647
-#23642 := (f85 f90 #23197)
-#23643 := (f7 #23642 #23477)
-#23644 := (= #23643 f1)
-#23757 := (= #23643 #4665)
-#23758 := (= #23642 #4664)
-#23759 := [monotonicity #23716]: #23758
-#23760 := [monotonicity #23759 #23720]: #23757
-#23761 := [trans #23760 #13251]: #23644
-#23648 := (not #23647)
-#23645 := (not #23644)
-#23786 := (or #23645 #23648 #23650 #23573)
-#23483 := (f48 #4661 #23197)
-#23625 := (= #23483 f51)
-#13250 := [and-elim #13248]: #4663
-#23762 := (= #23483 #4662)
-#23778 := [monotonicity #23716]: #23762
-#23779 := [trans #23778 #13250]: #23625
-#13249 := [and-elim #13248]: #4660
-#23777 := (= #23488 #4659)
-#23780 := [monotonicity #23716]: #23777
-#23781 := [trans #23780 #13249]: #23489
-#23478 := (f7 f45 #23477)
-#23481 := (= #23478 f1)
-#13254 := [and-elim #13248]: #4674
-#23782 := (= #23478 #4673)
-#23787 := [monotonicity #23720]: #23782
-#23788 := [trans #23787 #13254]: #23481
-#13256 := [not-or-elim #13236]: #4687
-#13258 := [and-elim #13256]: #4686
-#74 := (:var 1 S9)
-#2793 := (f308 f310 #74)
-#2794 := (f92 #2793 #40)
-#2795 := (f37 #2794 #40)
-#2796 := (pattern #2795)
-#2803 := (= #2795 f1)
-#44 := (f27 f28 #40)
-#110 := (f7 f45 #44)
-#111 := (= #110 f1)
-#9649 := (not #111)
-#45 := (f26 #44)
-#47 := (= #45 f29)
-#222 := (f46 f47 #74)
-#223 := (f37 #222 #40)
-#224 := (= #223 f1)
-#9775 := (not #224)
-#219 := (f85 f90 #40)
-#220 := (f7 #219 #44)
-#221 := (= #220 f1)
-#17712 := (not #221)
-#215 := (f49 f50 #74)
-#216 := (f48 #215 #40)
-#217 := (= #216 f51)
-#17711 := (not #217)
-#205 := (f46 f52 #74)
-#206 := (f37 #205 #40)
-#207 := (= #206 f1)
-#3098 := (not #207)
-#326 := (f82 f98 #74)
-#328 := (= #326 f1)
-#17838 := (not #328)
-#20291 := (or #17838 #3098 #17711 #17712 #9775 #47 #9649 #2803)
-#20296 := (forall (vars (?v0 S9) (?v1 S10)) (:pat #2796) #20291)
-#81 := (not #47)
-#9361 := (and #328 #207 #217 #221 #224 #81 #111)
-#9364 := (not #9361)
-#9367 := (or #9364 #2803)
-#9370 := (forall (vars (?v0 S9) (?v1 S10)) (:pat #2796) #9367)
-#20297 := (iff #9370 #20296)
-#20294 := (iff #9367 #20291)
-#20277 := (or #17838 #3098 #17711 #17712 #9775 #47 #9649)
-#20288 := (or #20277 #2803)
-#20292 := (iff #20288 #20291)
-#20293 := [rewrite]: #20292
-#20289 := (iff #9367 #20288)
-#20286 := (iff #9364 #20277)
-#20278 := (not #20277)
-#20281 := (not #20278)
-#20284 := (iff #20281 #20277)
-#20285 := [rewrite]: #20284
-#20282 := (iff #9364 #20281)
-#20279 := (iff #9361 #20278)
-#20280 := [rewrite]: #20279
-#20283 := [monotonicity #20280]: #20282
-#20287 := [trans #20283 #20285]: #20286
-#20290 := [monotonicity #20287]: #20289
-#20295 := [trans #20290 #20293]: #20294
-#20298 := [quant-intro #20295]: #20297
-#16317 := (~ #9370 #9370)
-#16315 := (~ #9367 #9367)
-#16316 := [refl]: #16315
-#16318 := [nnf-pos #16316]: #16317
-#2797 := (and #81 #111)
-#2798 := (and #224 #2797)
-#2799 := (and #221 #2798)
-#2800 := (and #217 #2799)
-#2801 := (and #207 #2800)
-#2802 := (and #328 #2801)
-#2804 := (implies #2802 #2803)
-#2805 := (forall (vars (?v0 S9) (?v1 S10)) (:pat #2796) #2804)
-#9373 := (iff #2805 #9370)
-#9353 := (not #2802)
-#9355 := (or #9353 #2803)
-#9358 := (forall (vars (?v0 S9) (?v1 S10)) (:pat #2796) #9355)
-#9371 := (iff #9358 #9370)
-#9368 := (iff #9355 #9367)
-#9365 := (iff #9353 #9364)
-#9362 := (iff #2802 #9361)
-#9363 := [rewrite]: #9362
-#9366 := [monotonicity #9363]: #9365
-#9369 := [monotonicity #9366]: #9368
-#9372 := [quant-intro #9369]: #9371
-#9359 := (iff #2805 #9358)
-#9356 := (iff #2804 #9355)
-#9357 := [rewrite]: #9356
-#9360 := [quant-intro #9357]: #9359
-#9374 := [trans #9360 #9372]: #9373
-#9352 := [asserted]: #2805
-#9375 := [mp #9352 #9374]: #9370
-#16319 := [mp~ #9375 #16318]: #9370
-#20299 := [mp #16319 #20298]: #20296
-#23482 := (not #23481)
-#23626 := (not #23625)
-#22212 := (not #4686)
-#23630 := (not #20296)
-#23628 := (or #23630 #22212 #23514 #23626 #23645 #23648 #23650 #23482 #23573)
-#23651 := (or #22212 #23514 #23626 #23645 #23648 #23650 #23482 #23573)
-#23631 := (or #23630 #23651)
-#23633 := (iff #23631 #23628)
-#23654 := [rewrite]: #23633
-#23632 := [quant-inst #4649 #23197]: #23631
-#23656 := [mp #23632 #23654]: #23628
-#23789 := [unit-resolution #23656 #20299 #13258 #23788 #23781 #23779]: #23786
-#23790 := [unit-resolution #23789 #23761 #23772 #23769 #23741]: false
-#23791 := [lemma #23790]: #4733
-#24205 := [trans #23750 #23791]: #23569
-#23578 := (or #23587 #23584 #23573)
-#23579 := [def-axiom]: #23578
-#24206 := [unit-resolution #23579 #24205 #23729]: #23573
-#23624 := (or #23581 #23623)
-#26 := (:var 2 S9)
-#2724 := (f308 f310 #26)
-#2725 := (f92 #2724 #28)
-#2726 := (f37 #2725 #40)
-#2806 := (pattern #2726)
-#2766 := (f32 f96 #24)
-#2807 := (f31 #2766 #37)
-#2808 := (f15 #49 #2807)
-#2809 := (pattern #2808)
-#36 := (f24 f25 #24)
-#2711 := (f23 #36 #28)
-#2816 := (f303 f304 #2711)
-#2817 := (f15 #2815 #2816)
-#2818 := (= #2817 f1)
-#2813 := (= #2808 f1)
-#20300 := (not #2813)
-#1680 := (f27 f28 #37)
-#2810 := (f7 f311 #1680)
-#2811 := (= #2810 f1)
-#20315 := (or #2811 #20300 #2818)
-#20320 := (forall (vars (?v3 S10)) (:pat #2809) #20315)
-#20326 := (not #20320)
-#265 := (f46 f52 #26)
-#266 := (f37 #265 #28)
-#271 := (= #266 f1)
-#2983 := (not #271)
-#66 := (f32 f33 #26)
-#2698 := (f31 #66 #40)
-#1994 := (f16 f30 #28)
-#2699 := (f15 #1994 #2698)
-#2704 := (= #2699 f1)
-#20135 := (not #2704)
-#20327 := (or #20135 #2983 #20326)
-#20328 := (not #20327)
-#2727 := (= #2726 f1)
-#9393 := (not #2727)
-#20333 := (or #9393 #20328)
-#20336 := (forall (vars (?v0 S9) (?v1 S10) (?v2 S10)) (:pat #2806) #20333)
-#2812 := (not #2811)
-#2814 := (and #2812 #2813)
-#9377 := (not #2814)
-#9378 := (or #9377 #2818)
-#9381 := (forall (vars (?v3 S10)) (:pat #2809) #9378)
-#9402 := (and #2704 #271 #9381)
-#9405 := (or #9393 #9402)
-#9408 := (forall (vars (?v0 S9) (?v1 S10) (?v2 S10)) (:pat #2806) #9405)
-#20337 := (iff #9408 #20336)
-#20334 := (iff #9405 #20333)
-#20331 := (iff #9402 #20328)
-#20323 := (and #2704 #271 #20320)
-#20329 := (iff #20323 #20328)
-#20330 := [rewrite]: #20329
-#20324 := (iff #9402 #20323)
-#20321 := (iff #9381 #20320)
-#20318 := (iff #9378 #20315)
-#20301 := (or #2811 #20300)
-#20312 := (or #20301 #2818)
-#20316 := (iff #20312 #20315)
-#20317 := [rewrite]: #20316
-#20313 := (iff #9378 #20312)
-#20310 := (iff #9377 #20301)
-#20302 := (not #20301)
-#20305 := (not #20302)
-#20308 := (iff #20305 #20301)
-#20309 := [rewrite]: #20308
-#20306 := (iff #9377 #20305)
-#20303 := (iff #2814 #20302)
-#20304 := [rewrite]: #20303
-#20307 := [monotonicity #20304]: #20306
-#20311 := [trans #20307 #20309]: #20310
-#20314 := [monotonicity #20311]: #20313
-#20319 := [trans #20314 #20317]: #20318
-#20322 := [quant-intro #20319]: #20321
-#20325 := [monotonicity #20322]: #20324
-#20332 := [trans #20325 #20330]: #20331
-#20335 := [monotonicity #20332]: #20334
-#20338 := [quant-intro #20335]: #20337
-#16334 := (~ #9408 #9408)
-#16332 := (~ #9405 #9405)
-#16330 := (~ #9402 #9402)
-#16328 := (~ #9381 #9381)
-#16326 := (~ #9378 #9378)
-#16327 := [refl]: #16326
-#16329 := [nnf-pos #16327]: #16328
-#16324 := (~ #271 #271)
-#16325 := [refl]: #16324
-#16322 := (~ #2704 #2704)
-#16323 := [refl]: #16322
-#16331 := [monotonicity #16323 #16325 #16329]: #16330
-#16320 := (~ #9393 #9393)
-#16321 := [refl]: #16320
-#16333 := [monotonicity #16321 #16331]: #16332
-#16335 := [nnf-pos #16333]: #16334
-#2819 := (implies #2814 #2818)
-#2820 := (forall (vars (?v3 S10)) (:pat #2809) #2819)
-#2821 := (and #271 #2820)
-#2822 := (and #2704 #2821)
-#2823 := (implies #2727 #2822)
-#2824 := (forall (vars (?v0 S9) (?v1 S10) (?v2 S10)) (:pat #2806) #2823)
-#9411 := (iff #2824 #9408)
-#9384 := (and #271 #9381)
-#9387 := (and #2704 #9384)
-#9394 := (or #9393 #9387)
-#9399 := (forall (vars (?v0 S9) (?v1 S10) (?v2 S10)) (:pat #2806) #9394)
-#9409 := (iff #9399 #9408)
-#9406 := (iff #9394 #9405)
-#9403 := (iff #9387 #9402)
-#9404 := [rewrite]: #9403
-#9407 := [monotonicity #9404]: #9406
-#9410 := [quant-intro #9407]: #9409
-#9400 := (iff #2824 #9399)
-#9397 := (iff #2823 #9394)
-#9390 := (implies #2727 #9387)
-#9395 := (iff #9390 #9394)
-#9396 := [rewrite]: #9395
-#9391 := (iff #2823 #9390)
-#9388 := (iff #2822 #9387)
-#9385 := (iff #2821 #9384)
-#9382 := (iff #2820 #9381)
-#9379 := (iff #2819 #9378)
-#9380 := [rewrite]: #9379
-#9383 := [quant-intro #9380]: #9382
-#9386 := [monotonicity #9383]: #9385
-#9389 := [monotonicity #9386]: #9388
-#9392 := [monotonicity #9389]: #9391
-#9398 := [trans #9392 #9396]: #9397
-#9401 := [quant-intro #9398]: #9400
-#9412 := [trans #9401 #9410]: #9411
-#9376 := [asserted]: #2824
-#9413 := [mp #9376 #9412]: #9408
-#16336 := [mp~ #9413 #16335]: #9408
-#20339 := [mp #16336 #20338]: #20336
-#23600 := (not #20336)
-#23601 := (or #23600 #23581 #23623)
-#23597 := (or #23600 #23624)
-#23634 := (iff #23597 #23601)
-#23635 := [rewrite]: #23634
-#23602 := [quant-inst #4649 #23197 #23197]: #23597
-#23657 := [mp #23602 #23635]: #23601
-#24193 := [unit-resolution #23657 #20339]: #23624
-#24194 := [unit-resolution #24193 #24206]: #23623
-#23658 := (or #23622 #23605)
-#23659 := [def-axiom]: #23658
-#24192 := [unit-resolution #23659 #24194]: #23605
-#24197 := (= #23870 #23604)
-#24195 := (= #23869 #23594)
-#24769 := (= #23868 #23197)
-#24767 := (= #23868 #4658)
-#24765 := (= f445 #4657)
-#23271 := (= #4657 f445)
-#3412 := (f20 f179 #3411)
-#3413 := (= #3412 #243)
-#21608 := (forall (vars (?v0 S6) (?v1 Int)) (:pat #21607) #3413)
-#3414 := (forall (vars (?v0 S6) (?v1 Int)) #3413)
-#21611 := (iff #3414 #21608)
-#21609 := (iff #3413 #3413)
-#21610 := [refl]: #21609
-#21612 := [quant-intro #21610]: #21611
-#16694 := (~ #3414 #3414)
-#16692 := (~ #3413 #3413)
-#16693 := [refl]: #16692
-#16695 := [nnf-pos #16693]: #16694
-#9823 := [asserted]: #3414
-#16696 := [mp~ #9823 #16695]: #3414
-#21613 := [mp #16696 #21612]: #21608
-#23244 := (not #21608)
-#23276 := (or #23244 #23271)
-#23277 := [quant-inst #20 #4655]: #23276
-#24764 := [unit-resolution #23277 #21613]: #23271
-#24766 := [symm #24764]: #24765
-#24768 := [monotonicity #24766]: #24767
-#24770 := [trans #24768 #24763]: #24769
-#24196 := [monotonicity #24770]: #24195
-#24198 := [monotonicity #24196]: #24197
-#24199 := [trans #24198 #24192]: #23871
-#23872 := (not #23871)
-#24204 := [hypothesis]: #23872
-#24200 := [unit-resolution #24204 #24199]: false
-#24213 := [lemma #24200]: #23871
-#20851 := (not #12426)
-#21753 := (or #20851 #12612 #11767 #11758 #12561 #20811 #21720)
-#21756 := (not #21753)
-#21735 := (or #16955 #16958 #21732)
-#21738 := (not #21735)
-#21741 := (or #16955 #16958 #21738)
-#21744 := (not #21741)
-#21747 := (or #12561 #20851 #12613 #21744)
-#21750 := (not #21747)
-#21759 := (or #21750 #21756)
-#21762 := (not #21759)
-#21765 := (or #16955 #16964 #12561 #20851 #21762)
-#21768 := (not #21765)
-#21771 := (or #16955 #16964 #21768)
-#21774 := (not #21771)
-#21777 := (or #16955 #16958 #21774)
-#21780 := (not #21777)
-#21783 := (or #16955 #16958 #21780)
-#21786 := (not #21783)
-#21789 := (or #12561 #20851 #12706 #21786)
-#21792 := (not #21789)
-#20942 := (not #4826)
-#20943 := (or #6155 #17965 #12734 #20942)
-#21803 := (forall (vars (?v0 Int)) (:pat #21662) #20943)
-#21808 := (not #21803)
-#20934 := (or #6155 #17965 #12734 #12748)
-#21795 := (forall (vars (?v0 Int)) (:pat #21662) #20934)
-#21800 := (not #21795)
-#21811 := (or #21800 #21808)
-#21814 := (not #21811)
-decl ?v0!15 :: Int
-#17138 := ?v0!15
-#17145 := (f107 #4734 ?v0!15)
-#17146 := (f106 #17145 f14)
-#17147 := (f20 #4748 #17146)
-#17462 := (* -1::Int #17147)
-#17463 := (+ f468 #17462)
-#17464 := (>= #17463 0::Int)
-#17449 := (* -1::Int ?v0!15)
-#17450 := (+ f443 #17449)
-#17451 := (<= #17450 0::Int)
-#17140 := (<= ?v0!15 4294967295::Int)
-#20908 := (not #17140)
-#17139 := (>= ?v0!15 0::Int)
-#20907 := (not #17139)
-#20923 := (or #20907 #20908 #17451 #17464)
-#20928 := (not #20923)
-#21817 := (or #20928 #21814)
-#21820 := (not #21817)
-#21823 := (or #12707 #12561 #20851 #11347 #11338 #11329 #11320 #21820)
-#21826 := (not #21823)
-#21829 := (or #21792 #21826)
-#21832 := (not #21829)
-#20995 := (not #4923)
-#20994 := (not #4918)
-#14815 := (not #4811)
-#20993 := (not #4806)
-#20726 := (or #6155 #17965 #12889 #12903)
-#21671 := (forall (vars (?v0 Int)) (:pat #21662) #20726)
-#21676 := (not #21671)
-#13535 := (<= f464 4294967295::Int)
-#20991 := (not #13535)
-#20990 := (not #12929)
-#13550 := (<= f463 4294967295::Int)
-#20989 := (not #13550)
-#1613 := 255::Int
-#13569 := (<= f462 255::Int)
-#20988 := (not #13569)
-#20987 := (not #12951)
-#16901 := (not #4780)
-#21835 := (or #12418 #16901 #20987 #20988 #20989 #20990 #20991 #12561 #20851 #12926 #21676 #12879 #20993 #12874 #14815 #11970 #11961 #11952 #11943 #20994 #20995 #21832)
-#21838 := (not #21835)
-#25425 := (iff #4750 #4780)
-#25421 := (iff #4780 #4750)
-#25422 := [commutativity]: #25421
-#25426 := [symm #25422]: #25425
-#21841 := (or #12418 #16901 #21838)
-#21844 := (not #21841)
-#20715 := (or #6155 #17965 #12385 #12397)
-#21663 := (forall (vars (?v0 Int)) (:pat #21662) #20715)
-#21668 := (not #21663)
-#21847 := (or #21668 #21844)
-#21850 := (not #21847)
-decl ?v0!13 :: Int
-#16874 := ?v0!13
-#16880 := (f107 #4734 ?v0!13)
-#16881 := (f106 #16880 f14)
-#16882 := (f20 #4748 #16881)
-#16883 := (* -1::Int #16882)
-#16884 := (+ f461 #16883)
-#16885 := (>= #16884 0::Int)
-#16879 := (>= ?v0!13 1::Int)
-#16876 := (<= ?v0!13 4294967295::Int)
-#20689 := (not #16876)
-#16875 := (>= ?v0!13 0::Int)
-#20688 := (not #16875)
-#20704 := (or #20688 #20689 #16879 #16885)
-#20709 := (not #20704)
-#21853 := (or #20709 #21850)
-#21856 := (not #21853)
-#21859 := (or #12382 #21856)
-#21862 := (not #21859)
-#21865 := (or #12382 #21862)
-#21868 := (not #21865)
-#16851 := (not #4745)
-#16842 := (not #4739)
-#21871 := (or #16842 #16851 #12177 #12168 #12159 #12150 #21868)
-#21874 := (not #21871)
-#23983 := (f37 #23974 #23775)
-#23984 := (= #23983 f1)
-#23981 := (f48 #4661 #23775)
-#23982 := (= #23981 f51)
-#23985 := (or #23982 #23984)
-#23986 := (not #23985)
-#23955 := (f27 f28 #23775)
-#23956 := (f26 #23955)
-#23957 := (= #23956 f29)
-#23987 := (or #23957 #23986)
-#23988 := (not #23987)
-#23959 := (f101 #4876 #23775)
-#23963 := (f208 f209 #23959)
-#23975 := (f37 #23974 #23963)
-#23976 := (= #23975 f1)
-#23972 := (f48 #4661 #23963)
-#23973 := (= #23972 f51)
-#23977 := (or #23973 #23976)
-#23978 := (not #23977)
-#23969 := (f27 f28 #23963)
-#23970 := (f26 #23969)
-#23971 := (= #23970 f29)
-#23964 := (f37 #4650 #23963)
-#23965 := (= #23964 f1)
-#23966 := (not #23965)
-#23960 := (f122 f210 #23959)
-#23961 := (= #23960 f1)
-#23962 := (not #23961)
-#23967 := (or #23962 #23966)
-#23968 := (not #23967)
-#23958 := (not #23957)
-#23979 := (or #23958 #23968 #23971 #23978)
-#23980 := (not #23979)
-#23989 := (or #23980 #23988)
-#23990 := (not #23989)
-#23952 := (f37 #4667 #23775)
-#23953 := (= #23952 f1)
-#23747 := (f101 #4876 #4736)
-#23877 := (f122 f210 #23747)
-#23878 := (= #23877 f1)
-#16845 := (not #4741)
-#23879 := (or #16845 #23878)
-#23880 := (not #23879)
-#24214 := [hypothesis]: #23879
-#13247 := [not-or-elim #13236]: #12419
-decl f44 :: S5
-#106 := f44
-#3534 := (f7 f44 f14)
-#3535 := (= #3534 f1)
-#9865 := [asserted]: #3535
-#337 := (f7 f44 #333)
-#352 := (:var 1 Int)
-#969 := (:var 4 Int)
-#2608 := (f53 #375 #969)
-#2609 := (f120 f121 #2608)
-#2610 := (f107 #2609 #352)
-#2611 := (f106 #2610 #333)
-#1666 := (:var 5 S9)
-#2629 := (f49 f50 #1666)
-#2630 := (f48 #2629 #2611)
-#1656 := (:var 3 S10)
-#2605 := (f32 f33 #1666)
-#2606 := (f31 #2605 #1656)
-#382 := (:var 2 Int)
-#2600 := (f189 f190 #333)
-#2601 := (f188 #2600 #382)
-#2602 := (f80 f81 #2601)
-#2603 := (f53 #2602 #969)
-#2604 := (f16 f30 #2603)
-#2607 := (f15 #2604 #2606)
-#2631 := (pattern #2607 #2630 #337)
-#1910 := (f104 f105 #1666)
-#1911 := (f102 f103 #1910)
-#2627 := (f101 #1911 #2611)
-#2628 := (pattern #2607 #2627 #337)
-#2634 := (f122 f210 #2627)
-#2635 := (= #2634 f1)
-#1923 := (f46 f47 #1666)
-#2632 := (f37 #1923 #2611)
-#2633 := (= #2632 f1)
-#20045 := (not #2633)
-#20046 := (or #20045 #2635)
-#20047 := (not #20046)
-#5437 := (* -1::Int #382)
-#6983 := (+ #352 #5437)
-#7537 := (>= #6983 0::Int)
-#5569 := (>= #352 0::Int)
-#17932 := (not #5569)
-#2616 := (= #2607 f1)
-#20021 := (not #2616)
-#338 := (= #337 f1)
-#9488 := (not #338)
-#2614 := (f82 f98 #1666)
-#2615 := (= #2614 f1)
-#20020 := (not #2615)
-#20053 := (or #20020 #9488 #20021 #17932 #7537 #20047)
-#20058 := (forall (vars (?v0 S9) (?v1 Int) (?v2 S10) (?v3 Int) (?v4 Int) (?v5 S6)) (:pat #2628 #2631) #20053)
-#2636 := (not #2635)
-#2637 := (and #2633 #2636)
-#8302 := (not #7537)
-#9127 := (and #2615 #338 #2616 #5569 #8302)
-#9132 := (not #9127)
-#9151 := (or #9132 #2637)
-#9154 := (forall (vars (?v0 S9) (?v1 Int) (?v2 S10) (?v3 Int) (?v4 Int) (?v5 S6)) (:pat #2628 #2631) #9151)
-#20059 := (iff #9154 #20058)
-#20056 := (iff #9151 #20053)
-#20022 := (or #20020 #9488 #20021 #17932 #7537)
-#20050 := (or #20022 #20047)
-#20054 := (iff #20050 #20053)
-#20055 := [rewrite]: #20054
-#20051 := (iff #9151 #20050)
-#20048 := (iff #2637 #20047)
-#20049 := [rewrite]: #20048
-#20031 := (iff #9132 #20022)
-#20023 := (not #20022)
-#20026 := (not #20023)
-#20029 := (iff #20026 #20022)
-#20030 := [rewrite]: #20029
-#20027 := (iff #9132 #20026)
-#20024 := (iff #9127 #20023)
-#20025 := [rewrite]: #20024
-#20028 := [monotonicity #20025]: #20027
-#20032 := [trans #20028 #20030]: #20031
-#20052 := [monotonicity #20032 #20049]: #20051
-#20057 := [trans #20052 #20055]: #20056
-#20060 := [quant-intro #20057]: #20059
-#16225 := (~ #9154 #9154)
-#16223 := (~ #9151 #9151)
-#16224 := [refl]: #16223
-#16226 := [nnf-pos #16224]: #16225
-#1758 := (< #352 #382)
-#553 := (<= 0::Int #352)
-#1759 := (and #553 #1758)
-#2617 := (and #2616 #1759)
-#2618 := (and #338 #2617)
-#2619 := (and #2615 #2618)
-#2638 := (implies #2619 #2637)
-#2639 := (forall (vars (?v0 S9) (?v1 Int) (?v2 S10) (?v3 Int) (?v4 Int) (?v5 S6)) (:pat #2628 #2631) #2638)
-#9157 := (iff #2639 #9154)
-#9111 := (not #2619)
-#9145 := (or #9111 #2637)
-#9148 := (forall (vars (?v0 S9) (?v1 Int) (?v2 S10) (?v3 Int) (?v4 Int) (?v5 S6)) (:pat #2628 #2631) #9145)
-#9155 := (iff #9148 #9154)
-#9152 := (iff #9145 #9151)
-#9133 := (iff #9111 #9132)
-#9130 := (iff #2619 #9127)
-#8305 := (and #5569 #8302)
-#9118 := (and #2616 #8305)
-#9121 := (and #338 #9118)
-#9124 := (and #2615 #9121)
-#9128 := (iff #9124 #9127)
-#9129 := [rewrite]: #9128
-#9125 := (iff #2619 #9124)
-#9122 := (iff #2618 #9121)
-#9119 := (iff #2617 #9118)
-#8306 := (iff #1759 #8305)
-#8303 := (iff #1758 #8302)
-#8304 := [rewrite]: #8303
-#5568 := (iff #553 #5569)
-#5570 := [rewrite]: #5568
-#8307 := [monotonicity #5570 #8304]: #8306
-#9120 := [monotonicity #8307]: #9119
-#9123 := [monotonicity #9120]: #9122
-#9126 := [monotonicity #9123]: #9125
-#9131 := [trans #9126 #9129]: #9130
-#9134 := [monotonicity #9131]: #9133
-#9153 := [monotonicity #9134]: #9152
-#9156 := [quant-intro #9153]: #9155
-#9149 := (iff #2639 #9148)
-#9146 := (iff #2638 #9145)
-#9147 := [rewrite]: #9146
-#9150 := [quant-intro #9147]: #9149
-#9158 := [trans #9150 #9156]: #9157
-#9144 := [asserted]: #2639
-#9159 := [mp #9144 #9158]: #9154
-#16227 := [mp~ #9159 #16226]: #9154
-#20061 := [mp #16227 #20060]: #20058
-#22593 := (not #3535)
-#24132 := (not #20058)
-#24133 := (or #24132 #22212 #22593 #23872 #12418 #23880)
-#23875 := (+ 0::Int #12352)
-#23876 := (>= #23875 0::Int)
-#23873 := (>= 0::Int 0::Int)
-#23874 := (not #23873)
-#23881 := (or #22212 #22593 #23872 #23874 #23876 #23880)
-#24156 := (or #24132 #23881)
-#24149 := (iff #24156 #24133)
-#23900 := (or #22212 #22593 #23872 #12418 #23880)
-#24201 := (or #24132 #23900)
-#24128 := (iff #24201 #24133)
-#24148 := [rewrite]: #24128
-#24202 := (iff #24156 #24201)
-#23903 := (iff #23881 #23900)
-#23897 := (or #22212 #22593 #23872 false #12418 #23880)
-#23901 := (iff #23897 #23900)
-#23902 := [rewrite]: #23901
-#23898 := (iff #23881 #23897)
-#23895 := (iff #23876 #12418)
-#23890 := (>= #12352 0::Int)
-#23893 := (iff #23890 #12418)
-#23894 := [rewrite]: #23893
-#23891 := (iff #23876 #23890)
-#23888 := (= #23875 #12352)
-#23889 := [rewrite]: #23888
-#23892 := [monotonicity #23889]: #23891
-#23896 := [trans #23892 #23894]: #23895
-#23886 := (iff #23874 false)
-#23884 := (iff #23874 #4808)
-#23882 := (iff #23873 true)
-#23883 := [rewrite]: #23882
-#23885 := [monotonicity #23883]: #23884
-#23887 := [trans #23885 #11287]: #23886
-#23899 := [monotonicity #23887 #23896]: #23898
-#23904 := [trans #23899 #23902]: #23903
-#24203 := [monotonicity #23904]: #24202
-#24150 := [trans #24203 #24148]: #24149
-#24157 := [quant-inst #4649 #4655 #23197 #4646 #184 #20]: #24156
-#24151 := [mp #24157 #24150]: #24133
-#24236 := [unit-resolution #24151 #20061 #9865 #13247 #13258 #24213 #24214]: false
-#24237 := [lemma #24236]: #23880
-#24109 := (or #23879 #4741)
-#24110 := [def-axiom]: #24109
-#24858 := [unit-resolution #24110 #24237]: #4741
-#24885 := (= #23952 #4740)
-#24881 := (= #23775 #4736)
-#23776 := (= #4736 #23775)
-#23784 := (f27 f28 #4736)
-#23785 := (= #23784 f14)
-#23267 := (f27 f28 #4656)
-#23268 := (= #23267 f14)
-#23273 := (or #23239 #23268)
-#23274 := [quant-inst #20 #4655]: #23273
-#24238 := [unit-resolution #23274 #21619]: #23268
-#24269 := (= #23784 #23267)
-#24243 := (= #4736 #4656)
-#23825 := (f53 #4654 #4657)
-#24241 := (= #23825 #4656)
-#24242 := [monotonicity #24764]: #24241
-#23828 := (= #4736 #23825)
-#23831 := (not #23828)
-decl f216 :: S54
-#1950 := f216
-#23793 := (f92 f216 #4736)
-#23794 := (f37 #23793 #4656)
-#23807 := (= #23794 f1)
-#23808 := (not #23807)
-#23834 := (or #23808 #23831)
-#23837 := (not #23834)
-#1678 := (f120 f121 #37)
-#1679 := (f107 #1678 #352)
-#1948 := (f106 #1679 #333)
-#1949 := (pattern #1948)
-#1956 := (f217 f218 #333)
-#1957 := (* #352 #1956)
-#1954 := (f20 f179 #37)
-#1958 := (+ #1954 #1957)
-#1959 := (f53 #375 #1958)
-#1960 := (= #1948 #1959)
-#19589 := (not #1960)
-#1951 := (f92 f216 #1948)
-#1952 := (f37 #1951 #37)
-#1953 := (= #1952 f1)
-#19588 := (not #1953)
-#19590 := (or #19588 #19589)
-#19591 := (not #19590)
-#19594 := (forall (vars (?v0 S10) (?v1 Int) (?v2 S6)) (:pat #1949) #19591)
-#1961 := (and #1953 #1960)
-#1962 := (forall (vars (?v0 S10) (?v1 Int) (?v2 S6)) (:pat #1949) #1961)
-#19595 := (iff #1962 #19594)
-#19592 := (iff #1961 #19591)
-#19593 := [rewrite]: #19592
-#19596 := [quant-intro #19593]: #19595
-#15868 := (~ #1962 #1962)
-#15866 := (~ #1961 #1961)
-#15867 := [refl]: #15866
-#15869 := [nnf-pos #15867]: #15868
-#8597 := [asserted]: #1962
-#15870 := [mp~ #8597 #15869]: #1962
-#19597 := [mp #15870 #19596]: #19594
-#24083 := (not #19594)
-#24120 := (or #24083 #23837)
-#23809 := (* 0::Int #3690)
-#23810 := (+ #4657 #23809)
-#23811 := (f53 #4654 #23810)
-#23812 := (= #4736 #23811)
-#23813 := (not #23812)
-#23814 := (or #23808 #23813)
-#23815 := (not #23814)
-#24119 := (or #24083 #23815)
-#24121 := (iff #24119 #24120)
-#24085 := (iff #24120 #24120)
-#24123 := [rewrite]: #24085
-#23838 := (iff #23815 #23837)
-#23835 := (iff #23814 #23834)
-#23832 := (iff #23813 #23831)
-#23829 := (iff #23812 #23828)
-#23826 := (= #23811 #23825)
-#23823 := (= #23810 #4657)
-#23818 := (+ #4657 0::Int)
-#23821 := (= #23818 #4657)
-#23822 := [rewrite]: #23821
-#23819 := (= #23810 #23818)
-#23816 := (= #23809 0::Int)
-#23817 := [rewrite]: #23816
-#23820 := [monotonicity #23817]: #23819
-#23824 := [trans #23820 #23822]: #23823
-#23827 := [monotonicity #23824]: #23826
-#23830 := [monotonicity #23827]: #23829
-#23833 := [monotonicity #23830]: #23832
-#23836 := [monotonicity #23833]: #23835
-#23839 := [monotonicity #23836]: #23838
-#24122 := [monotonicity #23839]: #24121
-#24127 := [trans #24122 #24123]: #24121
-#24084 := [quant-inst #4656 #184 #20]: #24119
-#24076 := [mp #24084 #24127]: #24120
-#24239 := [unit-resolution #24076 #19597]: #23837
-#24078 := (or #23834 #23828)
-#24134 := [def-axiom]: #24078
-#24240 := [unit-resolution #24134 #24239]: #23828
-#24268 := [trans #24240 #24242]: #24243
-#24270 := [monotonicity #24268]: #24269
-#24271 := [trans #24270 #24238]: #23785
-#24086 := (not #23785)
-#23792 := (iff #4739 #23785)
-#1692 := (f27 f28 #28)
-#2360 := (= #1692 #333)
-#3378 := (iff #3373 #2360)
-#21601 := (forall (vars (?v0 S10) (?v1 S6)) (:pat #3372) #3378)
-#3379 := (forall (vars (?v0 S10) (?v1 S6)) #3378)
-#21604 := (iff #3379 #21601)
-#21602 := (iff #3378 #3378)
-#21603 := [refl]: #21602
-#21605 := [quant-intro #21603]: #21604
-#16679 := (~ #3379 #3379)
-#16677 := (~ #3378 #3378)
-#16678 := [refl]: #16677
-#16680 := [nnf-pos #16678]: #16679
-#9818 := [asserted]: #3379
-#16681 := [mp~ #9818 #16680]: #3379
-#21606 := [mp #16681 #21605]: #21601
-#23224 := (not #21601)
-#24118 := (or #23224 #23792)
-#24087 := [quant-inst #4736 #20]: #24118
-#24152 := [unit-resolution #24087 #21606]: #23792
-#24093 := (not #23792)
-#24142 := (or #24093 #24086)
-#24129 := [hypothesis]: #16842
-#24094 := (or #24093 #4739 #24086)
-#24107 := [def-axiom]: #24094
-#24145 := [unit-resolution #24107 #24129]: #24142
-#24235 := [unit-resolution #24145 #24152]: #24086
-#24272 := [unit-resolution #24235 #24271]: false
-#24267 := [lemma #24272]: #4739
-#24310 := (or #23214 #16842 #23776)
-#23783 := (or #16842 #23776)
-#24311 := (or #23214 #23783)
-#24313 := (iff #24311 #24310)
-#24314 := [rewrite]: #24313
-#24312 := [quant-inst #4736 #20]: #24311
-#24309 := [mp #24312 #24314]: #24310
-#24868 := [unit-resolution #24309 #16676 #24267]: #23776
-#24884 := [symm #24868]: #24881
-#24886 := [monotonicity #24884]: #24885
-#24888 := [trans #24886 #24858]: #23953
-#23954 := (not #23953)
-#23991 := (or #23954 #23990)
-#23992 := (not #23991)
-#23947 := (f37 #4743 #23775)
-#23948 := (= #23947 f1)
-#23993 := (iff #23948 #23992)
-#419 := (f46 f124 #74)
-#3010 := (f37 #419 #40)
-#3011 := (pattern #3010)
-#3029 := (f46 f332 #74)
-#3036 := (f37 #3029 #40)
-#3037 := (= #3036 f1)
-#3038 := (or #217 #3037)
-#20442 := (not #3038)
-#20443 := (or #47 #20442)
-#20444 := (not #20443)
-#405 := (f104 f105 #74)
-#406 := (f102 f103 #405)
-#3013 := (f101 #406 #40)
-#3017 := (f208 f209 #3013)
-#3030 := (f37 #3029 #3017)
-#3031 := (= #3030 f1)
-#3026 := (f48 #215 #3017)
-#3027 := (= #3026 f51)
-#3032 := (or #3027 #3031)
-#20437 := (not #3032)
-#3022 := (f27 f28 #3017)
-#3023 := (f26 #3022)
-#3024 := (= #3023 f29)
-#3018 := (f37 #205 #3017)
-#3019 := (= #3018 f1)
-#3020 := (not #3019)
-#3014 := (f122 f210 #3013)
-#3015 := (= #3014 f1)
-#3016 := (not #3015)
-#3021 := (or #3016 #3020)
-#20436 := (not #3021)
-#20438 := (or #81 #20436 #3024 #20437)
-#20439 := (not #20438)
-#20447 := (or #20439 #20444)
-#20453 := (not #20447)
-#20454 := (or #9775 #20453)
-#20455 := (not #20454)
-#3012 := (= #3010 f1)
-#20460 := (iff #3012 #20455)
-#20463 := (forall (vars (?v0 S9) (?v1 S10)) (:pat #3011) #20460)
-#3039 := (and #81 #3038)
-#3025 := (not #3024)
-#9561 := (and #47 #3021 #3025 #3032)
-#9564 := (or #9561 #3039)
-#9567 := (and #224 #9564)
-#9570 := (iff #3012 #9567)
-#9573 := (forall (vars (?v0 S9) (?v1 S10)) (:pat #3011) #9570)
-#20464 := (iff #9573 #20463)
-#20461 := (iff #9570 #20460)
-#20458 := (iff #9567 #20455)
-#20450 := (and #224 #20447)
-#20456 := (iff #20450 #20455)
-#20457 := [rewrite]: #20456
-#20451 := (iff #9567 #20450)
-#20448 := (iff #9564 #20447)
-#20445 := (iff #3039 #20444)
-#20446 := [rewrite]: #20445
-#20440 := (iff #9561 #20439)
-#20441 := [rewrite]: #20440
-#20449 := [monotonicity #20441 #20446]: #20448
-#20452 := [monotonicity #20449]: #20451
-#20459 := [trans #20452 #20457]: #20458
-#20462 := [monotonicity #20459]: #20461
-#20465 := [quant-intro #20462]: #20464
-#16439 := (~ #9573 #9573)
-#16437 := (~ #9570 #9570)
-#16438 := [refl]: #16437
-#16440 := [nnf-pos #16438]: #16439
-#3033 := (and #3025 #3032)
-#3034 := (and #3021 #3033)
-#3035 := (and #47 #3034)
-#3040 := (or #3035 #3039)
-#3041 := (and #224 #3040)
-#3042 := (iff #3012 #3041)
-#3043 := (forall (vars (?v0 S9) (?v1 S10)) (:pat #3011) #3042)
-#9574 := (iff #3043 #9573)
-#9571 := (iff #3042 #9570)
-#9568 := (iff #3041 #9567)
-#9565 := (iff #3040 #9564)
-#9562 := (iff #3035 #9561)
-#9563 := [rewrite]: #9562
-#9566 := [monotonicity #9563]: #9565
-#9569 := [monotonicity #9566]: #9568
-#9572 := [monotonicity #9569]: #9571
-#9575 := [quant-intro #9572]: #9574
-#9557 := [asserted]: #3043
-#9576 := [mp #9557 #9575]: #9573
-#16441 := [mp~ #9576 #16440]: #9573
-#20466 := [mp #16441 #20465]: #20463
-#24578 := (not #20463)
-#24587 := (or #24578 #23993)
-#24588 := [quant-inst #4649 #23775]: #24587
-#24576 := [unit-resolution #24588 #20466]: #23993
-#24328 := (not #23948)
-#24805 := (iff #16851 #24328)
-#24799 := (iff #4745 #23948)
-#24744 := (iff #23948 #4745)
-#24742 := (= #23947 #4744)
-#24743 := [monotonicity #24884]: #24742
-#24802 := [monotonicity #24743]: #24744
-#24800 := [symm #24802]: #24799
-#24806 := [monotonicity #24800]: #24805
-#24577 := [hypothesis]: #16851
-#24788 := [mp #24577 #24806]: #24328
-#24325 := (not #23993)
-#24326 := (or #24325 #23948 #23991)
-#24327 := [def-axiom]: #24326
-#24835 := [unit-resolution #24327 #24788 #24576]: #23991
-#24535 := (or #23992 #23954 #23990)
-#24324 := [def-axiom]: #24535
-#24836 := [unit-resolution #24324 #24835 #24888]: #23990
-#22576 := (f26 f14)
-#22577 := (= #22576 f29)
-#22584 := (iff #3535 #22577)
-#2909 := (pattern #337)
-#3604 := (= #3591 f29)
-#3605 := (iff #338 #3604)
-#3606 := (forall (vars (?v0 S6)) (:pat #2909) #3605)
-#16794 := (~ #3606 #3606)
-#16792 := (~ #3605 #3605)
-#16793 := [refl]: #16792
-#16795 := [nnf-pos #16793]: #16794
-#9915 := [asserted]: #3606
-#16796 := [mp~ #9915 #16795]: #3606
-#22308 := (not #3606)
-#22587 := (or #22308 #22584)
-#22588 := [quant-inst #20]: #22587
-#24801 := [unit-resolution #22588 #16796]: #22584
-#22589 := (not #22584)
-#24703 := (or #22589 #22577)
-#22594 := (or #22589 #22593 #22577)
-#22595 := [def-axiom]: #22594
-#24704 := [unit-resolution #22595 #9865]: #24703
-#24323 := [unit-resolution #24704 #24801]: #22577
-#24839 := (= #23956 #22576)
-#24847 := (= #23955 f14)
-#24833 := (or #24093 #23785)
-#24315 := (or #24093 #16842 #23785)
-#24316 := [def-axiom]: #24315
-#24834 := [unit-resolution #24316 #24267]: #24833
-#24837 := [unit-resolution #24834 #24152]: #23785
-#24838 := (= #23955 #23784)
-#24427 := [monotonicity #24884]: #24838
-#24848 := [trans #24427 #24837]: #24847
-#24840 := [monotonicity #24848]: #24839
-#24832 := [trans #24840 #24323]: #23957
-#24080 := (not #23878)
-#24854 := (iff #24080 #23962)
-#24515 := (iff #23878 #23961)
-#24841 := (iff #23961 #23878)
-#24432 := (= #23960 #23877)
-#25070 := (= #23959 #23747)
-#25071 := [monotonicity #24884]: #25070
-#24852 := [monotonicity #25071]: #24432
-#24831 := [monotonicity #24852]: #24841
-#24413 := [symm #24831]: #24515
-#24856 := [monotonicity #24413]: #24854
-#24081 := (or #23879 #24080)
-#24079 := [def-axiom]: #24081
-#24431 := [unit-resolution #24079 #24237]: #24080
-#24857 := [mp #24431 #24856]: #23962
-#24589 := (or #23967 #23961)
-#24590 := [def-axiom]: #24589
-#24892 := [unit-resolution #24590 #24857]: #23967
-#24901 := (or #23980 #23958 #23968)
-#24674 := (f20 f179 #23197)
-#25003 := (f53 #4654 #24674)
-#24977 := (f120 f121 #23197)
-#24978 := (f107 #24977 0::Int)
-#24985 := (f106 #24978 f14)
-#25006 := (= #24985 #25003)
-#25009 := (not #25006)
-#24986 := (f92 f216 #24985)
-#24987 := (f37 #24986 #23197)
-#24988 := (= #24987 f1)
-#24989 := (not #24988)
-#25012 := (or #24989 #25009)
-#25015 := (not #25012)
-#25121 := [hypothesis]: #25012
-#25018 := (or #24083 #25015)
-#24990 := (+ #24674 #23809)
-#24991 := (f53 #4654 #24990)
-#24992 := (= #24985 #24991)
-#24993 := (not #24992)
-#24994 := (or #24989 #24993)
-#24995 := (not #24994)
-#25019 := (or #24083 #24995)
-#25021 := (iff #25019 #25018)
-#25023 := (iff #25018 #25018)
-#25024 := [rewrite]: #25023
-#25016 := (iff #24995 #25015)
-#25013 := (iff #24994 #25012)
-#25010 := (iff #24993 #25009)
-#25007 := (iff #24992 #25006)
-#25004 := (= #24991 #25003)
-#25001 := (= #24990 #24674)
-#24996 := (+ #24674 0::Int)
-#24999 := (= #24996 #24674)
-#25000 := [rewrite]: #24999
-#24997 := (= #24990 #24996)
-#24998 := [monotonicity #23817]: #24997
-#25002 := [trans #24998 #25000]: #25001
-#25005 := [monotonicity #25002]: #25004
-#25008 := [monotonicity #25005]: #25007
-#25011 := [monotonicity #25008]: #25010
-#25014 := [monotonicity #25011]: #25013
-#25017 := [monotonicity #25014]: #25016
-#25022 := [monotonicity #25017]: #25021
-#25025 := [trans #25022 #25024]: #25021
-#25020 := [quant-inst #23197 #184 #20]: #25019
-#25026 := [mp #25020 #25025]: #25018
-#25122 := [unit-resolution #25026 #19597 #25121]: false
-#25123 := [lemma #25122]: #25015
-#25029 := (or #25012 #25006)
-#25030 := [def-axiom]: #25029
-#24893 := [unit-resolution #25030 #25123]: #25006
-#25119 := (or #25009 #23973)
-#25115 := (= #23972 #4662)
-#25082 := (= #23963 #4658)
-#25080 := (= #23963 #23868)
-#24904 := (f120 f121 #23868)
-#24905 := (f107 #24904 0::Int)
-#24906 := (f106 #24905 f14)
-#24907 := (f101 #4876 #24906)
-#24908 := (f208 f209 #24907)
-#24909 := (= #24908 #23868)
-#24916 := (f37 #4667 #24906)
-#24917 := (= #24916 f1)
-#24918 := (not #24917)
-decl f123 :: S69
-#412 := f123
-#24913 := (f122 f123 #24907)
-#24914 := (= #24913 f1)
-#24915 := (not #24914)
-#24911 := (f122 f210 #24907)
-#24912 := (= #24911 f1)
-#24910 := (not #24909)
-#24919 := (or #24910 #24912 #24915 #24918)
-#24920 := (not #24919)
-#24974 := [hypothesis]: #24919
-#24895 := (f37 #4667 #23868)
-#24896 := (= #24895 f1)
-#24966 := (= #24895 #4668)
-#24967 := [monotonicity #24768]: #24966
-#24968 := [trans #24967 #13252]: #24896
-#24903 := (not #24896)
-#24965 := [hypothesis]: #24903
-#24969 := [unit-resolution #24965 #24968]: false
-#24970 := [lemma #24969]: #24896
-#397 := (:var 3 Int)
-#444 := (:var 2 S6)
-#1810 := (f189 f190 #444)
-#1811 := (f188 #1810 #352)
-#1812 := (f80 f81 #1811)
-#1813 := (f53 #1812 #397)
-#1655 := (f80 f81 #444)
-#1805 := (f53 #1655 #397)
-#1806 := (f120 f121 #1805)
-#1807 := (f107 #1806 #243)
-#1808 := (f106 #1807 #444)
-#55 := (:var 4 S9)
-#1815 := (f104 f105 #55)
-#1816 := (f102 f103 #1815)
-#1817 := (f101 #1816 #1808)
-#1818 := (pattern #1817 #1813)
-#1803 := (f77 f78 #55)
-#1804 := (f75 f76 #1803)
-#1809 := (f74 #1804 #1808)
-#1814 := (pattern #1809 #1813)
-#1823 := (f120 f121 #1813)
-#1824 := (f107 #1823 #243)
-#1825 := (f106 #1824 #444)
-#1819 := (f46 f47 #55)
-#1835 := (f37 #1819 #1825)
-#1836 := (= #1835 f1)
-#19365 := (not #1836)
-#1826 := (f101 #1816 #1825)
-#1833 := (f122 f123 #1826)
-#1834 := (= #1833 f1)
-#19364 := (not #1834)
-#1830 := (f122 f210 #1826)
-#1831 := (= #1830 f1)
-#1827 := (f208 f209 #1826)
-#1828 := (= #1827 #1813)
-#19363 := (not #1828)
-#19366 := (or #19363 #1831 #19364 #19365)
-#19367 := (not #19366)
-#6377 := (* -1::Int #352)
-#6988 := (+ #243 #6377)
-#6989 := (>= #6988 0::Int)
-#1820 := (f37 #1819 #1813)
-#1821 := (= #1820 f1)
-#8411 := (not #1821)
-#19373 := (or #8411 #6155 #6989 #19367)
-#19378 := (forall (vars (?v0 S9) (?v1 Int) (?v2 S6) (?v3 Int) (?v4 Int)) (:pat #1814 #1818) #19373)
-#1832 := (not #1831)
-#8420 := (and #1828 #1832 #1834 #1836)
-#8206 := (not #6989)
-#8209 := (and #5433 #8206)
-#8212 := (not #8209)
-#8429 := (or #8411 #8212 #8420)
-#8434 := (forall (vars (?v0 S9) (?v1 Int) (?v2 S6) (?v3 Int) (?v4 Int)) (:pat #1814 #1818) #8429)
-#19379 := (iff #8434 #19378)
-#19376 := (iff #8429 #19373)
-#19248 := (or #6155 #6989)
-#19370 := (or #8411 #19248 #19367)
-#19374 := (iff #19370 #19373)
-#19375 := [rewrite]: #19374
-#19371 := (iff #8429 #19370)
-#19368 := (iff #8420 #19367)
-#19369 := [rewrite]: #19368
-#19257 := (iff #8212 #19248)
-#19249 := (not #19248)
-#19252 := (not #19249)
-#19255 := (iff #19252 #19248)
-#19256 := [rewrite]: #19255
-#19253 := (iff #8212 #19252)
-#19250 := (iff #8209 #19249)
-#19251 := [rewrite]: #19250
-#19254 := [monotonicity #19251]: #19253
-#19258 := [trans #19254 #19256]: #19257
-#19372 := [monotonicity #19258 #19369]: #19371
-#19377 := [trans #19372 #19375]: #19376
-#19380 := [quant-intro #19377]: #19379
-#15655 := (~ #8434 #8434)
-#15653 := (~ #8429 #8429)
-#15654 := [refl]: #15653
-#15656 := [nnf-pos #15654]: #15655
-#1837 := (and #1834 #1836)
-#1838 := (and #1832 #1837)
-#1839 := (and #1828 #1838)
-#1664 := (< #243 #352)
-#1665 := (and #409 #1664)
-#1840 := (implies #1665 #1839)
-#1841 := (implies #1821 #1840)
-#1842 := (forall (vars (?v0 S9) (?v1 Int) (?v2 S6) (?v3 Int) (?v4 Int)) (:pat #1814 #1818) #1841)
-#8437 := (iff #1842 #8434)
-#8178 := (not #1665)
-#8405 := (or #8178 #1839)
-#8412 := (or #8411 #8405)
-#8417 := (forall (vars (?v0 S9) (?v1 Int) (?v2 S6) (?v3 Int) (?v4 Int)) (:pat #1814 #1818) #8412)
-#8435 := (iff #8417 #8434)
-#8432 := (iff #8412 #8429)
-#8423 := (or #8212 #8420)
-#8426 := (or #8411 #8423)
-#8430 := (iff #8426 #8429)
-#8431 := [rewrite]: #8430
-#8427 := (iff #8412 #8426)
-#8424 := (iff #8405 #8423)
-#8421 := (iff #1839 #8420)
-#8422 := [rewrite]: #8421
-#8213 := (iff #8178 #8212)
-#8210 := (iff #1665 #8209)
-#8207 := (iff #1664 #8206)
-#8208 := [rewrite]: #8207
-#8211 := [monotonicity #5432 #8208]: #8210
-#8214 := [monotonicity #8211]: #8213
-#8425 := [monotonicity #8214 #8422]: #8424
-#8428 := [monotonicity #8425]: #8427
-#8433 := [trans #8428 #8431]: #8432
-#8436 := [quant-intro #8433]: #8435
-#8418 := (iff #1842 #8417)
-#8415 := (iff #1841 #8412)
-#8408 := (implies #1821 #8405)
-#8413 := (iff #8408 #8412)
-#8414 := [rewrite]: #8413
-#8409 := (iff #1841 #8408)
-#8406 := (iff #1840 #8405)
-#8407 := [rewrite]: #8406
-#8410 := [monotonicity #8407]: #8409
-#8416 := [trans #8410 #8414]: #8415
-#8419 := [quant-intro #8416]: #8418
-#8438 := [trans #8419 #8436]: #8437
-#8404 := [asserted]: #1842
-#8439 := [mp #8404 #8438]: #8434
-#15657 := [mp~ #8439 #15656]: #8434
-#19381 := [mp #15657 #19380]: #19378
-#24899 := (not #19378)
-#24897 := (or #24899 #24903 #12418 #24920)
-#24921 := (or #24903 #23874 #23876 #24920)
-#24931 := (or #24899 #24921)
-#24944 := (iff #24931 #24897)
-#24925 := (or #24903 #12418 #24920)
-#24938 := (or #24899 #24925)
-#24941 := (iff #24938 #24897)
-#24942 := [rewrite]: #24941
-#24939 := (iff #24931 #24938)
-#24928 := (iff #24921 #24925)
-#24922 := (or #24903 false #12418 #24920)
-#24926 := (iff #24922 #24925)
-#24927 := [rewrite]: #24926
-#24923 := (iff #24921 #24922)
-#24924 := [monotonicity #23887 #23896]: #24923
-#24929 := [trans #24924 #24927]: #24928
-#24940 := [monotonicity #24929]: #24939
-#24945 := [trans #24940 #24942]: #24944
-#24937 := [quant-inst #4649 #4655 #20 #4646 #184]: #24931
-#24946 := [mp #24937 #24945]: #24897
-#24960 := [unit-resolution #24946 #19381 #13247 #24970 #24974]: false
-#24961 := [lemma #24960]: #24920
-#24519 := (or #24919 #24909)
-#24506 := [def-axiom]: #24519
-#25103 := [unit-resolution #24506 #24961]: #24909
-#25078 := (= #23963 #24908)
-#25076 := (= #23959 #24907)
-#25074 := (= #23747 #24907)
-#25072 := (= #24907 #23747)
-#25068 := (= #24906 #4736)
-#25060 := (= #23825 #4736)
-#25061 := [symm #24240]: #25060
-#25066 := (= #24906 #23825)
-#25058 := (= #4656 #23825)
-#25059 := [monotonicity #24766]: #25058
-#25064 := (= #24906 #4656)
-#25056 := (= #25003 #4656)
-#25040 := (= #24674 f445)
-#25038 := (= #24674 #4657)
-#23272 := (= #22274 #4657)
-#23281 := (or #23244 #23272)
-#23282 := [quant-inst #4652 #4657]: #23281
-#25035 := [unit-resolution #23282 #21613]: #23272
-#25036 := (= #24674 #22274)
-#25037 := [monotonicity #23716]: #25036
-#25039 := [trans #25037 #25035]: #25038
-#25041 := [trans #25039 #24764]: #25040
-#25057 := [monotonicity #25041]: #25056
-#25062 := (= #24906 #25003)
-#25104 := [hypothesis]: #25006
-#25054 := (= #24906 #24985)
-#25052 := (= #24905 #24978)
-#25050 := (= #24978 #24905)
-#25048 := (= #24977 #24904)
-#25046 := (= #23197 #23868)
-#25044 := (= #4658 #23868)
-#25045 := [symm #24768]: #25044
-#25047 := [trans #23716 #25045]: #25046
-#25049 := [monotonicity #25047]: #25048
-#25051 := [monotonicity #25049]: #25050
-#25053 := [symm #25051]: #25052
-#25055 := [monotonicity #25053]: #25054
-#25105 := [trans #25055 #25104]: #25062
-#25106 := [trans #25105 #25057]: #25064
-#25107 := [trans #25106 #25059]: #25066
-#25108 := [trans #25107 #25061]: #25068
-#25109 := [monotonicity #25108]: #25072
-#25110 := [symm #25109]: #25074
-#25111 := [trans #25071 #25110]: #25076
-#25112 := [monotonicity #25111]: #25078
-#25113 := [trans #25112 #25103]: #25080
-#25114 := [trans #25113 #24768]: #25082
-#25116 := [monotonicity #25114]: #25115
-#25117 := [trans #25116 #13250]: #23973
-#24573 := (not #23973)
-#25102 := [hypothesis]: #24573
-#25118 := [unit-resolution #25102 #25117]: false
-#25120 := [lemma #25118]: #25119
-#24898 := [unit-resolution #25120 #24893]: #23973
-#24724 := (or #23977 #24573)
-#24725 := [def-axiom]: #24724
-#24900 := [unit-resolution #24725 #24898]: #23977
-#24729 := (not #23971)
-#24091 := (f208 f209 #23747)
-#24092 := (f27 f28 #24091)
-#24095 := (f26 #24092)
-#24096 := (= #24095 f29)
-#24803 := [hypothesis]: #23971
-#24789 := (= #24095 #23970)
-#24739 := (= #24092 #23969)
-#24618 := (= #24091 #23963)
-#24619 := (= #23747 #23959)
-#24785 := [symm #25071]: #24619
-#24738 := [monotonicity #24785]: #24618
-#24756 := [monotonicity #24738]: #24739
-#24790 := [monotonicity #24756]: #24789
-#24807 := [trans #24790 #24803]: #24096
-#24723 := (not #24096)
-#24097 := (f7 f45 #24092)
-#24098 := (= #24097 f1)
-#24099 := (not #24098)
-#24100 := (or #24096 #24099)
-#24101 := (not #24100)
-#3329 := (:var 0 S56)
-#3330 := (f208 f209 #3329)
-#3331 := (pattern #3330)
-#3332 := (f27 f28 #3330)
-#3336 := (f7 f45 #3332)
-#3337 := (= #3336 f1)
-#20620 := (not #3337)
-#3333 := (f26 #3332)
-#3334 := (= #3333 f29)
-#20621 := (or #3334 #20620)
-#20622 := (not #20621)
-#20625 := (forall (vars (?v0 S56)) (:pat #3331) #20622)
-#3335 := (not #3334)
-#3338 := (and #3335 #3337)
-#3339 := (forall (vars (?v0 S56)) (:pat #3331) #3338)
-#20626 := (iff #3339 #20625)
-#20623 := (iff #3338 #20622)
-#20624 := [rewrite]: #20623
-#20627 := [quant-intro #20624]: #20626
-#16654 := (~ #3339 #3339)
-#16652 := (~ #3338 #3338)
-#16653 := [refl]: #16652
-#16655 := [nnf-pos #16653]: #16654
-#9783 := [asserted]: #3339
-#16656 := [mp~ #9783 #16655]: #3339
-#20628 := [mp #16656 #20627]: #20625
-#24322 := (not #20625)
-#24734 := (or #24322 #24101)
-#24722 := [quant-inst #23747]: #24734
-#24571 := [unit-resolution #24722 #20628]: #24101
-#24786 := (or #24100 #24723)
-#24787 := [def-axiom]: #24786
-#24572 := [unit-resolution #24787 #24571]: #24723
-#24808 := [unit-resolution #24572 #24807]: false
-#24804 := [lemma #24808]: #24729
-#24412 := (or #23980 #23958 #23968 #23971 #23978)
-#24428 := [def-axiom]: #24412
-#24855 := [unit-resolution #24428 #24804 #24900]: #24901
-#24853 := [unit-resolution #24855 #24892 #24832]: #23980
-#23866 := (or #23989 #23979)
-#23867 := [def-axiom]: #23866
-#24514 := [unit-resolution #23867 #24853 #24836]: false
-#24533 := [lemma #24514]: #4745
-#25242 := (or #16851 #21874)
-#21877 := (or #16842 #16851 #21874)
-#21880 := (not #21877)
-#21883 := (or #16842 #16845 #21880)
-#21886 := (not #21883)
-#21889 := (or #16842 #16845 #21886)
-#21892 := (not #21889)
-#21895 := (or #12210 #21892)
-#21898 := (not #21895)
-#21901 := (or #12210 #21898)
-#20954 := (forall (vars (?v0 Int)) #20943)
-#20961 := (not #20954)
-#20939 := (forall (vars (?v0 Int)) #20934)
-#20960 := (not #20939)
-#20962 := (or #20960 #20961)
-#20963 := (not #20962)
-#20968 := (or #20928 #20963)
-#20974 := (not #20968)
-#20975 := (or #12707 #12561 #20851 #11347 #11338 #11329 #11320 #20974)
-#20976 := (not #20975)
-#20766 := (forall (vars (?v0 Int)) #20761)
-#20784 := (not #20766)
-#20785 := (or #20784 #20771)
-#20786 := (not #20785)
-#20791 := (or #20755 #20786)
-#20797 := (not #20791)
-#20798 := (or #12465 #20797)
-#20799 := (not #20798)
-#20804 := (or #12465 #20799)
-#20812 := (not #20804)
-#20813 := (or #16993 #16996 #12527 #11608 #20810 #20811 #20812)
-#20814 := (not #20813)
-#20819 := (or #16993 #16996 #20814)
-#20825 := (not #20819)
-#20862 := (or #20851 #12612 #11767 #11758 #12561 #20811 #20825)
-#20863 := (not #20862)
-#20826 := (or #16955 #16964 #11692 #11683 #11674 #11656 #11647 #12561 #20811 #20825)
-#20827 := (not #20826)
-#20832 := (or #16955 #16964 #20827)
-#20838 := (not #20832)
-#20839 := (or #16955 #16958 #20838)
-#20840 := (not #20839)
-#20845 := (or #16955 #16958 #20840)
-#20852 := (not #20845)
-#20853 := (or #12561 #20851 #12613 #20852)
-#20854 := (not #20853)
-#20868 := (or #20854 #20863)
-#20874 := (not #20868)
-#20875 := (or #16955 #16964 #12561 #20851 #20874)
-#20876 := (not #20875)
-#20881 := (or #16955 #16964 #20876)
-#20887 := (not #20881)
-#20888 := (or #16955 #16958 #20887)
-#20889 := (not #20888)
-#20894 := (or #16955 #16958 #20889)
-#20900 := (not #20894)
-#20901 := (or #12561 #20851 #12706 #20900)
-#20902 := (not #20901)
-#20981 := (or #20902 #20976)
-#20996 := (not #20981)
-#20731 := (forall (vars (?v0 Int)) #20726)
-#20992 := (not #20731)
-#20997 := (or #12418 #16901 #20987 #20988 #20989 #20990 #20991 #12561 #20851 #12926 #20992 #12879 #20993 #12874 #14815 #11970 #11961 #11952 #11943 #20994 #20995 #20996)
-#20998 := (not #20997)
-#21003 := (or #12418 #16901 #20998)
-#21010 := (not #21003)
-#20720 := (forall (vars (?v0 Int)) #20715)
-#21009 := (not #20720)
-#21011 := (or #21009 #21010)
-#21012 := (not #21011)
-#21017 := (or #20709 #21012)
-#21023 := (not #21017)
-#21024 := (or #12382 #21023)
-#21025 := (not #21024)
-#21030 := (or #12382 #21025)
-#21036 := (not #21030)
-#21037 := (or #16842 #16851 #12177 #12168 #12159 #12150 #21036)
-#21038 := (not #21037)
-#21043 := (or #16842 #16851 #21038)
-#21049 := (not #21043)
-#21050 := (or #16842 #16845 #21049)
-#21051 := (not #21050)
-#21056 := (or #16842 #16845 #21051)
-#21062 := (not #21056)
-#21063 := (or #12210 #21062)
-#21064 := (not #21063)
-#21069 := (or #12210 #21064)
-#21902 := (iff #21069 #21901)
-#21899 := (iff #21064 #21898)
-#21896 := (iff #21063 #21895)
-#21893 := (iff #21062 #21892)
-#21890 := (iff #21056 #21889)
-#21887 := (iff #21051 #21886)
-#21884 := (iff #21050 #21883)
-#21881 := (iff #21049 #21880)
-#21878 := (iff #21043 #21877)
-#21875 := (iff #21038 #21874)
-#21872 := (iff #21037 #21871)
-#21869 := (iff #21036 #21868)
-#21866 := (iff #21030 #21865)
-#21863 := (iff #21025 #21862)
-#21860 := (iff #21024 #21859)
-#21857 := (iff #21023 #21856)
-#21854 := (iff #21017 #21853)
-#21851 := (iff #21012 #21850)
-#21848 := (iff #21011 #21847)
-#21845 := (iff #21010 #21844)
-#21842 := (iff #21003 #21841)
-#21839 := (iff #20998 #21838)
-#21836 := (iff #20997 #21835)
-#21833 := (iff #20996 #21832)
-#21830 := (iff #20981 #21829)
-#21827 := (iff #20976 #21826)
-#21824 := (iff #20975 #21823)
-#21821 := (iff #20974 #21820)
-#21818 := (iff #20968 #21817)
-#21815 := (iff #20963 #21814)
-#21812 := (iff #20962 #21811)
-#21809 := (iff #20961 #21808)
-#21806 := (iff #20954 #21803)
-#21804 := (iff #20943 #20943)
-#21805 := [refl]: #21804
-#21807 := [quant-intro #21805]: #21806
-#21810 := [monotonicity #21807]: #21809
-#21801 := (iff #20960 #21800)
-#21798 := (iff #20939 #21795)
-#21796 := (iff #20934 #20934)
-#21797 := [refl]: #21796
-#21799 := [quant-intro #21797]: #21798
-#21802 := [monotonicity #21799]: #21801
-#21813 := [monotonicity #21802 #21810]: #21812
-#21816 := [monotonicity #21813]: #21815
-#21819 := [monotonicity #21816]: #21818
-#21822 := [monotonicity #21819]: #21821
-#21825 := [monotonicity #21822]: #21824
-#21828 := [monotonicity #21825]: #21827
-#21793 := (iff #20902 #21792)
-#21790 := (iff #20901 #21789)
-#21787 := (iff #20900 #21786)
-#21784 := (iff #20894 #21783)
-#21781 := (iff #20889 #21780)
-#21778 := (iff #20888 #21777)
-#21775 := (iff #20887 #21774)
-#21772 := (iff #20881 #21771)
-#21769 := (iff #20876 #21768)
-#21766 := (iff #20875 #21765)
-#21763 := (iff #20874 #21762)
-#21760 := (iff #20868 #21759)
-#21757 := (iff #20863 #21756)
-#21754 := (iff #20862 #21753)
-#21721 := (iff #20825 #21720)
-#21718 := (iff #20819 #21717)
-#21715 := (iff #20814 #21714)
-#21712 := (iff #20813 #21711)
-#21709 := (iff #20812 #21708)
-#21706 := (iff #20804 #21705)
-#21703 := (iff #20799 #21702)
-#21700 := (iff #20798 #21699)
-#21697 := (iff #20797 #21696)
-#21694 := (iff #20791 #21693)
-#21691 := (iff #20786 #21690)
-#21688 := (iff #20785 #21687)
-#21685 := (iff #20784 #21684)
-#21682 := (iff #20766 #21679)
-#21680 := (iff #20761 #20761)
-#21681 := [refl]: #21680
-#21683 := [quant-intro #21681]: #21682
-#21686 := [monotonicity #21683]: #21685
-#21689 := [monotonicity #21686]: #21688
-#21692 := [monotonicity #21689]: #21691
-#21695 := [monotonicity #21692]: #21694
-#21698 := [monotonicity #21695]: #21697
-#21701 := [monotonicity #21698]: #21700
-#21704 := [monotonicity #21701]: #21703
-#21707 := [monotonicity #21704]: #21706
-#21710 := [monotonicity #21707]: #21709
-#21713 := [monotonicity #21710]: #21712
-#21716 := [monotonicity #21713]: #21715
-#21719 := [monotonicity #21716]: #21718
-#21722 := [monotonicity #21719]: #21721
-#21755 := [monotonicity #21722]: #21754
-#21758 := [monotonicity #21755]: #21757
-#21751 := (iff #20854 #21750)
-#21748 := (iff #20853 #21747)
-#21745 := (iff #20852 #21744)
-#21742 := (iff #20845 #21741)
-#21739 := (iff #20840 #21738)
-#21736 := (iff #20839 #21735)
-#21733 := (iff #20838 #21732)
-#21730 := (iff #20832 #21729)
-#21727 := (iff #20827 #21726)
-#21724 := (iff #20826 #21723)
-#21725 := [monotonicity #21722]: #21724
-#21728 := [monotonicity #21725]: #21727
-#21731 := [monotonicity #21728]: #21730
-#21734 := [monotonicity #21731]: #21733
-#21737 := [monotonicity #21734]: #21736
-#21740 := [monotonicity #21737]: #21739
-#21743 := [monotonicity #21740]: #21742
-#21746 := [monotonicity #21743]: #21745
-#21749 := [monotonicity #21746]: #21748
-#21752 := [monotonicity #21749]: #21751
-#21761 := [monotonicity #21752 #21758]: #21760
-#21764 := [monotonicity #21761]: #21763
-#21767 := [monotonicity #21764]: #21766
-#21770 := [monotonicity #21767]: #21769
-#21773 := [monotonicity #21770]: #21772
-#21776 := [monotonicity #21773]: #21775
-#21779 := [monotonicity #21776]: #21778
-#21782 := [monotonicity #21779]: #21781
-#21785 := [monotonicity #21782]: #21784
-#21788 := [monotonicity #21785]: #21787
-#21791 := [monotonicity #21788]: #21790
-#21794 := [monotonicity #21791]: #21793
-#21831 := [monotonicity #21794 #21828]: #21830
-#21834 := [monotonicity #21831]: #21833
-#21677 := (iff #20992 #21676)
-#21674 := (iff #20731 #21671)
-#21672 := (iff #20726 #20726)
-#21673 := [refl]: #21672
-#21675 := [quant-intro #21673]: #21674
-#21678 := [monotonicity #21675]: #21677
-#21837 := [monotonicity #21678 #21834]: #21836
-#21840 := [monotonicity #21837]: #21839
-#21843 := [monotonicity #21840]: #21842
-#21846 := [monotonicity #21843]: #21845
-#21669 := (iff #21009 #21668)
-#21666 := (iff #20720 #21663)
-#21664 := (iff #20715 #20715)
-#21665 := [refl]: #21664
-#21667 := [quant-intro #21665]: #21666
-#21670 := [monotonicity #21667]: #21669
-#21849 := [monotonicity #21670 #21846]: #21848
-#21852 := [monotonicity #21849]: #21851
-#21855 := [monotonicity #21852]: #21854
-#21858 := [monotonicity #21855]: #21857
-#21861 := [monotonicity #21858]: #21860
-#21864 := [monotonicity #21861]: #21863
-#21867 := [monotonicity #21864]: #21866
-#21870 := [monotonicity #21867]: #21869
-#21873 := [monotonicity #21870]: #21872
-#21876 := [monotonicity #21873]: #21875
-#21879 := [monotonicity #21876]: #21878
-#21882 := [monotonicity #21879]: #21881
-#21885 := [monotonicity #21882]: #21884
-#21888 := [monotonicity #21885]: #21887
-#21891 := [monotonicity #21888]: #21890
-#21894 := [monotonicity #21891]: #21893
-#21897 := [monotonicity #21894]: #21896
-#21900 := [monotonicity #21897]: #21899
-#21903 := [monotonicity #21900]: #21902
-#13426 := (and #5433 #14701 #12736 #4826)
-#17163 := (not #13426)
-#17166 := (forall (vars (?v0 Int)) #17163)
-#14059 := (and #5433 #14701)
-#14054 := (not #14059)
-#13437 := (or #14054 #12734 #12748)
-#13432 := (forall (vars (?v0 Int)) #13437)
-#17170 := (and #13432 #17166)
-#17141 := (and #17139 #17140)
-#17142 := (not #17141)
-#17469 := (or #17142 #17451 #17464)
-#17472 := (not #17469)
-#17475 := (or #17472 #17170)
-#17481 := (and #12706 #12423 #12426 #4814 #4816 #4818 #4820 #17475)
-#17041 := (not #12508)
-#13498 := (or #14054 #12468 #12482)
-#13497 := (forall (vars (?v0 Int)) #13498)
-#17044 := (and #13497 #17041)
-#17019 := (and #17017 #17018)
-#17020 := (not #17019)
-#17334 := (or #17020 #17316 #17329)
-#17337 := (not #17334)
-#17340 := (or #17337 #17044)
-#17343 := (and #12460 #17340)
-#17346 := (or #12465 #17343)
-#17352 := (and #12444 #13511 #12523 #4971 #12452 #12435 #17346)
-#17357 := (or #16993 #16996 #17352)
-#17397 := (and #12426 #12613 #5012 #5013 #12423 #12435 #17357)
-#17363 := (and #4933 #4938 #4943 #4948 #4953 #4956 #4958 #12423 #12435 #17357)
-#17368 := (or #16955 #16964 #17363)
-#17374 := (and #4933 #4935 #17368)
-#17379 := (or #16955 #16958 #17374)
-#17385 := (and #12423 #12426 #12612 #17379)
-#17402 := (or #17385 #17397)
-#17408 := (and #4933 #4938 #12423 #12426 #17402)
-#17413 := (or #16955 #16964 #17408)
-#17419 := (and #4933 #4935 #17413)
-#17424 := (or #16955 #16958 #17419)
-#17430 := (and #12423 #12426 #12707 #17424)
-#17486 := (or #17430 #17481)
-#13522 := (or #14054 #12889 #12903)
-#13521 := (forall (vars (?v0 Int)) #13522)
-#17492 := (and #12419 #4780 #12951 #13569 #13550 #12929 #13535 #12423 #12426 #12923 #13521 #12880 #4806 #4891 #4811 #4902 #4906 #4910 #4914 #4918 #4923 #17486)
-#17497 := (or #12418 #16901 #17492)
-#13582 := (or #14054 #12385 #12397)
-#13581 := (forall (vars (?v0 Int)) #13582)
-#17500 := (and #13581 #17497)
-#16877 := (and #16875 #16876)
-#16878 := (not #16877)
-#16886 := (or #16878 #16879 #16885)
-#16887 := (not #16886)
-#17503 := (or #16887 #17500)
-#17506 := (and #12379 #17503)
-#17509 := (or #12382 #17506)
-#17515 := (and #4739 #4745 #4750 #4755 #4760 #4765 #17509)
-#17520 := (or #16842 #16851 #17515)
-#17526 := (and #4739 #4741 #17520)
-#17531 := (or #16842 #16845 #17526)
-#17534 := (and #4733 #17531)
-#17537 := (or #12210 #17534)
-#21070 := (iff #17537 #21069)
-#21067 := (iff #17534 #21064)
-#21059 := (and #4733 #21056)
-#21065 := (iff #21059 #21064)
-#21066 := [rewrite]: #21065
-#21060 := (iff #17534 #21059)
-#21057 := (iff #17531 #21056)
-#21054 := (iff #17526 #21051)
-#21046 := (and #4739 #4741 #21043)
-#21052 := (iff #21046 #21051)
-#21053 := [rewrite]: #21052
-#21047 := (iff #17526 #21046)
-#21044 := (iff #17520 #21043)
-#21041 := (iff #17515 #21038)
-#21033 := (and #4739 #4745 #4750 #4755 #4760 #4765 #21030)
-#21039 := (iff #21033 #21038)
-#21040 := [rewrite]: #21039
-#21034 := (iff #17515 #21033)
-#21031 := (iff #17509 #21030)
-#21028 := (iff #17506 #21025)
-#21020 := (and #12379 #21017)
-#21026 := (iff #21020 #21025)
-#21027 := [rewrite]: #21026
-#21021 := (iff #17506 #21020)
-#21018 := (iff #17503 #21017)
-#21015 := (iff #17500 #21012)
-#21006 := (and #20720 #21003)
-#21013 := (iff #21006 #21012)
-#21014 := [rewrite]: #21013
-#21007 := (iff #17500 #21006)
-#21004 := (iff #17497 #21003)
-#21001 := (iff #17492 #20998)
-#20984 := (and #12419 #4780 #12951 #13569 #13550 #12929 #13535 #12423 #12426 #12923 #20731 #12880 #4806 #4891 #4811 #4902 #4906 #4910 #4914 #4918 #4923 #20981)
-#20999 := (iff #20984 #20998)
-#21000 := [rewrite]: #20999
-#20985 := (iff #17492 #20984)
-#20982 := (iff #17486 #20981)
-#20979 := (iff #17481 #20976)
-#20971 := (and #12706 #12423 #12426 #4814 #4816 #4818 #4820 #20968)
-#20977 := (iff #20971 #20976)
-#20978 := [rewrite]: #20977
-#20972 := (iff #17481 #20971)
-#20969 := (iff #17475 #20968)
-#20966 := (iff #17170 #20963)
-#20957 := (and #20939 #20954)
-#20964 := (iff #20957 #20963)
-#20965 := [rewrite]: #20964
-#20958 := (iff #17170 #20957)
-#20955 := (iff #17166 #20954)
-#20952 := (iff #17163 #20943)
-#20944 := (not #20943)
-#20947 := (not #20944)
-#20950 := (iff #20947 #20943)
-#20951 := [rewrite]: #20950
-#20948 := (iff #17163 #20947)
-#20945 := (iff #13426 #20944)
-#20946 := [rewrite]: #20945
-#20949 := [monotonicity #20946]: #20948
-#20953 := [trans #20949 #20951]: #20952
-#20956 := [quant-intro #20953]: #20955
-#20940 := (iff #13432 #20939)
-#20937 := (iff #13437 #20934)
-#18608 := (or #6155 #17965)
-#20931 := (or #18608 #12734 #12748)
-#20935 := (iff #20931 #20934)
-#20936 := [rewrite]: #20935
-#20932 := (iff #13437 #20931)
-#18617 := (iff #14054 #18608)
-#18609 := (not #18608)
-#18612 := (not #18609)
-#18615 := (iff #18612 #18608)
-#18616 := [rewrite]: #18615
-#18613 := (iff #14054 #18612)
-#18610 := (iff #14059 #18609)
-#18611 := [rewrite]: #18610
-#18614 := [monotonicity #18611]: #18613
-#18618 := [trans #18614 #18616]: #18617
-#20933 := [monotonicity #18618]: #20932
-#20938 := [trans #20933 #20936]: #20937
-#20941 := [quant-intro #20938]: #20940
-#20959 := [monotonicity #20941 #20956]: #20958
-#20967 := [trans #20959 #20965]: #20966
-#20929 := (iff #17472 #20928)
-#20926 := (iff #17469 #20923)
-#20909 := (or #20907 #20908)
-#20920 := (or #20909 #17451 #17464)
-#20924 := (iff #20920 #20923)
-#20925 := [rewrite]: #20924
-#20921 := (iff #17469 #20920)
-#20918 := (iff #17142 #20909)
-#20910 := (not #20909)
-#20913 := (not #20910)
-#20916 := (iff #20913 #20909)
-#20917 := [rewrite]: #20916
-#20914 := (iff #17142 #20913)
-#20911 := (iff #17141 #20910)
-#20912 := [rewrite]: #20911
-#20915 := [monotonicity #20912]: #20914
-#20919 := [trans #20915 #20917]: #20918
-#20922 := [monotonicity #20919]: #20921
-#20927 := [trans #20922 #20925]: #20926
-#20930 := [monotonicity #20927]: #20929
-#20970 := [monotonicity #20930 #20967]: #20969
-#20973 := [monotonicity #20970]: #20972
-#20980 := [trans #20973 #20978]: #20979
-#20905 := (iff #17430 #20902)
-#20897 := (and #12423 #12426 #12707 #20894)
-#20903 := (iff #20897 #20902)
-#20904 := [rewrite]: #20903
-#20898 := (iff #17430 #20897)
-#20895 := (iff #17424 #20894)
-#20892 := (iff #17419 #20889)
-#20884 := (and #4933 #4935 #20881)
-#20890 := (iff #20884 #20889)
-#20891 := [rewrite]: #20890
-#20885 := (iff #17419 #20884)
-#20882 := (iff #17413 #20881)
-#20879 := (iff #17408 #20876)
-#20871 := (and #4933 #4938 #12423 #12426 #20868)
-#20877 := (iff #20871 #20876)
-#20878 := [rewrite]: #20877
-#20872 := (iff #17408 #20871)
-#20869 := (iff #17402 #20868)
-#20866 := (iff #17397 #20863)
-#20859 := (and #12426 #12613 #5012 #5013 #12423 #12435 #20819)
-#20864 := (iff #20859 #20863)
-#20865 := [rewrite]: #20864
-#20860 := (iff #17397 #20859)
-#20820 := (iff #17357 #20819)
-#20817 := (iff #17352 #20814)
-#20807 := (and #12444 #13511 #12523 #4971 #12452 #12435 #20804)
-#20815 := (iff #20807 #20814)
-#20816 := [rewrite]: #20815
-#20808 := (iff #17352 #20807)
-#20805 := (iff #17346 #20804)
-#20802 := (iff #17343 #20799)
-#20794 := (and #12460 #20791)
-#20800 := (iff #20794 #20799)
-#20801 := [rewrite]: #20800
-#20795 := (iff #17343 #20794)
-#20792 := (iff #17340 #20791)
-#20789 := (iff #17044 #20786)
-#20781 := (and #20766 #20770)
-#20787 := (iff #20781 #20786)
-#20788 := [rewrite]: #20787
-#20782 := (iff #17044 #20781)
-#20779 := (iff #17041 #20770)
-#20774 := (not #20771)
-#20777 := (iff #20774 #20770)
-#20778 := [rewrite]: #20777
-#20775 := (iff #17041 #20774)
-#20772 := (iff #12508 #20771)
-#20773 := [rewrite]: #20772
-#20776 := [monotonicity #20773]: #20775
-#20780 := [trans #20776 #20778]: #20779
-#20767 := (iff #13497 #20766)
-#20764 := (iff #13498 #20761)
-#20758 := (or #18608 #12468 #12482)
-#20762 := (iff #20758 #20761)
-#20763 := [rewrite]: #20762
-#20759 := (iff #13498 #20758)
-#20760 := [monotonicity #18618]: #20759
-#20765 := [trans #20760 #20763]: #20764
-#20768 := [quant-intro #20765]: #20767
-#20783 := [monotonicity #20768 #20780]: #20782
-#20790 := [trans #20783 #20788]: #20789
-#20756 := (iff #17337 #20755)
-#20753 := (iff #17334 #20750)
-#20736 := (or #20734 #20735)
-#20747 := (or #20736 #17316 #17329)
-#20751 := (iff #20747 #20750)
-#20752 := [rewrite]: #20751
-#20748 := (iff #17334 #20747)
-#20745 := (iff #17020 #20736)
-#20737 := (not #20736)
-#20740 := (not #20737)
-#20743 := (iff #20740 #20736)
-#20744 := [rewrite]: #20743
-#20741 := (iff #17020 #20740)
-#20738 := (iff #17019 #20737)
-#20739 := [rewrite]: #20738
-#20742 := [monotonicity #20739]: #20741
-#20746 := [trans #20742 #20744]: #20745
-#20749 := [monotonicity #20746]: #20748
-#20754 := [trans #20749 #20752]: #20753
-#20757 := [monotonicity #20754]: #20756
-#20793 := [monotonicity #20757 #20790]: #20792
-#20796 := [monotonicity #20793]: #20795
-#20803 := [trans #20796 #20801]: #20802
-#20806 := [monotonicity #20803]: #20805
-#20809 := [monotonicity #20806]: #20808
-#20818 := [trans #20809 #20816]: #20817
-#20821 := [monotonicity #20818]: #20820
-#20861 := [monotonicity #20821]: #20860
-#20867 := [trans #20861 #20865]: #20866
-#20857 := (iff #17385 #20854)
-#20848 := (and #12423 #12426 #12612 #20845)
-#20855 := (iff #20848 #20854)
-#20856 := [rewrite]: #20855
-#20849 := (iff #17385 #20848)
-#20846 := (iff #17379 #20845)
-#20843 := (iff #17374 #20840)
-#20835 := (and #4933 #4935 #20832)
-#20841 := (iff #20835 #20840)
-#20842 := [rewrite]: #20841
-#20836 := (iff #17374 #20835)
-#20833 := (iff #17368 #20832)
-#20830 := (iff #17363 #20827)
-#20822 := (and #4933 #4938 #4943 #4948 #4953 #4956 #4958 #12423 #12435 #20819)
-#20828 := (iff #20822 #20827)
-#20829 := [rewrite]: #20828
-#20823 := (iff #17363 #20822)
-#20824 := [monotonicity #20821]: #20823
-#20831 := [trans #20824 #20829]: #20830
-#20834 := [monotonicity #20831]: #20833
-#20837 := [monotonicity #20834]: #20836
-#20844 := [trans #20837 #20842]: #20843
-#20847 := [monotonicity #20844]: #20846
-#20850 := [monotonicity #20847]: #20849
-#20858 := [trans #20850 #20856]: #20857
-#20870 := [monotonicity #20858 #20867]: #20869
-#20873 := [monotonicity #20870]: #20872
-#20880 := [trans #20873 #20878]: #20879
-#20883 := [monotonicity #20880]: #20882
-#20886 := [monotonicity #20883]: #20885
-#20893 := [trans #20886 #20891]: #20892
-#20896 := [monotonicity #20893]: #20895
-#20899 := [monotonicity #20896]: #20898
-#20906 := [trans #20899 #20904]: #20905
-#20983 := [monotonicity #20906 #20980]: #20982
-#20732 := (iff #13521 #20731)
-#20729 := (iff #13522 #20726)
-#20723 := (or #18608 #12889 #12903)
-#20727 := (iff #20723 #20726)
-#20728 := [rewrite]: #20727
-#20724 := (iff #13522 #20723)
-#20725 := [monotonicity #18618]: #20724
-#20730 := [trans #20725 #20728]: #20729
-#20733 := [quant-intro #20730]: #20732
-#20986 := [monotonicity #20733 #20983]: #20985
-#21002 := [trans #20986 #21000]: #21001
-#21005 := [monotonicity #21002]: #21004
-#20721 := (iff #13581 #20720)
-#20718 := (iff #13582 #20715)
-#20712 := (or #18608 #12385 #12397)
-#20716 := (iff #20712 #20715)
-#20717 := [rewrite]: #20716
-#20713 := (iff #13582 #20712)
-#20714 := [monotonicity #18618]: #20713
-#20719 := [trans #20714 #20717]: #20718
-#20722 := [quant-intro #20719]: #20721
-#21008 := [monotonicity #20722 #21005]: #21007
-#21016 := [trans #21008 #21014]: #21015
-#20710 := (iff #16887 #20709)
-#20707 := (iff #16886 #20704)
-#20690 := (or #20688 #20689)
-#20701 := (or #20690 #16879 #16885)
-#20705 := (iff #20701 #20704)
-#20706 := [rewrite]: #20705
-#20702 := (iff #16886 #20701)
-#20699 := (iff #16878 #20690)
-#20691 := (not #20690)
-#20694 := (not #20691)
-#20697 := (iff #20694 #20690)
-#20698 := [rewrite]: #20697
-#20695 := (iff #16878 #20694)
-#20692 := (iff #16877 #20691)
-#20693 := [rewrite]: #20692
-#20696 := [monotonicity #20693]: #20695
-#20700 := [trans #20696 #20698]: #20699
-#20703 := [monotonicity #20700]: #20702
-#20708 := [trans #20703 #20706]: #20707
-#20711 := [monotonicity #20708]: #20710
-#21019 := [monotonicity #20711 #21016]: #21018
-#21022 := [monotonicity #21019]: #21021
-#21029 := [trans #21022 #21027]: #21028
-#21032 := [monotonicity #21029]: #21031
-#21035 := [monotonicity #21032]: #21034
-#21042 := [trans #21035 #21040]: #21041
-#21045 := [monotonicity #21042]: #21044
-#21048 := [monotonicity #21045]: #21047
-#21055 := [trans #21048 #21053]: #21054
-#21058 := [monotonicity #21055]: #21057
-#21061 := [monotonicity #21058]: #21060
-#21068 := [trans #21061 #21066]: #21067
-#21071 := [monotonicity #21068]: #21070
-#17148 := (+ #17147 #12746)
-#17149 := (<= #17148 0::Int)
-#17143 := (+ ?v0!15 #12352)
-#17144 := (>= #17143 0::Int)
-#17150 := (or #17142 #17144 #17149)
-#17151 := (not #17150)
-#17174 := (or #17151 #17170)
-#17135 := (not #11320)
-#17132 := (not #11329)
-#17129 := (not #11338)
-#17126 := (not #11347)
-#16916 := (not #12431)
-#17178 := (and #12710 #16916 #17126 #17129 #17132 #17135 #17174)
-#17026 := (+ #17025 #12480)
-#17027 := (<= #17026 0::Int)
-#17021 := (+ ?v0!14 #12461)
-#17022 := (>= #17021 0::Int)
-#17028 := (or #17020 #17022 #17027)
-#17029 := (not #17028)
-#17048 := (or #17029 #17044)
-#17013 := (not #12465)
-#17052 := (and #17013 #17048)
-#17056 := (or #12465 #17052)
-#17008 := (not #12457)
-#17005 := (not #11608)
-#17002 := (not #12527)
-#13504 := (and #12444 #13511)
-#13503 := (not #13504)
-#16999 := (not #13503)
-#17060 := (and #16999 #17002 #17005 #17008 #17056)
-#17064 := (or #16993 #16996 #17060)
-#16990 := (not #12440)
-#17093 := (not #11758)
-#17090 := (not #11767)
-#17096 := (and #16916 #12618 #17090 #17093 #16990 #17064)
-#16987 := (not #11647)
-#16984 := (not #11656)
-#16981 := (not #12561)
-#16978 := (not #11674)
-#16975 := (not #11683)
-#16972 := (not #11692)
-#16967 := (not #11701)
-#17068 := (and #16967 #16972 #16975 #16978 #16981 #16984 #16987 #16990 #17064)
-#17072 := (or #16955 #16964 #17068)
-#16961 := (not #11713)
-#17076 := (and #16961 #17072)
-#17080 := (or #16955 #16958 #17076)
-#17084 := (and #16916 #12612 #17080)
-#17100 := (or #17084 #17096)
-#17104 := (and #16967 #16916 #17100)
-#17108 := (or #16955 #16964 #17104)
-#17112 := (and #16961 #17108)
-#17116 := (or #16955 #16958 #17112)
-#17120 := (and #16916 #12707 #17116)
-#17182 := (or #17120 #17178)
-#16950 := (not #11934)
-#16947 := (not #11943)
-#16944 := (not #11952)
-#16941 := (not #11961)
-#16938 := (not #11970)
-#16935 := (not #14815)
-#16932 := (not #12874)
-#16929 := (not #12886)
-#16919 := (not #12926)
-#13528 := (and #12929 #13535)
-#13527 := (not #13528)
-#16913 := (not #13527)
-#13547 := (and #12426 #13550)
-#13542 := (not #13547)
-#16910 := (not #13542)
-#13562 := (and #12951 #13569)
-#13561 := (not #13562)
-#16907 := (not #13561)
-#16904 := (not #12967)
-#17186 := (and #16904 #16907 #16910 #16913 #16916 #16919 #13521 #16929 #16932 #16935 #16938 #16941 #16944 #16947 #16950 #17182)
-#17190 := (or #13124 #16901 #17186)
-#17194 := (and #13581 #17190)
-#17198 := (or #16887 #17194)
-#16871 := (not #12382)
-#17202 := (and #16871 #17198)
-#17206 := (or #12382 #17202)
-#16866 := (not #12150)
-#16863 := (not #12159)
-#16860 := (not #12168)
-#16857 := (not #12177)
-#16854 := (not #12186)
-#17210 := (and #16854 #16857 #16860 #16863 #16866 #17206)
-#17214 := (or #16842 #16851 #17210)
-#16848 := (not #12198)
-#17218 := (and #16848 #17214)
-#17222 := (or #16842 #16845 #17218)
-#16839 := (not #12210)
-#17226 := (and #16839 #17222)
-#17230 := (or #12210 #17226)
-#17538 := (iff #17230 #17537)
-#17535 := (iff #17226 #17534)
-#17532 := (iff #17222 #17531)
-#17529 := (iff #17218 #17526)
-#17523 := (and #4742 #17520)
-#17527 := (iff #17523 #17526)
-#17528 := [rewrite]: #17527
-#17524 := (iff #17218 #17523)
-#17521 := (iff #17214 #17520)
-#17518 := (iff #17210 #17515)
-#17512 := (and #4746 #4750 #4755 #4760 #4765 #17509)
-#17516 := (iff #17512 #17515)
-#17517 := [rewrite]: #17516
-#17513 := (iff #17210 #17512)
-#17510 := (iff #17206 #17509)
-#17507 := (iff #17202 #17506)
-#17504 := (iff #17198 #17503)
-#17501 := (iff #17194 #17500)
-#17498 := (iff #17190 #17497)
-#17495 := (iff #17186 #17492)
-#17489 := (and #12964 #13562 #13547 #13528 #12428 #12923 #13521 #12883 #4891 #4811 #4902 #4906 #4910 #4914 #4924 #17486)
-#17493 := (iff #17489 #17492)
-#17494 := [rewrite]: #17493
-#17490 := (iff #17186 #17489)
-#17487 := (iff #17182 #17486)
-#17484 := (iff #17178 #17481)
-#17478 := (and #12706 #12428 #4814 #4816 #4818 #4820 #17475)
-#17482 := (iff #17478 #17481)
-#17483 := [rewrite]: #17482
-#17479 := (iff #17178 #17478)
-#17476 := (iff #17174 #17475)
-#17473 := (iff #17151 #17472)
-#17470 := (iff #17150 #17469)
-#17467 := (iff #17149 #17464)
-#17456 := (+ #12746 #17147)
-#17459 := (<= #17456 0::Int)
-#17465 := (iff #17459 #17464)
-#17466 := [rewrite]: #17465
-#17460 := (iff #17149 #17459)
-#17457 := (= #17148 #17456)
-#17458 := [rewrite]: #17457
-#17461 := [monotonicity #17458]: #17460
-#17468 := [trans #17461 #17466]: #17467
-#17454 := (iff #17144 #17451)
-#17443 := (+ #12352 ?v0!15)
-#17446 := (>= #17443 0::Int)
-#17452 := (iff #17446 #17451)
-#17453 := [rewrite]: #17452
-#17447 := (iff #17144 #17446)
-#17444 := (= #17143 #17443)
-#17445 := [rewrite]: #17444
-#17448 := [monotonicity #17445]: #17447
-#17455 := [trans #17448 #17453]: #17454
-#17471 := [monotonicity #17455 #17468]: #17470
-#17474 := [monotonicity #17471]: #17473
-#17477 := [monotonicity #17474]: #17476
-#17441 := (iff #17135 #4820)
-#17442 := [rewrite]: #17441
-#17439 := (iff #17132 #4818)
-#17440 := [rewrite]: #17439
-#17437 := (iff #17129 #4816)
-#17438 := [rewrite]: #17437
-#17435 := (iff #17126 #4814)
-#17436 := [rewrite]: #17435
-#17260 := (iff #16916 #12428)
-#17261 := [rewrite]: #17260
-#17480 := [monotonicity #12714 #17261 #17436 #17438 #17440 #17442 #17477]: #17479
-#17485 := [trans #17480 #17483]: #17484
-#17433 := (iff #17120 #17430)
-#17427 := (and #12428 #12707 #17424)
-#17431 := (iff #17427 #17430)
-#17432 := [rewrite]: #17431
-#17428 := (iff #17120 #17427)
-#17425 := (iff #17116 #17424)
-#17422 := (iff #17112 #17419)
-#17416 := (and #4936 #17413)
-#17420 := (iff #17416 #17419)
-#17421 := [rewrite]: #17420
-#17417 := (iff #17112 #17416)
-#17414 := (iff #17108 #17413)
-#17411 := (iff #17104 #17408)
-#17405 := (and #4939 #12428 #17402)
-#17409 := (iff #17405 #17408)
-#17410 := [rewrite]: #17409
-#17406 := (iff #17104 #17405)
-#17403 := (iff #17100 #17402)
-#17400 := (iff #17096 #17397)
-#17394 := (and #12428 #12613 #5012 #5013 #12437 #17357)
-#17398 := (iff #17394 #17397)
-#17399 := [rewrite]: #17398
-#17395 := (iff #17096 #17394)
-#17358 := (iff #17064 #17357)
-#17355 := (iff #17060 #17352)
-#17349 := (and #13504 #12523 #4971 #12454 #17346)
-#17353 := (iff #17349 #17352)
-#17354 := [rewrite]: #17353
-#17350 := (iff #17060 #17349)
-#17347 := (iff #17056 #17346)
-#17344 := (iff #17052 #17343)
-#17341 := (iff #17048 #17340)
-#17338 := (iff #17029 #17337)
-#17335 := (iff #17028 #17334)
-#17332 := (iff #17027 #17329)
-#17321 := (+ #12480 #17025)
-#17324 := (<= #17321 0::Int)
-#17330 := (iff #17324 #17329)
-#17331 := [rewrite]: #17330
-#17325 := (iff #17027 #17324)
-#17322 := (= #17026 #17321)
-#17323 := [rewrite]: #17322
-#17326 := [monotonicity #17323]: #17325
-#17333 := [trans #17326 #17331]: #17332
-#17319 := (iff #17022 #17316)
-#17308 := (+ #12461 ?v0!14)
-#17311 := (>= #17308 0::Int)
-#17317 := (iff #17311 #17316)
-#17318 := [rewrite]: #17317
-#17312 := (iff #17022 #17311)
-#17309 := (= #17021 #17308)
-#17310 := [rewrite]: #17309
-#17313 := [monotonicity #17310]: #17312
-#17320 := [trans #17313 #17318]: #17319
-#17336 := [monotonicity #17320 #17333]: #17335
-#17339 := [monotonicity #17336]: #17338
-#17342 := [monotonicity #17339]: #17341
-#17306 := (iff #17013 #12460)
-#17307 := [rewrite]: #17306
-#17345 := [monotonicity #17307 #17342]: #17344
-#17348 := [monotonicity #17345]: #17347
-#17304 := (iff #17008 #12454)
-#17305 := [rewrite]: #17304
-#17302 := (iff #17005 #4971)
-#17303 := [rewrite]: #17302
-#17300 := (iff #17002 #12523)
-#17301 := [rewrite]: #17300
-#17298 := (iff #16999 #13504)
-#17299 := [rewrite]: #17298
-#17351 := [monotonicity #17299 #17301 #17303 #17305 #17348]: #17350
-#17356 := [trans #17351 #17354]: #17355
-#17359 := [monotonicity #17356]: #17358
-#17296 := (iff #16990 #12437)
-#17297 := [rewrite]: #17296
-#17392 := (iff #17093 #5013)
-#17393 := [rewrite]: #17392
-#17390 := (iff #17090 #5012)
-#17391 := [rewrite]: #17390
-#17396 := [monotonicity #17261 #12622 #17391 #17393 #17297 #17359]: #17395
-#17401 := [trans #17396 #17399]: #17400
-#17388 := (iff #17084 #17385)
-#17382 := (and #12428 #12612 #17379)
-#17386 := (iff #17382 #17385)
-#17387 := [rewrite]: #17386
-#17383 := (iff #17084 #17382)
-#17380 := (iff #17080 #17379)
-#17377 := (iff #17076 #17374)
-#17371 := (and #4936 #17368)
-#17375 := (iff #17371 #17374)
-#17376 := [rewrite]: #17375
-#17372 := (iff #17076 #17371)
-#17369 := (iff #17072 #17368)
-#17366 := (iff #17068 #17363)
-#17360 := (and #4939 #4943 #4948 #4953 #12423 #4956 #4958 #12437 #17357)
-#17364 := (iff #17360 #17363)
-#17365 := [rewrite]: #17364
-#17361 := (iff #17068 #17360)
-#17294 := (iff #16987 #4958)
-#17295 := [rewrite]: #17294
-#17292 := (iff #16984 #4956)
-#17293 := [rewrite]: #17292
-#17290 := (iff #16981 #12423)
-#17291 := [rewrite]: #17290
-#17288 := (iff #16978 #4953)
-#17289 := [rewrite]: #17288
-#17286 := (iff #16975 #4948)
-#17287 := [rewrite]: #17286
-#17284 := (iff #16972 #4943)
-#17285 := [rewrite]: #17284
-#17282 := (iff #16967 #4939)
-#17283 := [rewrite]: #17282
-#17362 := [monotonicity #17283 #17285 #17287 #17289 #17291 #17293 #17295 #17297 #17359]: #17361
-#17367 := [trans #17362 #17365]: #17366
-#17370 := [monotonicity #17367]: #17369
-#17280 := (iff #16961 #4936)
-#17281 := [rewrite]: #17280
-#17373 := [monotonicity #17281 #17370]: #17372
-#17378 := [trans #17373 #17376]: #17377
-#17381 := [monotonicity #17378]: #17380
-#17384 := [monotonicity #17261 #17381]: #17383
-#17389 := [trans #17384 #17387]: #17388
-#17404 := [monotonicity #17389 #17401]: #17403
-#17407 := [monotonicity #17283 #17261 #17404]: #17406
-#17412 := [trans #17407 #17410]: #17411
-#17415 := [monotonicity #17412]: #17414
-#17418 := [monotonicity #17281 #17415]: #17417
-#17423 := [trans #17418 #17421]: #17422
-#17426 := [monotonicity #17423]: #17425
-#17429 := [monotonicity #17261 #17426]: #17428
-#17434 := [trans #17429 #17432]: #17433
-#17488 := [monotonicity #17434 #17485]: #17487
-#17278 := (iff #16950 #4924)
-#17279 := [rewrite]: #17278
-#17276 := (iff #16947 #4914)
-#17277 := [rewrite]: #17276
-#17274 := (iff #16944 #4910)
-#17275 := [rewrite]: #17274
-#17272 := (iff #16941 #4906)
-#17273 := [rewrite]: #17272
-#17270 := (iff #16938 #4902)
-#17271 := [rewrite]: #17270
-#17268 := (iff #16935 #4811)
-#17269 := [rewrite]: #17268
-#17266 := (iff #16932 #4891)
-#17267 := [rewrite]: #17266
-#17264 := (iff #16929 #12883)
-#17265 := [rewrite]: #17264
-#17262 := (iff #16919 #12923)
-#17263 := [rewrite]: #17262
-#17258 := (iff #16913 #13528)
-#17259 := [rewrite]: #17258
-#17256 := (iff #16910 #13547)
-#17257 := [rewrite]: #17256
-#17254 := (iff #16907 #13562)
-#17255 := [rewrite]: #17254
-#17252 := (iff #16904 #12964)
-#17253 := [rewrite]: #17252
-#17491 := [monotonicity #17253 #17255 #17257 #17259 #17261 #17263 #17265 #17267 #17269 #17271 #17273 #17275 #17277 #17279 #17488]: #17490
-#17496 := [trans #17491 #17494]: #17495
-#17499 := [monotonicity #13128 #17496]: #17498
-#17502 := [monotonicity #17499]: #17501
-#17505 := [monotonicity #17502]: #17504
-#17250 := (iff #16871 #12379)
-#17251 := [rewrite]: #17250
-#17508 := [monotonicity #17251 #17505]: #17507
-#17511 := [monotonicity #17508]: #17510
-#17248 := (iff #16866 #4765)
-#17249 := [rewrite]: #17248
-#17246 := (iff #16863 #4760)
-#17247 := [rewrite]: #17246
-#17244 := (iff #16860 #4755)
-#17245 := [rewrite]: #17244
-#17242 := (iff #16857 #4750)
-#17243 := [rewrite]: #17242
-#17240 := (iff #16854 #4746)
-#17241 := [rewrite]: #17240
-#17514 := [monotonicity #17241 #17243 #17245 #17247 #17249 #17511]: #17513
-#17519 := [trans #17514 #17517]: #17518
-#17522 := [monotonicity #17519]: #17521
-#17238 := (iff #16848 #4742)
-#17239 := [rewrite]: #17238
-#17525 := [monotonicity #17239 #17522]: #17524
-#17530 := [trans #17525 #17528]: #17529
-#17533 := [monotonicity #17530]: #17532
-#17236 := (iff #16839 #4733)
-#17237 := [rewrite]: #17236
-#17536 := [monotonicity #17237 #17533]: #17535
-#17539 := [monotonicity #17536]: #17538
-#13425 := (exists (vars (?v0 Int)) #13426)
-#13431 := (not #13432)
-#13420 := (or #13431 #13425)
-#13419 := (and #13432 #13420)
-#13414 := (or #12707 #12431 #11347 #11338 #11329 #11320 #13419)
-#13492 := (not #13497)
-#13491 := (or #13492 #12508)
-#13486 := (and #13497 #13491)
-#13485 := (or #12465 #13486)
-#13480 := (and #12460 #13485)
-#13479 := (or #13503 #12527 #11608 #12457 #13480)
-#13474 := (and #12444 #13511 #13479)
-#13456 := (or #12431 #12612 #11767 #11758 #12440 #13474)
-#13473 := (or #11701 #11692 #11683 #11674 #12561 #11656 #11647 #12440 #13474)
-#13468 := (and #4933 #4938 #13473)
-#13467 := (or #11713 #13468)
-#13462 := (and #4933 #4935 #13467)
-#13461 := (or #12431 #12613 #13462)
-#13455 := (and #13461 #13456)
-#13450 := (or #11701 #12431 #13455)
-#13449 := (and #4933 #4938 #13450)
-#13444 := (or #11713 #13449)
-#13443 := (and #4933 #4935 #13444)
-#13438 := (or #12431 #12706 #13443)
-#13413 := (and #13438 #13414)
-#13516 := (not #13521)
-#13408 := (or #12967 #13561 #13542 #13527 #12431 #12926 #13516 #12886 #12874 #14815 #11970 #11961 #11952 #11943 #11934 #13413)
-#13407 := (and #12419 #4780 #13408)
-#13576 := (not #13581)
-#13402 := (or #13576 #13407)
-#13401 := (and #13581 #13402)
-#13396 := (or #12382 #13401)
-#13395 := (and #12379 #13396)
-#13390 := (or #12186 #12177 #12168 #12159 #12150 #13395)
-#13389 := (and #4739 #4745 #13390)
-#13384 := (or #12198 #13389)
-#13383 := (and #4739 #4741 #13384)
-#13378 := (or #12210 #13383)
-#13377 := (and #4733 #13378)
-#13372 := (not #13377)
-#17231 := (~ #13372 #17230)
-#17227 := (not #13378)
-#17228 := (~ #17227 #17226)
-#17223 := (not #13383)
-#17224 := (~ #17223 #17222)
-#17219 := (not #13384)
-#17220 := (~ #17219 #17218)
-#17215 := (not #13389)
-#17216 := (~ #17215 #17214)
-#17211 := (not #13390)
-#17212 := (~ #17211 #17210)
-#17207 := (not #13395)
-#17208 := (~ #17207 #17206)
-#17203 := (not #13396)
-#17204 := (~ #17203 #17202)
-#17199 := (not #13401)
-#17200 := (~ #17199 #17198)
-#17195 := (not #13402)
-#17196 := (~ #17195 #17194)
-#17191 := (not #13407)
-#17192 := (~ #17191 #17190)
-#17187 := (not #13408)
-#17188 := (~ #17187 #17186)
-#17183 := (not #13413)
-#17184 := (~ #17183 #17182)
-#17179 := (not #13414)
-#17180 := (~ #17179 #17178)
-#17175 := (not #13419)
-#17176 := (~ #17175 #17174)
-#17171 := (not #13420)
-#17172 := (~ #17171 #17170)
-#17167 := (not #13425)
-#17168 := (~ #17167 #17166)
-#17164 := (~ #17163 #17163)
-#17165 := [refl]: #17164
-#17169 := [nnf-neg #17165]: #17168
-#17160 := (not #13431)
-#17161 := (~ #17160 #13432)
-#17158 := (~ #13432 #13432)
-#17156 := (~ #13437 #13437)
-#17157 := [refl]: #17156
-#17159 := [nnf-pos #17157]: #17158
-#17162 := [nnf-neg #17159]: #17161
-#17173 := [nnf-neg #17162 #17169]: #17172
-#17152 := (~ #13431 #17151)
-#17153 := [sk]: #17152
-#17177 := [nnf-neg #17153 #17173]: #17176
-#17136 := (~ #17135 #17135)
-#17137 := [refl]: #17136
-#17133 := (~ #17132 #17132)
-#17134 := [refl]: #17133
-#17130 := (~ #17129 #17129)
-#17131 := [refl]: #17130
-#17127 := (~ #17126 #17126)
-#17128 := [refl]: #17127
-#16917 := (~ #16916 #16916)
-#16918 := [refl]: #16917
-#17124 := (~ #12710 #12710)
-#17125 := [refl]: #17124
-#17181 := [nnf-neg #17125 #16918 #17128 #17131 #17134 #17137 #17177]: #17180
-#17121 := (not #13438)
-#17122 := (~ #17121 #17120)
-#17117 := (not #13443)
-#17118 := (~ #17117 #17116)
-#17113 := (not #13444)
-#17114 := (~ #17113 #17112)
-#17109 := (not #13449)
-#17110 := (~ #17109 #17108)
-#17105 := (not #13450)
-#17106 := (~ #17105 #17104)
-#17101 := (not #13455)
-#17102 := (~ #17101 #17100)
-#17097 := (not #13456)
-#17098 := (~ #17097 #17096)
-#17065 := (not #13474)
-#17066 := (~ #17065 #17064)
-#17061 := (not #13479)
-#17062 := (~ #17061 #17060)
-#17057 := (not #13480)
-#17058 := (~ #17057 #17056)
-#17053 := (not #13485)
-#17054 := (~ #17053 #17052)
-#17049 := (not #13486)
-#17050 := (~ #17049 #17048)
-#17045 := (not #13491)
-#17046 := (~ #17045 #17044)
-#17042 := (~ #17041 #17041)
-#17043 := [refl]: #17042
-#17038 := (not #13492)
-#17039 := (~ #17038 #13497)
-#17036 := (~ #13497 #13497)
-#17034 := (~ #13498 #13498)
-#17035 := [refl]: #17034
-#17037 := [nnf-pos #17035]: #17036
-#17040 := [nnf-neg #17037]: #17039
-#17047 := [nnf-neg #17040 #17043]: #17046
-#17030 := (~ #13492 #17029)
-#17031 := [sk]: #17030
-#17051 := [nnf-neg #17031 #17047]: #17050
-#17014 := (~ #17013 #17013)
-#17015 := [refl]: #17014
-#17055 := [nnf-neg #17015 #17051]: #17054
-#17011 := (~ #12465 #12465)
-#17012 := [refl]: #17011
-#17059 := [nnf-neg #17012 #17055]: #17058
-#17009 := (~ #17008 #17008)
-#17010 := [refl]: #17009
-#17006 := (~ #17005 #17005)
-#17007 := [refl]: #17006
-#17003 := (~ #17002 #17002)
-#17004 := [refl]: #17003
-#17000 := (~ #16999 #16999)
-#17001 := [refl]: #17000
-#17063 := [nnf-neg #17001 #17004 #17007 #17010 #17059]: #17062
-#16997 := (~ #16996 #16996)
-#16998 := [refl]: #16997
-#16994 := (~ #16993 #16993)
-#16995 := [refl]: #16994
-#17067 := [nnf-neg #16995 #16998 #17063]: #17066
-#16991 := (~ #16990 #16990)
-#16992 := [refl]: #16991
-#17094 := (~ #17093 #17093)
-#17095 := [refl]: #17094
-#17091 := (~ #17090 #17090)
-#17092 := [refl]: #17091
-#17088 := (~ #12618 #12618)
-#17089 := [refl]: #17088
-#17099 := [nnf-neg #16918 #17089 #17092 #17095 #16992 #17067]: #17098
-#17085 := (not #13461)
-#17086 := (~ #17085 #17084)
-#17081 := (not #13462)
-#17082 := (~ #17081 #17080)
-#17077 := (not #13467)
-#17078 := (~ #17077 #17076)
-#17073 := (not #13468)
-#17074 := (~ #17073 #17072)
-#17069 := (not #13473)
-#17070 := (~ #17069 #17068)
-#16988 := (~ #16987 #16987)
-#16989 := [refl]: #16988
-#16985 := (~ #16984 #16984)
-#16986 := [refl]: #16985
-#16982 := (~ #16981 #16981)
-#16983 := [refl]: #16982
-#16979 := (~ #16978 #16978)
-#16980 := [refl]: #16979
-#16976 := (~ #16975 #16975)
-#16977 := [refl]: #16976
-#16973 := (~ #16972 #16972)
-#16974 := [refl]: #16973
-#16968 := (~ #16967 #16967)
-#16969 := [refl]: #16968
-#17071 := [nnf-neg #16969 #16974 #16977 #16980 #16983 #16986 #16989 #16992 #17067]: #17070
-#16965 := (~ #16964 #16964)
-#16966 := [refl]: #16965
-#16956 := (~ #16955 #16955)
-#16957 := [refl]: #16956
-#17075 := [nnf-neg #16957 #16966 #17071]: #17074
-#16962 := (~ #16961 #16961)
-#16963 := [refl]: #16962
-#17079 := [nnf-neg #16963 #17075]: #17078
-#16959 := (~ #16958 #16958)
-#16960 := [refl]: #16959
-#17083 := [nnf-neg #16957 #16960 #17079]: #17082
-#16970 := (~ #12612 #12612)
-#16971 := [refl]: #16970
-#17087 := [nnf-neg #16918 #16971 #17083]: #17086
-#17103 := [nnf-neg #17087 #17099]: #17102
-#17107 := [nnf-neg #16969 #16918 #17103]: #17106
-#17111 := [nnf-neg #16957 #16966 #17107]: #17110
-#17115 := [nnf-neg #16963 #17111]: #17114
-#17119 := [nnf-neg #16957 #16960 #17115]: #17118
-#16953 := (~ #12707 #12707)
-#16954 := [refl]: #16953
-#17123 := [nnf-neg #16918 #16954 #17119]: #17122
-#17185 := [nnf-neg #17123 #17181]: #17184
-#16951 := (~ #16950 #16950)
-#16952 := [refl]: #16951
-#16948 := (~ #16947 #16947)
-#16949 := [refl]: #16948
-#16945 := (~ #16944 #16944)
-#16946 := [refl]: #16945
-#16942 := (~ #16941 #16941)
-#16943 := [refl]: #16942
-#16939 := (~ #16938 #16938)
-#16940 := [refl]: #16939
-#16936 := (~ #16935 #16935)
-#16937 := [refl]: #16936
-#16933 := (~ #16932 #16932)
-#16934 := [refl]: #16933
-#16930 := (~ #16929 #16929)
-#16931 := [refl]: #16930
-#16926 := (not #13516)
-#16927 := (~ #16926 #13521)
-#16924 := (~ #13521 #13521)
-#16922 := (~ #13522 #13522)
-#16923 := [refl]: #16922
-#16925 := [nnf-pos #16923]: #16924
-#16928 := [nnf-neg #16925]: #16927
-#16920 := (~ #16919 #16919)
-#16921 := [refl]: #16920
-#16914 := (~ #16913 #16913)
-#16915 := [refl]: #16914
-#16911 := (~ #16910 #16910)
-#16912 := [refl]: #16911
-#16908 := (~ #16907 #16907)
-#16909 := [refl]: #16908
-#16905 := (~ #16904 #16904)
-#16906 := [refl]: #16905
-#17189 := [nnf-neg #16906 #16909 #16912 #16915 #16918 #16921 #16928 #16931 #16934 #16937 #16940 #16943 #16946 #16949 #16952 #17185]: #17188
-#16902 := (~ #16901 #16901)
-#16903 := [refl]: #16902
-#16899 := (~ #13124 #13124)
-#16900 := [refl]: #16899
-#17193 := [nnf-neg #16900 #16903 #17189]: #17192
-#16896 := (not #13576)
-#16897 := (~ #16896 #13581)
-#16894 := (~ #13581 #13581)
-#16892 := (~ #13582 #13582)
-#16893 := [refl]: #16892
-#16895 := [nnf-pos #16893]: #16894
-#16898 := [nnf-neg #16895]: #16897
-#17197 := [nnf-neg #16898 #17193]: #17196
-#16888 := (~ #13576 #16887)
-#16889 := [sk]: #16888
-#17201 := [nnf-neg #16889 #17197]: #17200
-#16872 := (~ #16871 #16871)
-#16873 := [refl]: #16872
-#17205 := [nnf-neg #16873 #17201]: #17204
-#16869 := (~ #12382 #12382)
-#16870 := [refl]: #16869
-#17209 := [nnf-neg #16870 #17205]: #17208
-#16867 := (~ #16866 #16866)
-#16868 := [refl]: #16867
-#16864 := (~ #16863 #16863)
-#16865 := [refl]: #16864
-#16861 := (~ #16860 #16860)
-#16862 := [refl]: #16861
-#16858 := (~ #16857 #16857)
-#16859 := [refl]: #16858
-#16855 := (~ #16854 #16854)
-#16856 := [refl]: #16855
-#17213 := [nnf-neg #16856 #16859 #16862 #16865 #16868 #17209]: #17212
-#16852 := (~ #16851 #16851)
-#16853 := [refl]: #16852
-#16843 := (~ #16842 #16842)
-#16844 := [refl]: #16843
-#17217 := [nnf-neg #16844 #16853 #17213]: #17216
-#16849 := (~ #16848 #16848)
-#16850 := [refl]: #16849
-#17221 := [nnf-neg #16850 #17217]: #17220
-#16846 := (~ #16845 #16845)
-#16847 := [refl]: #16846
-#17225 := [nnf-neg #16844 #16847 #17221]: #17224
-#16840 := (~ #16839 #16839)
-#16841 := [refl]: #16840
-#17229 := [nnf-neg #16841 #17225]: #17228
-#16837 := (~ #12210 #12210)
-#16838 := [refl]: #16837
-#17232 := [nnf-neg #16838 #17229]: #17231
-#14840 := (or #12707 #12431 #11347 #11338 #11329 #11320 #12788)
-#14845 := (and #12729 #14840)
-#14848 := (or #12967 #12961 #12948 #12938 #12431 #12926 #12920 #12886 #12874 #14815 #11970 #11961 #11952 #11943 #11934 #14845)
-#14851 := (and #12419 #4780 #14848)
-#14854 := (or #12415 #14851)
-#14857 := (and #12412 #14854)
-#14860 := (or #12382 #14857)
-#14863 := (and #12379 #14860)
-#14866 := (or #12186 #12177 #12168 #12159 #12150 #14863)
-#14869 := (and #4739 #4745 #14866)
-#14872 := (or #12198 #14869)
-#14875 := (and #4739 #4741 #14872)
-#14878 := (or #12210 #14875)
-#14881 := (and #4733 #14878)
-#14884 := (not #14881)
-#13373 := (iff #14884 #13372)
-#13374 := (iff #14881 #13377)
-#13379 := (iff #14878 #13378)
-#13380 := (iff #14875 #13383)
-#13385 := (iff #14872 #13384)
-#13386 := (iff #14869 #13389)
-#13391 := (iff #14866 #13390)
-#13392 := (iff #14863 #13395)
-#13397 := (iff #14860 #13396)
-#13398 := (iff #14857 #13401)
-#13403 := (iff #14854 #13402)
-#13404 := (iff #14851 #13407)
-#13409 := (iff #14848 #13408)
-#13410 := (iff #14845 #13413)
-#13415 := (iff #14840 #13414)
-#13416 := (iff #12788 #13419)
-#13421 := (iff #12785 #13420)
-#13422 := (iff #12782 #13425)
-#13427 := (iff #12777 #13426)
-#14696 := (iff #5617 #14701)
-#14733 := -4294967295::Int
-#14709 := (+ -4294967295::Int #243)
-#14702 := (<= #14709 0::Int)
-#14698 := (iff #14702 #14701)
-#14699 := [rewrite]: #14698
-#14703 := (iff #5617 #14702)
-#14704 := (= #5616 #14709)
-#14710 := (+ #243 -4294967295::Int)
-#14706 := (= #14710 #14709)
-#14707 := [rewrite]: #14706
-#14711 := (= #5616 #14710)
-#14728 := (= #5615 -4294967295::Int)
-#14734 := (* -1::Int 4294967295::Int)
-#14730 := (= #14734 -4294967295::Int)
-#14731 := [rewrite]: #14730
-#14735 := (= #5615 #14734)
-#8091 := (= f135 4294967295::Int)
-#1205 := 65536::Int
-#1604 := (* 65536::Int 65536::Int)
-#1609 := (- #1604 1::Int)
-#1610 := (= f135 #1609)
-#8092 := (iff #1610 #8091)
-#8089 := (= #1609 4294967295::Int)
-#1268 := 4294967296::Int
-#8082 := (- 4294967296::Int 1::Int)
-#8087 := (= #8082 4294967295::Int)
-#8088 := [rewrite]: #8087
-#8084 := (= #1609 #8082)
-#8053 := (= #1604 4294967296::Int)
-#8054 := [rewrite]: #8053
-#8085 := [monotonicity #8054]: #8084
-#8090 := [trans #8085 #8088]: #8089
-#8093 := [monotonicity #8090]: #8092
-#8081 := [asserted]: #1610
-#8096 := [mp #8081 #8093]: #8091
-#14732 := [monotonicity #8096]: #14735
-#14729 := [trans #14732 #14731]: #14728
-#14708 := [monotonicity #14729]: #14711
-#14705 := [trans #14708 #14707]: #14704
-#14700 := [monotonicity #14705]: #14703
-#14697 := [trans #14700 #14699]: #14696
-#13424 := [monotonicity #14697]: #13427
-#13423 := [quant-intro #13424]: #13422
-#13428 := (iff #12765 #13431)
-#13433 := (iff #12762 #13432)
-#13434 := (iff #12757 #13437)
-#14055 := (iff #6637 #14054)
-#14056 := (iff #5624 #14059)
-#14057 := [monotonicity #14697]: #14056
-#14052 := [monotonicity #14057]: #14055
-#13435 := [monotonicity #14052]: #13434
-#13430 := [quant-intro #13435]: #13433
-#13429 := [monotonicity #13430]: #13428
-#13418 := [monotonicity #13429 #13423]: #13421
-#13417 := [monotonicity #13430 #13418]: #13416
-#13412 := [monotonicity #13417]: #13415
-#13439 := (iff #12729 #13438)
-#13440 := (iff #12700 #13443)
-#13445 := (iff #12694 #13444)
-#13446 := (iff #12689 #13449)
-#13451 := (iff #12681 #13450)
-#13452 := (iff #12672 #13455)
-#13457 := (iff #12667 #13456)
-#13475 := (iff #12556 #13474)
-#13476 := (iff #12548 #13479)
-#13481 := (iff #12520 #13480)
-#13482 := (iff #12517 #13485)
-#13487 := (iff #12514 #13486)
-#13488 := (iff #12511 #13491)
-#13493 := (iff #12499 #13492)
-#13494 := (iff #12496 #13497)
-#13499 := (iff #12491 #13498)
-#13496 := [monotonicity #14052]: #13499
-#13495 := [quant-intro #13496]: #13494
-#13490 := [monotonicity #13495]: #13493
-#13489 := [monotonicity #13490]: #13488
-#13484 := [monotonicity #13495 #13489]: #13487
-#13483 := [monotonicity #13484]: #13482
-#13478 := [monotonicity #13483]: #13481
-#13500 := (iff #12533 #13503)
-#13505 := (iff #12530 #13504)
-#13506 := (iff #12446 #13511)
-#13541 := (+ 4294967295::Int #12447)
-#13515 := (>= #13541 1::Int)
-#13508 := (iff #13515 #13511)
-#13509 := [rewrite]: #13508
-#13512 := (iff #12446 #13515)
-#13538 := (= #12448 #13541)
-#13539 := [monotonicity #8096]: #13538
-#13513 := [monotonicity #13539]: #13512
-#13507 := [trans #13513 #13509]: #13506
-#13502 := [monotonicity #13507]: #13505
-#13501 := [monotonicity #13502]: #13500
-#13477 := [monotonicity #13501 #13478]: #13476
-#13472 := [monotonicity #13507 #13477]: #13475
-#13454 := [monotonicity #13472]: #13457
-#13458 := (iff #12637 #13461)
-#13463 := (iff #12607 #13462)
-#13464 := (iff #12601 #13467)
-#13469 := (iff #12596 #13468)
-#13470 := (iff #12588 #13473)
-#13471 := [monotonicity #13472]: #13470
-#13466 := [monotonicity #13471]: #13469
-#13465 := [monotonicity #13466]: #13464
-#13460 := [monotonicity #13465]: #13463
-#13459 := [monotonicity #13460]: #13458
-#13453 := [monotonicity #13459 #13454]: #13452
-#13448 := [monotonicity #13453]: #13451
-#13447 := [monotonicity #13448]: #13446
-#13442 := [monotonicity #13447]: #13445
-#13441 := [monotonicity #13442]: #13440
-#13436 := [monotonicity #13441]: #13439
-#13411 := [monotonicity #13436 #13412]: #13410
-#13517 := (iff #12920 #13516)
-#13518 := (iff #12917 #13521)
-#13523 := (iff #12912 #13522)
-#13520 := [monotonicity #14052]: #13523
-#13519 := [quant-intro #13520]: #13518
-#13514 := [monotonicity #13519]: #13517
-#13524 := (iff #12938 #13527)
-#13529 := (iff #12935 #13528)
-#13530 := (iff #12932 #13535)
-#13536 := (>= #13541 0::Int)
-#13532 := (iff #13536 #13535)
-#13533 := [rewrite]: #13532
-#13537 := (iff #12932 #13536)
-#13534 := [monotonicity #13539]: #13537
-#13531 := [trans #13534 #13533]: #13530
-#13526 := [monotonicity #13531]: #13529
-#13525 := [monotonicity #13526]: #13524
-#13543 := (iff #12948 #13542)
-#13544 := (iff #12945 #13547)
-#13549 := (iff #12941 #13550)
-#13556 := (+ 4294967295::Int #12877)
-#13555 := (>= #13556 0::Int)
-#13551 := (iff #13555 #13550)
-#13548 := [rewrite]: #13551
-#13552 := (iff #12941 #13555)
-#13557 := (= #12942 #13556)
-#13554 := [monotonicity #8096]: #13557
-#13553 := [monotonicity #13554]: #13552
-#13546 := [trans #13553 #13548]: #13549
-#13545 := [monotonicity #13546]: #13544
-#13540 := [monotonicity #13545]: #13543
-#13558 := (iff #12961 #13561)
-#13563 := (iff #12958 #13562)
-#13564 := (iff #12954 #13569)
-#13575 := (+ 255::Int #12901)
-#13570 := (>= #13575 0::Int)
-#13566 := (iff #13570 #13569)
-#13567 := [rewrite]: #13566
-#13571 := (iff #12954 #13570)
-#13572 := (= #12955 #13575)
-#1614 := (= f137 255::Int)
-#8095 := [asserted]: #1614
-#13573 := [monotonicity #8095]: #13572
-#13568 := [monotonicity #13573]: #13571
-#13565 := [trans #13568 #13567]: #13564
-#13560 := [monotonicity #13565]: #13563
-#13559 := [monotonicity #13560]: #13558
-#13406 := [monotonicity #13559 #13540 #13525 #13514 #13411]: #13409
-#13405 := [monotonicity #13406]: #13404
-#13577 := (iff #12415 #13576)
-#13578 := (iff #12412 #13581)
-#13583 := (iff #12407 #13582)
-#13580 := [monotonicity #14052]: #13583
-#13579 := [quant-intro #13580]: #13578
-#13574 := [monotonicity #13579]: #13577
-#13400 := [monotonicity #13574 #13405]: #13403
-#13399 := [monotonicity #13579 #13400]: #13398
-#13394 := [monotonicity #13399]: #13397
-#13393 := [monotonicity #13394]: #13392
-#13388 := [monotonicity #13393]: #13391
-#13387 := [monotonicity #13388]: #13386
-#13382 := [monotonicity #13387]: #13385
-#13381 := [monotonicity #13382]: #13380
-#13376 := [monotonicity #13381]: #13379
-#13375 := [monotonicity #13376]: #13374
-#13370 := [monotonicity #13375]: #13373
-#13269 := (not #13103)
-#14885 := (iff #13269 #14884)
-#14882 := (iff #13103 #14881)
-#14879 := (iff #13100 #14878)
-#14876 := (iff #13095 #14875)
-#14873 := (iff #13089 #14872)
-#14870 := (iff #13084 #14869)
-#14867 := (iff #13076 #14866)
-#14864 := (iff #13055 #14863)
-#14861 := (iff #13052 #14860)
-#14858 := (iff #13049 #14857)
-#14855 := (iff #13046 #14854)
-#14852 := (iff #13041 #14851)
-#14849 := (iff #13033 #14848)
-#14846 := (iff #12850 #14845)
-#14843 := (iff #12845 #14840)
-#14825 := (or #12431 #11347 #11338 #11329 #11320 #12788)
-#14837 := (or #12431 #12707 #14825)
-#14841 := (iff #14837 #14840)
-#14842 := [rewrite]: #14841
-#14838 := (iff #12845 #14837)
-#14835 := (iff #12820 #14825)
-#14830 := (and true #14825)
-#14833 := (iff #14830 #14825)
-#14834 := [rewrite]: #14833
-#14831 := (iff #12820 #14830)
-#14828 := (iff #12815 #14825)
-#14822 := (or false #12431 #11347 #11338 #11329 #11320 #12788)
-#14826 := (iff #14822 #14825)
-#14827 := [rewrite]: #14826
-#14823 := (iff #12815 #14822)
-#14820 := (iff #11381 false)
-#14818 := (iff #11381 #4808)
-#13920 := (iff #3125 true)
-#9647 := [asserted]: #3125
-#13921 := [iff-true #9647]: #13920
-#14819 := [monotonicity #13921]: #14818
-#14821 := [trans #14819 #11287]: #14820
-#14824 := [monotonicity #14821]: #14823
-#14829 := [trans #14824 #14827]: #14828
-#14832 := [monotonicity #13921 #14829]: #14831
-#14836 := [trans #14832 #14834]: #14835
-#14839 := [monotonicity #14836]: #14838
-#14844 := [trans #14839 #14842]: #14843
-#14847 := [monotonicity #14844]: #14846
-#14816 := (iff #11409 #14815)
-#14813 := (iff #4812 #4811)
-#14808 := (and #4811 true)
-#14811 := (iff #14808 #4811)
-#14812 := [rewrite]: #14811
-#14809 := (iff #4812 #14808)
-#14790 := (iff #4686 true)
-#14791 := [iff-true #13258]: #14790
-#14810 := [monotonicity #14791]: #14809
-#14814 := [trans #14810 #14812]: #14813
-#14817 := [monotonicity #14814]: #14816
-#14850 := [monotonicity #14817 #14847]: #14849
-#14853 := [monotonicity #14850]: #14852
-#14856 := [monotonicity #14853]: #14855
-#14859 := [monotonicity #14856]: #14858
-#14862 := [monotonicity #14859]: #14861
-#14865 := [monotonicity #14862]: #14864
-#14868 := [monotonicity #14865]: #14867
-#14871 := [monotonicity #14868]: #14870
-#14874 := [monotonicity #14871]: #14873
-#14877 := [monotonicity #14874]: #14876
-#14880 := [monotonicity #14877]: #14879
-#14883 := [monotonicity #14880]: #14882
-#14886 := [monotonicity #14883]: #14885
-#13270 := [not-or-elim #13236]: #13269
-#14887 := [mp #13270 #14886]: #14884
-#13371 := [mp #14887 #13370]: #13372
-#17233 := [mp~ #13371 #17232]: #17230
-#17234 := [mp #17233 #17539]: #17537
-#21072 := [mp #17234 #21071]: #21069
-#21904 := [mp #21072 #21903]: #21901
-#25238 := [unit-resolution #21904 #23791]: #21898
-#22129 := (or #21895 #21889)
-#22130 := [def-axiom]: #22129
-#25239 := [unit-resolution #22130 #25238]: #21889
-#22125 := (or #21892 #16842 #16845 #21886)
-#22126 := [def-axiom]: #22125
-#25240 := [unit-resolution #22126 #24267 #24858 #25239]: #21886
-#22115 := (or #21883 #21877)
-#22116 := [def-axiom]: #22115
-#25241 := [unit-resolution #22116 #25240]: #21877
-#22109 := (or #21880 #16842 #16851 #21874)
-#22110 := [def-axiom]: #22109
-#25243 := [unit-resolution #22110 #24267 #25241]: #25242
-#25244 := [unit-resolution #25243 #24533]: #21874
-#22091 := (or #21871 #4750)
-#22092 := [def-axiom]: #22091
-#25245 := [unit-resolution #22092 #25244]: #4750
-#25427 := [mp #25245 #25426]: #4780
-#22099 := (or #21871 #21865)
-#22100 := [def-axiom]: #22099
-#25428 := [unit-resolution #22100 #25244]: #21865
-#25429 := (or #21868 #21862)
-#24430 := [hypothesis]: #12382
-#24463 := [th-lemma arith farkas 1 1 #13247 #24430]: false
-#24464 := [lemma #24463]: #12379
-#22085 := (or #21868 #12382 #21862)
-#22086 := [def-axiom]: #22085
-#25430 := [unit-resolution #22086 #24464]: #25429
-#25431 := [unit-resolution #25430 #25428]: #21862
-#22077 := (or #21859 #21853)
-#22078 := [def-axiom]: #22077
-#25432 := [unit-resolution #22078 #25431]: #21853
-#25227 := (= f461 #16882)
-#25248 := (= #4749 #16882)
-#25246 := (= #16882 #4749)
-#25236 := (= #16881 #4736)
-#25234 := (= #16880 #4735)
-#25232 := (= ?v0!13 0::Int)
-#21324 := (not #16879)
-#25229 := [hypothesis]: #20709
-#21355 := (or #20704 #21324)
-#21358 := [def-axiom]: #21355
-#25230 := [unit-resolution #21358 #25229]: #21324
-#21402 := (or #20704 #16875)
-#21382 := [def-axiom]: #21402
-#25231 := [unit-resolution #21382 #25229]: #16875
-#25233 := [th-lemma arith eq-propagate 0 0 #25231 #25230]: #25232
-#25235 := [monotonicity #25233]: #25234
-#25237 := [monotonicity #25235]: #25236
-#25247 := [monotonicity #25237]: #25246
-#25249 := [symm #25247]: #25248
-#25250 := [trans #25245 #25249]: #25227
-#21334 := (not #16885)
-#21317 := (or #20704 #21334)
-#21335 := [def-axiom]: #21317
-#25251 := [unit-resolution #21335 #25229]: #21334
-#25252 := (not #25227)
-#25253 := (or #25252 #16885)
-#25254 := [th-lemma arith triangle-eq]: #25253
-#25255 := [unit-resolution #25254 #25251 #25250]: false
-#25256 := [lemma #25255]: #20704
-#22073 := (or #21856 #20709 #21850)
-#22074 := [def-axiom]: #22073
-#25433 := [unit-resolution #22074 #25256 #25432]: #21850
-#22065 := (or #21847 #21841)
-#22066 := [def-axiom]: #22065
-#25434 := [unit-resolution #22066 #25433]: #21841
-#25435 := (or #21844 #16901 #21838)
-#22061 := (or #21844 #12418 #16901 #21838)
-#22062 := [def-axiom]: #22061
-#25436 := [unit-resolution #22062 #13247]: #25435
-#25437 := [unit-resolution #25436 #25434 #25427]: #21838
-#22051 := (or #21835 #21829)
-#22052 := [def-axiom]: #22051
-#25701 := [unit-resolution #22052 #25437]: #21829
-#24559 := (+ f462 #17462)
-#24560 := (>= #24559 0::Int)
-#24547 := (+ f464 #17449)
-#24548 := (<= #24547 0::Int)
-#25467 := (not #24548)
-#21964 := (not #17451)
-#25470 := [hypothesis]: #21826
-#21999 := (or #21823 #21817)
-#22000 := [def-axiom]: #21999
-#25471 := [unit-resolution #22000 #25470]: #21817
-#22017 := (or #21835 #13550)
-#22018 := [def-axiom]: #22017
-#25472 := [unit-resolution #22018 #25437]: #13550
-#22033 := (or #21835 #4806)
-#22034 := [def-axiom]: #22033
-#25473 := [unit-resolution #22034 #25437]: #4806
-#22031 := (or #21835 #12880)
-#22032 := [def-axiom]: #22031
-#25474 := [unit-resolution #22032 #25437]: #12880
-#22025 := (or #21835 #12426)
-#22026 := [def-axiom]: #22025
-#25475 := [unit-resolution #22026 #25437]: #12426
-#21997 := (or #21823 #4820)
-#21998 := [def-axiom]: #21997
-#25476 := [unit-resolution #21998 #25470]: #4820
-#24481 := (or #21808 #20989 #20851 #12879 #20993 #11320)
-#24437 := (= #4805 f468)
-#24386 := (= f462 f468)
-#24472 := [hypothesis]: #4820
-#24474 := [symm #24472]: #24386
-#24473 := [hypothesis]: #4806
-#24475 := [trans #24473 #24474]: #24437
-#24476 := [hypothesis]: #21803
-#24477 := [hypothesis]: #12880
-#24478 := [hypothesis]: #12426
-#24479 := [hypothesis]: #13550
-#24438 := (not #24437)
-#24443 := (or #21808 #20851 #20989 #12879 #24438)
-#24330 := (+ f463 #12352)
-#24331 := (>= #24330 0::Int)
-#24439 := (or #20851 #20989 #24331 #24438)
-#24444 := (or #21808 #24439)
-#24451 := (iff #24444 #24443)
-#24440 := (or #20851 #20989 #12879 #24438)
-#24446 := (or #21808 #24440)
-#24449 := (iff #24446 #24443)
-#24450 := [rewrite]: #24449
-#24447 := (iff #24444 #24446)
-#24441 := (iff #24439 #24440)
-#24343 := (iff #24331 #12879)
-#24335 := (+ #12352 f463)
-#24338 := (>= #24335 0::Int)
-#24341 := (iff #24338 #12879)
-#24342 := [rewrite]: #24341
-#24339 := (iff #24331 #24338)
-#24336 := (= #24330 #24335)
-#24337 := [rewrite]: #24336
-#24340 := [monotonicity #24337]: #24339
-#24344 := [trans #24340 #24342]: #24343
-#24442 := [monotonicity #24344]: #24441
-#24448 := [monotonicity #24442]: #24447
-#24452 := [trans #24448 #24450]: #24451
-#24445 := [quant-inst #4786]: #24444
-#24453 := [mp #24445 #24452]: #24443
-#24480 := [unit-resolution #24453 #24479 #24478 #24477 #24476 #24475]: false
-#24482 := [lemma #24480]: #24481
-#25477 := [unit-resolution #24482 #25476 #25475 #25474 #25473 #25472]: #21808
-#21975 := (or #21811 #21803)
-#21976 := [def-axiom]: #21975
-#25478 := [unit-resolution #21976 #25477]: #21811
-#21983 := (or #21820 #20928 #21814)
-#21984 := [def-axiom]: #21983
-#25479 := [unit-resolution #21984 #25478 #25471]: #20928
-#21965 := (or #20923 #21964)
-#21966 := [def-axiom]: #21965
-#25480 := [unit-resolution #21966 #25479]: #21964
-#21985 := (or #21823 #12706)
-#21986 := [def-axiom]: #21985
-#25481 := [unit-resolution #21986 #25470]: #12706
-#25468 := (or #25467 #12707 #17451)
-#25463 := [hypothesis]: #21964
-#25464 := [hypothesis]: #12706
-#25465 := [hypothesis]: #24548
-#25466 := [th-lemma arith farkas -1 -1 1 #25465 #25464 #25463]: false
-#25469 := [lemma #25466]: #25468
-#25482 := [unit-resolution #25469 #25481 #25480]: #25467
-#25485 := (or #24548 #24560)
-#21962 := (or #20923 #17140)
-#21963 := [def-axiom]: #21962
-#25483 := [unit-resolution #21963 #25479]: #17140
-#21960 := (or #20923 #17139)
-#21961 := [def-axiom]: #21960
-#25484 := [unit-resolution #21961 #25479]: #17139
-#22029 := (or #21835 #21671)
-#22030 := [def-axiom]: #22029
-#25443 := [unit-resolution #22030 #25437]: #21671
-#25377 := (or #21676 #20907 #20908 #24548 #24560)
-#24538 := (+ #17147 #12901)
-#24539 := (<= #24538 0::Int)
-#24530 := (+ ?v0!15 #12447)
-#24531 := (>= #24530 0::Int)
-#24540 := (or #20907 #20908 #24531 #24539)
-#25378 := (or #21676 #24540)
-#25393 := (iff #25378 #25377)
-#24565 := (or #20907 #20908 #24548 #24560)
-#25388 := (or #21676 #24565)
-#25391 := (iff #25388 #25377)
-#25392 := [rewrite]: #25391
-#25389 := (iff #25378 #25388)
-#24566 := (iff #24540 #24565)
-#24563 := (iff #24539 #24560)
-#24553 := (+ #12901 #17147)
-#24556 := (<= #24553 0::Int)
-#24561 := (iff #24556 #24560)
-#24562 := [rewrite]: #24561
-#24557 := (iff #24539 #24556)
-#24554 := (= #24538 #24553)
-#24555 := [rewrite]: #24554
-#24558 := [monotonicity #24555]: #24557
-#24564 := [trans #24558 #24562]: #24563
-#24551 := (iff #24531 #24548)
-#24541 := (+ #12447 ?v0!15)
-#24544 := (>= #24541 0::Int)
-#24549 := (iff #24544 #24548)
-#24550 := [rewrite]: #24549
-#24545 := (iff #24531 #24544)
-#24542 := (= #24530 #24541)
-#24543 := [rewrite]: #24542
-#24546 := [monotonicity #24543]: #24545
-#24552 := [trans #24546 #24550]: #24551
-#24567 := [monotonicity #24552 #24564]: #24566
-#25390 := [monotonicity #24567]: #25389
-#25394 := [trans #25390 #25392]: #25393
-#25387 := [quant-inst #17138]: #25378
-#25395 := [mp #25387 #25394]: #25377
-#25486 := [unit-resolution #25395 #25443 #25484 #25483]: #25485
-#25487 := [unit-resolution #25486 #25482]: #24560
-#21967 := (not #17464)
-#21968 := (or #20923 #21967)
-#21969 := [def-axiom]: #21968
-#25488 := [unit-resolution #21969 #25479]: #21967
-#25343 := (+ f462 #12746)
-#25346 := (<= #25343 0::Int)
-#25489 := [symm #25476]: #24386
-#25490 := (not #24386)
-#25491 := (or #25490 #25346)
-#25492 := [th-lemma arith triangle-eq]: #25491
-#25493 := [unit-resolution #25492 #25489]: #25346
-#25494 := [th-lemma arith farkas -1 -1 1 #25493 #25488 #25487]: false
-#25495 := [lemma #25494]: #21823
-#22007 := (or #21832 #21792 #21826)
-#22008 := [def-axiom]: #22007
-#25702 := [unit-resolution #22008 #25495 #25701]: #21792
-#21954 := (or #21789 #12707)
-#21955 := [def-axiom]: #21954
-#25703 := [unit-resolution #21955 #25702]: #12707
-#22019 := (or #21835 #12929)
-#22020 := [def-axiom]: #22019
-#25704 := [unit-resolution #22020 #25437]: #12929
-#25505 := (or #24132 #22212 #22593 #23872 #20990 #12706 #25569)
-#25564 := (+ f464 #12352)
-#25565 := (>= #25564 0::Int)
-#25570 := (or #22212 #22593 #23872 #20990 #25565 #25569)
-#25510 := (or #24132 #25570)
-#25678 := (iff #25510 #25505)
-#25581 := (or #22212 #22593 #23872 #20990 #12706 #25569)
-#25617 := (or #24132 #25581)
-#25627 := (iff #25617 #25505)
-#25677 := [rewrite]: #25627
-#25618 := (iff #25510 #25617)
-#25582 := (iff #25570 #25581)
-#25579 := (iff #25565 #12706)
-#25571 := (+ #12352 f464)
-#25574 := (>= #25571 0::Int)
-#25577 := (iff #25574 #12706)
-#25578 := [rewrite]: #25577
-#25575 := (iff #25565 #25574)
-#25572 := (= #25564 #25571)
-#25573 := [rewrite]: #25572
-#25576 := [monotonicity #25573]: #25575
-#25580 := [trans #25576 #25578]: #25579
-#25583 := [monotonicity #25580]: #25582
-#25626 := [monotonicity #25583]: #25618
-#25679 := [trans #25626 #25677]: #25678
-#25531 := [quant-inst #4649 #4655 #23197 #4646 #4790 #20]: #25510
-#25680 := [mp #25531 #25679]: #25505
-#25705 := [unit-resolution #25680 #20061 #9865 #13258 #25704 #25703 #24213 #25700]: false
-#25706 := [lemma #25705]: #25569
-#25399 := (or #25568 #4935)
-#25400 := [def-axiom]: #25399
-#25731 := [unit-resolution #25400 #25706]: #4935
-#26020 := (= #25632 #4934)
-#26018 := (= #25423 #4930)
-#24154 := (f120 f121 #23775)
-#25370 := (f107 #24154 f464)
-#25379 := (f106 #25370 f14)
-#26016 := (= #25379 #4930)
-#25593 := (= #4930 #25379)
-#25591 := (= #4929 #25370)
-#25584 := (= #25370 #4929)
-#25588 := (= #24154 #4734)
-#25586 := (= #23775 #4656)
-#25600 := [symm #25059]: #24241
-#25601 := (= #23775 #23825)
-#25585 := [trans #24884 #24240]: #25601
-#25587 := [trans #25585 #25600]: #25586
-#25589 := [monotonicity #25587]: #25588
-#25590 := [monotonicity #25589]: #25584
-#25592 := [symm #25590]: #25591
-#25594 := [monotonicity #25592]: #25593
-#26017 := [symm #25594]: #26016
-#26000 := (= #25423 #25379)
-#25396 := (= #25379 #25423)
-#25496 := (not #25396)
-#25397 := (f92 f216 #25379)
-#25403 := (f37 #25397 #23775)
-#25404 := (= #25403 f1)
-#25405 := (not #25404)
-#25499 := (or #25405 #25496)
-#25502 := (not #25499)
-#25352 := (or #24083 #25502)
-#25406 := (* f464 #3690)
-#25407 := (+ #24163 #25406)
-#25410 := (f53 #4654 #25407)
-#25411 := (= #25379 #25410)
-#25409 := (not #25411)
-#25412 := (or #25405 #25409)
-#25408 := (not #25412)
-#24059 := (or #24083 #25408)
-#25364 := (iff #24059 #25352)
-#25367 := (iff #25352 #25352)
-#25368 := [rewrite]: #25367
-#25503 := (iff #25408 #25502)
-#25500 := (iff #25412 #25499)
-#25497 := (iff #25409 #25496)
-#25419 := (iff #25411 #25396)
-#25424 := (= #25410 #25423)
-#25417 := (= #25407 #25416)
-#25414 := (= #25406 #25413)
-#25415 := [rewrite]: #25414
-#25418 := [monotonicity #25415]: #25417
-#25347 := [monotonicity #25418]: #25424
-#25420 := [monotonicity #25347]: #25419
-#25498 := [monotonicity #25420]: #25497
-#25501 := [monotonicity #25498]: #25500
-#25504 := [monotonicity #25501]: #25503
-#25366 := [monotonicity #25504]: #25364
-#25383 := [trans #25366 #25368]: #25364
-#25365 := [quant-inst #23775 #4790 #20]: #24059
-#25381 := [mp #25365 #25383]: #25352
-#25598 := [unit-resolution #25381 #19597]: #25502
-#25386 := (or #25499 #25396)
-#25398 := [def-axiom]: #25386
-#25599 := [unit-resolution #25398 #25598]: #25396
-#26001 := [symm #25599]: #26000
-#26019 := [trans #26001 #26017]: #26018
-#26021 := [monotonicity #26019]: #26020
-#26022 := [trans #26021 #25731]: #25633
-#25634 := (not #25633)
-#25670 := (or #25634 #25669)
-#25671 := (not #25670)
-#25630 := (f37 #4743 #25423)
-#25631 := (= #25630 f1)
-#25672 := (iff #25631 #25671)
-#25841 := (or #24578 #25672)
-#25833 := [quant-inst #4649 #25423]: #25841
-#25839 := [unit-resolution #25833 #20466]: #25672
-#25972 := (not #25631)
-#25982 := (iff #16964 #25972)
-#25996 := (iff #4938 #25631)
-#25628 := (iff #25631 #4938)
-#25993 := (= #25630 #4937)
-#25994 := [monotonicity #26019]: #25993
-#25995 := [monotonicity #25994]: #25628
-#25997 := [symm #25995]: #25996
-#25983 := [monotonicity #25997]: #25982
-#25840 := [hypothesis]: #16964
-#25981 := [mp #25840 #25983]: #25972
-#25821 := (not #25672)
-#25822 := (or #25821 #25631 #25670)
-#25971 := [def-axiom]: #25822
-#25629 := [unit-resolution #25971 #25981 #25839]: #25670
-#25963 := (or #25671 #25634 #25669)
-#25563 := [def-axiom]: #25963
-#25991 := [unit-resolution #25563 #25629 #26022]: #25669
-#26064 := (= #25636 #22576)
-#25401 := (= #25635 f14)
-#25376 := (f27 f28 #4930)
-#25349 := (= #25376 f14)
-#25369 := (iff #4933 #25349)
-#24058 := (or #23224 #25369)
-#24053 := [quant-inst #4930 #20]: #24058
-#25534 := [unit-resolution #24053 #21606]: #25369
-#24057 := (not #25369)
-#25992 := (or #24057 #25349)
-#25508 := (or #23239 #25401)
-#25509 := [quant-inst #20 #25416]: #25508
-#25597 := [unit-resolution #25509 #21619]: #25401
-#25610 := (= #25376 #25635)
-#25611 := (= #4930 #25423)
-#25612 := [trans #25594 #25599]: #25611
-#25613 := [monotonicity #25612]: #25610
-#25614 := [trans #25613 #25597]: #25349
-#24054 := (not #25349)
-#25536 := (or #24057 #24054)
-#25535 := [hypothesis]: #16955
-#23996 := (or #24057 #4933 #24054)
-#25363 := [def-axiom]: #23996
-#25595 := [unit-resolution #25363 #25535]: #25536
-#25596 := [unit-resolution #25595 #25534]: #24054
-#25615 := [unit-resolution #25596 #25614]: false
-#25616 := [lemma #25615]: #4933
-#25689 := (or #24057 #16955 #25349)
-#25692 := [def-axiom]: #25689
-#26003 := [unit-resolution #25692 #25616]: #25992
-#26004 := [unit-resolution #26003 #25534]: #25349
-#26002 := (= #25635 #25376)
-#26005 := [monotonicity #26019]: #26002
-#26063 := [trans #26005 #26004]: #25401
-#26065 := [monotonicity #26063]: #26064
-#25959 := [trans #26065 #24323]: #25637
-#25380 := (not #25567)
-#26060 := (iff #25380 #25642)
-#26066 := (iff #25567 #25641)
-#25950 := (iff #25641 #25567)
-#25948 := (= #25640 #25566)
-#26055 := (= #25639 #24364)
-#25958 := (= #24364 #25639)
-#26068 := [monotonicity #25612]: #25958
-#26056 := [symm #26068]: #26055
-#25949 := [monotonicity #26056]: #25948
-#26058 := [monotonicity #25949]: #25950
-#26067 := [symm #26058]: #26066
-#26061 := [monotonicity #26067]: #26060
-#25382 := (or #25568 #25380)
-#25515 := [def-axiom]: #25382
-#25960 := [unit-resolution #25515 #25706]: #25380
-#26059 := [mp #25960 #26061]: #25642
-#25834 := (or #25647 #25641)
-#25850 := [def-axiom]: #25834
-#26054 := [unit-resolution #25850 #26059]: #25647
-#26062 := (or #25659 #25638 #25648)
-#26337 := (+ #24674 #25413)
-#26535 := (= #25416 #26337)
-#26536 := (* -1::Int #26337)
-#26537 := (+ #25416 #26536)
-#26538 := (<= #26537 0::Int)
-#24353 := (* -1::Int #23755)
-#24356 := (+ #22274 #24353)
-#24358 := (>= #24356 0::Int)
-#24352 := (= #22274 #23755)
-#26548 := (= #4657 #23755)
-#26546 := (= #23755 #4657)
-#26545 := [trans #24240 #25600]: #24243
-#26547 := [monotonicity #26545]: #26546
-#26549 := [symm #26547]: #26548
-#26550 := [trans #25035 #26549]: #24352
-#26551 := (not #24352)
-#26573 := (or #26551 #24358)
-#26574 := [th-lemma arith triangle-eq]: #26573
-#26575 := [unit-resolution #26574 #26550]: #24358
-#25314 := (* -1::Int #24163)
-#25315 := (+ #23755 #25314)
-#25317 := (>= #25315 0::Int)
-#25313 := (= #23755 #24163)
-#26555 := (= #24163 #23755)
-#26556 := [monotonicity #24884]: #26555
-#26557 := [symm #26556]: #25313
-#26558 := (not #25313)
-#26576 := (or #26558 #25317)
-#26577 := [th-lemma arith triangle-eq]: #26576
-#26578 := [unit-resolution #26577 #26557]: #25317
-#26442 := (* -1::Int #24674)
-#26443 := (+ #22274 #26442)
-#26444 := (<= #26443 0::Int)
-#26437 := (= #22274 #24674)
-#26562 := [symm #25037]: #26437
-#26563 := (not #26437)
-#26579 := (or #26563 #26444)
-#26580 := [th-lemma arith triangle-eq]: #26579
-#26581 := [unit-resolution #26580 #26562]: #26444
-#26584 := (not #24358)
-#26583 := (not #26444)
-#26582 := (not #25317)
-#26585 := (or #26538 #26582 #26583 #26584)
-#26586 := [th-lemma arith assign-bounds 1 -1 1]: #26585
-#26587 := [unit-resolution #26586 #26581 #26578 #26575]: #26538
-#26539 := (>= #26537 0::Int)
-#24357 := (<= #24356 0::Int)
-#26552 := (or #26551 #24357)
-#26553 := [th-lemma arith triangle-eq]: #26552
-#26554 := [unit-resolution #26553 #26550]: #24357
-#25316 := (<= #25315 0::Int)
-#26559 := (or #26558 #25316)
-#26560 := [th-lemma arith triangle-eq]: #26559
-#26561 := [unit-resolution #26560 #26557]: #25316
-#26457 := (>= #26443 0::Int)
-#26564 := (or #26563 #26457)
-#26565 := [th-lemma arith triangle-eq]: #26564
-#26566 := [unit-resolution #26565 #26562]: #26457
-#26569 := (not #24357)
-#26568 := (not #26457)
-#26567 := (not #25316)
-#26570 := (or #26539 #26567 #26568 #26569)
-#26571 := [th-lemma arith assign-bounds 1 -1 1]: #26570
-#26572 := [unit-resolution #26571 #26566 #26561 #26554]: #26539
-#26589 := (not #26539)
-#26588 := (not #26538)
-#26590 := (or #26535 #26588 #26589)
-#26591 := [th-lemma arith triangle-eq]: #26590
-#26172 := [unit-resolution #26591 #26572 #26587]: #26535
-#26692 := (not #26535)
-#26693 := (or #26692 #25653)
-#26688 := (= #25652 #4662)
-#26629 := (= #25643 #4658)
-#26627 := (= #25643 #23868)
-#26231 := (f107 #24904 f464)
-#26232 := (f106 #26231 f14)
-#26233 := (f101 #4876 #26232)
-#26234 := (f208 f209 #26233)
-#26235 := (= #26234 #23868)
-#26242 := (f37 #4667 #26232)
-#26243 := (= #26242 f1)
-#26244 := (not #26243)
-#26239 := (f122 f123 #26233)
-#26240 := (= #26239 f1)
-#26241 := (not #26240)
-#26237 := (f122 f210 #26233)
-#26238 := (= #26237 f1)
-#26236 := (not #26235)
-#26245 := (or #26236 #26238 #26241 #26244)
-#26246 := (not #26245)
-#26164 := (or #24899 #24903 #20990 #12706 #26246)
-#26247 := (or #24903 #20990 #25565 #26246)
-#26165 := (or #24899 #26247)
-#25624 := (iff #26165 #26164)
-#26248 := (or #24903 #20990 #12706 #26246)
-#25804 := (or #24899 #26248)
-#25619 := (iff #25804 #26164)
-#25620 := [rewrite]: #25619
-#25961 := (iff #26165 #25804)
-#26249 := (iff #26247 #26248)
-#26250 := [monotonicity #25580]: #26249
-#25962 := [monotonicity #26250]: #25961
-#26136 := [trans #25962 #25620]: #25624
-#26197 := [quant-inst #4649 #4655 #20 #4646 #4790]: #26165
-#26189 := [mp #26197 #26136]: #26164
-#26541 := [unit-resolution #26189 #19381 #25704 #25703 #24970]: #26246
-#26190 := (or #26245 #26235)
-#26201 := [def-axiom]: #26190
-#26542 := [unit-resolution #26201 #26541]: #26235
-#26625 := (= #25643 #26234)
-#26623 := (= #25639 #26233)
-#26621 := (= #26233 #25639)
-#26619 := (= #26232 #25423)
-#26617 := (= #26232 #25379)
-#26615 := (= #26232 #4930)
-#24594 := (f20 f179 #4930)
-#25354 := (f53 #4654 #24594)
-#26609 := (= #25354 #4930)
-#25355 := (= #4930 #25354)
-#25684 := (or #23214 #16955 #25355)
-#25375 := (or #16955 #25355)
-#25685 := (or #23214 #25375)
-#25688 := (iff #25685 #25684)
-#25690 := [rewrite]: #25688
-#25687 := [quant-inst #4930 #20]: #25685
-#25691 := [mp #25687 #25690]: #25684
-#26543 := [unit-resolution #25691 #16676 #25616]: #25355
-#26610 := [symm #26543]: #26609
-#26613 := (= #26232 #25354)
-#26340 := (f53 #4654 #26337)
-#26607 := (= #26340 #25354)
-#26597 := (= #26337 #24594)
-#26595 := (= #25416 #24594)
-#26530 := (= #24594 #25416)
-#26531 := (* -1::Int #25416)
-#26532 := (+ #24594 #26531)
-#26533 := (<= #26532 0::Int)
-#25602 := (f20 f179 #25423)
-#25607 := (* -1::Int #25602)
-#25608 := (+ #25413 #25607)
-#25609 := (+ #24163 #25608)
-#25724 := (>= #25609 0::Int)
-#25605 := (= #25609 0::Int)
-#25699 := (or #23244 #25605)
-#25603 := (= #25602 #25416)
-#25707 := (or #23244 #25603)
-#25717 := (iff #25707 #25699)
-#25719 := (iff #25699 #25699)
-#25720 := [rewrite]: #25719
-#25604 := (iff #25603 #25605)
-#25606 := [rewrite]: #25604
-#25718 := [monotonicity #25606]: #25717
-#25721 := [trans #25718 #25720]: #25717
-#25716 := [quant-inst #20 #25416]: #25707
-#25722 := [mp #25716 #25721]: #25699
-#26321 := [unit-resolution #25722 #21613]: #25605
-#26301 := (not #25605)
-#26304 := (or #26301 #25724)
-#26303 := [th-lemma arith triangle-eq]: #26304
-#26305 := [unit-resolution #26303 #26321]: #25724
-#25726 := (+ #24594 #25607)
-#25729 := (<= #25726 0::Int)
-#25725 := (= #24594 #25602)
-#26306 := (= #25602 #24594)
-#26322 := [monotonicity #26019]: #26306
-#26364 := [symm #26322]: #25725
-#26365 := (not #25725)
-#26355 := (or #26365 #25729)
-#26424 := [th-lemma arith triangle-eq]: #26355
-#26276 := [unit-resolution #26424 #26364]: #25729
-#26290 := (not #25724)
-#26361 := (not #25729)
-#26362 := (or #26533 #26361 #26290)
-#26363 := [th-lemma arith assign-bounds -1 1]: #26362
-#26282 := [unit-resolution #26363 #26276 #26305]: #26533
-#26534 := (>= #26532 0::Int)
-#25723 := (<= #25609 0::Int)
-#26366 := (or #26301 #25723)
-#26281 := [th-lemma arith triangle-eq]: #26366
-#26360 := [unit-resolution #26281 #26321]: #25723
-#25730 := (>= #25726 0::Int)
-#26368 := (or #26365 #25730)
-#26359 := [th-lemma arith triangle-eq]: #26368
-#26369 := [unit-resolution #26359 #26364]: #25730
-#26357 := (not #25723)
-#26356 := (not #25730)
-#26223 := (or #26534 #26356 #26357)
-#26371 := [th-lemma arith assign-bounds -1 1]: #26223
-#26367 := [unit-resolution #26371 #26369 #26360]: #26534
-#26370 := (not #26534)
-#26372 := (not #26533)
-#26421 := (or #26530 #26372 #26370)
-#26378 := [th-lemma arith triangle-eq]: #26421
-#26379 := [unit-resolution #26378 #26367 #26282]: #26530
-#26676 := [symm #26379]: #26595
-#26593 := (= #26337 #25416)
-#26674 := [hypothesis]: #26535
-#26675 := [symm #26674]: #26593
-#26677 := [trans #26675 #26676]: #26597
-#26678 := [monotonicity #26677]: #26607
-#26611 := (= #26232 #26340)
-#26319 := (f107 #24977 f464)
-#26320 := (f106 #26319 f14)
-#26343 := (= #26320 #26340)
-#26346 := (not #26343)
-#26327 := (f92 f216 #26320)
-#26328 := (f37 #26327 #23197)
-#26329 := (= #26328 f1)
-#26330 := (not #26329)
-#26349 := (or #26330 #26346)
-#26352 := (not #26349)
-#26326 := (or #24083 #26352)
-#26331 := (+ #24674 #25406)
-#26332 := (f53 #4654 #26331)
-#26333 := (= #26320 #26332)
-#26334 := (not #26333)
-#26335 := (or #26330 #26334)
-#26336 := (not #26335)
-#26358 := (or #24083 #26336)
-#26413 := (iff #26358 #26326)
-#26449 := (iff #26326 #26326)
-#26455 := [rewrite]: #26449
-#26353 := (iff #26336 #26352)
-#26350 := (iff #26335 #26349)
-#26347 := (iff #26334 #26346)
-#26344 := (iff #26333 #26343)
-#26341 := (= #26332 #26340)
-#26338 := (= #26331 #26337)
-#26339 := [monotonicity #25415]: #26338
-#26342 := [monotonicity #26339]: #26341
-#26345 := [monotonicity #26342]: #26344
-#26348 := [monotonicity #26345]: #26347
-#26351 := [monotonicity #26348]: #26350
-#26354 := [monotonicity #26351]: #26353
-#26414 := [monotonicity #26354]: #26413
-#26426 := [trans #26414 #26455]: #26413
-#26412 := [quant-inst #23197 #4790 #20]: #26358
-#26429 := [mp #26412 #26426]: #26326
-#26599 := [unit-resolution #26429 #19597]: #26352
-#26438 := (or #26349 #26343)
-#26439 := [def-axiom]: #26438
-#26600 := [unit-resolution #26439 #26599]: #26343
-#26605 := (= #26232 #26320)
-#26603 := (= #26231 #26319)
-#26601 := (= #26319 #26231)
-#26602 := [monotonicity #25049]: #26601
-#26604 := [symm #26602]: #26603
-#26606 := [monotonicity #26604]: #26605
-#26612 := [trans #26606 #26600]: #26611
-#26679 := [trans #26612 #26678]: #26613
-#26680 := [trans #26679 #26610]: #26615
-#26681 := [trans #26680 #25594]: #26617
-#26682 := [trans #26681 #25599]: #26619
-#26683 := [monotonicity #26682]: #26621
-#26684 := [symm #26683]: #26623
-#26685 := [monotonicity #26684]: #26625
-#26686 := [trans #26685 #26542]: #26627
-#26687 := [trans #26686 #24768]: #26629
-#26689 := [monotonicity #26687]: #26688
-#26690 := [trans #26689 #13250]: #25653
-#25784 := (not #25653)
-#26673 := [hypothesis]: #25784
-#26691 := [unit-resolution #26673 #26690]: false
-#26694 := [lemma #26691]: #26693
-#26057 := [unit-resolution #26694 #26172]: #25653
-#25854 := (or #25656 #25784)
-#25858 := [def-axiom]: #25854
-#26075 := [unit-resolution #25858 #26057]: #25656
-#25823 := (not #25651)
-#25831 := (f7 f45 #25649)
-#25832 := (= #25831 f1)
-#25842 := (not #25832)
-#25793 := (or #25651 #25842)
-#25794 := (not #25793)
-#25819 := [hypothesis]: #25793
-#25965 := (or #24322 #25794)
-#25966 := [quant-inst #25639]: #25965
-#25820 := [unit-resolution #25966 #20628 #25819]: false
-#25978 := [lemma #25820]: #25794
-#25783 := (or #25793 #25823)
-#25785 := [def-axiom]: #25783
-#26076 := [unit-resolution #25785 #25978]: #25823
-#25902 := (or #25659 #25638 #25648 #25651 #25657)
-#25903 := [def-axiom]: #25902
-#26079 := [unit-resolution #25903 #26076 #26075]: #26062
-#26080 := [unit-resolution #26079 #26054 #25959]: #25659
-#25951 := (or #25668 #25658)
-#25952 := [def-axiom]: #25951
-#26084 := [unit-resolution #25952 #26080 #25991]: false
-#26077 := [lemma #26084]: #4938
-#26462 := (= f464 ?v0!14)
-#26496 := (not #26462)
-#26464 := (= #4940 #17025)
-#26470 := (not #26464)
-#26469 := (+ #4940 #17327)
-#26471 := (>= #26469 0::Int)
-#26480 := (not #26471)
-#25802 := (+ #4940 #12480)
-#25803 := (<= #25802 0::Int)
-#26651 := [hypothesis]: #12613
-#21146 := (+ f462 #12480)
-#21147 := (<= #21146 0::Int)
-#21143 := (= f462 f470)
-#26522 := (iff #5012 #21143)
-#26520 := (iff #21143 #5012)
-#26521 := [commutativity]: #26520
-#26523 := [symm #26521]: #26522
-#26652 := (or #16964 #21768)
-#21956 := (or #21789 #21783)
-#21957 := [def-axiom]: #21956
-#25732 := [unit-resolution #21957 #25702]: #21783
-#21948 := (or #21786 #16955 #16958 #21780)
-#21949 := [def-axiom]: #21948
-#25733 := [unit-resolution #21949 #25732]: #21783
-#25734 := [unit-resolution #25733 #25731 #25616]: #21780
-#21938 := (or #21777 #21771)
-#21939 := [def-axiom]: #21938
-#25735 := [unit-resolution #21939 #25734]: #21771
-#21932 := (or #21774 #16955 #16964 #21768)
-#21933 := [def-axiom]: #21932
-#26659 := [unit-resolution #21933 #25616 #25735]: #26652
-#26660 := [unit-resolution #26659 #26077]: #21768
-#21922 := (or #21765 #21759)
-#21923 := [def-axiom]: #21922
-#26669 := [unit-resolution #21923 #26660]: #21759
-#21154 := (or #21747 #12612)
-#21156 := [def-axiom]: #21154
-#26670 := [unit-resolution #21156 #26651]: #21747
-#21912 := (or #21762 #21750 #21756)
-#21913 := [def-axiom]: #21912
-#26518 := [unit-resolution #21913 #26670 #26669]: #21756
-#21140 := (or #21753 #5012)
-#21141 := [def-axiom]: #21140
-#26519 := [unit-resolution #21141 #26518]: #5012
-#26524 := [mp #26519 #26523]: #21143
-#26525 := (not #21143)
-#26528 := (or #26525 #21147)
-#26661 := [th-lemma arith triangle-eq]: #26528
-#26662 := [unit-resolution #26661 #26524]: #21147
-#26663 := (not #21147)
-#26664 := (or #25803 #12612 #26663)
-#26665 := [th-lemma arith assign-bounds 1 -1]: #26664
-#26666 := [unit-resolution #26665 #26662 #26651]: #25803
-#21300 := (not #17329)
-#25851 := [hypothesis]: #21720
-#21134 := (or #21753 #21717)
-#21905 := [def-axiom]: #21134
-#26163 := [unit-resolution #21905 #25851]: #21753
-#25791 := (or #21741 #21756)
-#25736 := [hypothesis]: #21744
-#21163 := (or #21741 #21735)
-#21164 := [def-axiom]: #21163
-#25737 := [unit-resolution #21164 #25736]: #21735
-#21169 := (or #21738 #16955 #16958 #21732)
-#21171 := [def-axiom]: #21169
-#25675 := [unit-resolution #21171 #25737 #25616 #25731]: #21732
-#21195 := (or #21729 #4938)
-#21196 := [def-axiom]: #21195
-#25676 := [unit-resolution #21196 #25675]: #4938
-#25786 := [hypothesis]: #21753
-#21157 := (or #21747 #21741)
-#21158 := [def-axiom]: #21157
-#25787 := [unit-resolution #21158 #25736]: #21747
-#25788 := [unit-resolution #21913 #25787 #25786]: #21762
-#25789 := [unit-resolution #21923 #25788]: #21765
-#25790 := [unit-resolution #21933 #25789 #25676 #25616 #25735]: false
-#25792 := [lemma #25790]: #25791
-#25809 := [unit-resolution #25792 #26163]: #21741
-#21207 := (or #21723 #21717)
-#21208 := [def-axiom]: #21207
-#26450 := [unit-resolution #21208 #25851]: #21723
-#26398 := (or #21732 #16964 #21726)
-#21182 := (or #21732 #16955 #16964 #21726)
-#21183 := [def-axiom]: #21182
-#26397 := [unit-resolution #21183 #25616]: #26398
-#26400 := [unit-resolution #26397 #26450 #26077]: #21732
-#21176 := (or #21735 #21729)
-#21188 := [def-axiom]: #21176
-#26401 := [unit-resolution #21188 #26400]: #21735
-#26423 := (or #21744 #21738)
-#21167 := (or #21744 #16955 #16958 #21738)
-#21162 := [def-axiom]: #21167
-#26375 := [unit-resolution #21162 #25616 #25731]: #26423
-#26373 := [unit-resolution #26375 #26401 #25809]: false
-#26409 := [lemma #26373]: #21717
-#26211 := (or #21720 #21714)
-#13588 := (<= f443 4294967295::Int)
-#13587 := (iff #12351 #13588)
-#13594 := (+ 4294967295::Int #12352)
-#13593 := (>= #13594 0::Int)
-#13589 := (iff #13593 #13588)
-#13586 := [rewrite]: #13589
-#13590 := (iff #12351 #13593)
-#13595 := (= #12353 #13594)
-#13592 := [monotonicity #8096]: #13595
-#13591 := [monotonicity #13592]: #13590
-#13584 := [trans #13591 #13586]: #13587
-#13266 := [not-or-elim #13236]: #12356
-#13268 := [and-elim #13266]: #12351
-#13585 := [mp #13268 #13584]: #13588
-#26206 := (not #13588)
-#26207 := (or #13511 #26206 #12706)
-#26208 := [th-lemma arith assign-bounds -1 1]: #26207
-#26195 := [unit-resolution #26208 #25703 #13585]: #13511
-#26209 := (or #20990 #12444)
-#26200 := [th-lemma arith farkas 1 1]: #26209
-#26210 := [unit-resolution #26200 #25704]: #12444
-#21242 := (or #21720 #16993 #16996 #21714)
-#21234 := [def-axiom]: #21242
-#26198 := [unit-resolution #21234 #26210 #26195]: #26211
-#26667 := [unit-resolution #26198 #26409]: #21714
-#21253 := (or #21711 #21705)
-#21261 := [def-axiom]: #21253
-#26668 := [unit-resolution #21261 #26667]: #21705
-#21308 := (>= #12524 -1::Int)
-#21252 := (or #21711 #12523)
-#21254 := [def-axiom]: #21252
-#26671 := [unit-resolution #21254 #26667]: #12523
-#26218 := (or #12527 #21308)
-#26219 := [th-lemma arith triangle-eq]: #26218
-#26672 := [unit-resolution #26219 #26671]: #21308
-#26226 := (not #21308)
-#26695 := (or #12460 #26226)
-#26220 := (or #12460 #26226 #12706)
-#26227 := [th-lemma arith assign-bounds -1 -1]: #26220
-#26696 := [unit-resolution #26227 #25703]: #26695
-#26697 := [unit-resolution #26696 #26672]: #12460
-#21271 := (or #21708 #12465 #21702)
-#21272 := [def-axiom]: #21271
-#26698 := [unit-resolution #21272 #26697 #26668]: #21702
-#21262 := (or #21699 #21693)
-#21264 := [def-axiom]: #21262
-#26699 := [unit-resolution #21264 #26698]: #21693
-#26707 := [symm #26519]: #21143
-#26708 := (= #4983 f462)
-#26705 := (= #4983 #4805)
-#26703 := (= #4982 #4804)
-#26701 := (= #4981 #4803)
-#21137 := (or #21753 #5013)
-#21142 := [def-axiom]: #21137
-#26700 := [unit-resolution #21142 #26518]: #5013
-#26702 := [monotonicity #26700]: #26701
-#26704 := [monotonicity #26702]: #26703
-#26706 := [monotonicity #26704]: #26705
-#26709 := [trans #26706 #25473]: #26708
-#26710 := [trans #26709 #26707]: #4984
-#21152 := (+ f463 #12502)
-#21153 := (>= #21152 0::Int)
-#21151 := (= f463 f471)
-#26713 := (iff #5013 #21151)
-#26711 := (iff #21151 #5013)
-#26712 := [commutativity]: #26711
-#26714 := [symm #26712]: #26713
-#26715 := [mp #26700 #26714]: #21151
-#26716 := (not #21151)
-#26717 := (or #26716 #21153)
-#26718 := [th-lemma arith triangle-eq]: #26717
-#26719 := [unit-resolution #26718 #26715]: #21153
-#26720 := (not #21153)
-#26721 := (or #12505 #12879 #26720)
-#26722 := [th-lemma arith assign-bounds -1 -1]: #26721
-#26723 := [unit-resolution #26722 #26719 #25474]: #12505
-#21289 := (or #20771 #12504 #20769)
-#21281 := [def-axiom]: #21289
-#26724 := [unit-resolution #21281 #26723 #26710]: #20771
-#21286 := (or #21687 #20770)
-#21290 := [def-axiom]: #21286
-#26725 := [unit-resolution #21290 #26724]: #21687
-#21278 := (or #21696 #20755 #21690)
-#21279 := [def-axiom]: #21278
-#26726 := [unit-resolution #21279 #26725 #26699]: #20755
-#21303 := (or #20750 #21300)
-#21301 := [def-axiom]: #21303
-#26727 := [unit-resolution #21301 #26726]: #21300
-#26481 := (not #25803)
-#26482 := (or #26480 #17329 #26481)
-#26476 := [hypothesis]: #26471
-#26477 := [hypothesis]: #25803
-#26478 := [hypothesis]: #21300
-#26479 := [th-lemma arith farkas -1 -1 1 #26478 #26477 #26476]: false
-#26483 := [lemma #26479]: #26482
-#26728 := [unit-resolution #26483 #26727 #26666]: #26480
-#26472 := (or #26470 #26471)
-#26473 := [th-lemma arith triangle-eq]: #26472
-#26729 := [unit-resolution #26473 #26728]: #26470
-#26497 := (or #26496 #26464)
-#26492 := (= #17025 #4940)
-#26490 := (= #17024 #4930)
-#26488 := (= #17023 #4929)
-#26486 := (= ?v0!14 f464)
-#26485 := [hypothesis]: #26462
-#26487 := [symm #26485]: #26486
-#26489 := [monotonicity #26487]: #26488
-#26491 := [monotonicity #26489]: #26490
-#26493 := [monotonicity #26491]: #26492
-#26494 := [symm #26493]: #26464
-#26484 := [hypothesis]: #26470
-#26495 := [unit-resolution #26484 #26494]: false
-#26498 := [lemma #26495]: #26497
-#26730 := [unit-resolution #26498 #26729]: #26496
-#26104 := (+ f464 #17314)
-#26204 := (>= #26104 0::Int)
-#21314 := (not #17316)
-#21293 := (or #20750 #21314)
-#21296 := [def-axiom]: #21293
-#26731 := [unit-resolution #21296 #26726]: #21314
-#26732 := (or #26204 #26226 #17316)
-#26733 := [th-lemma arith assign-bounds -1 -1]: #26732
-#26734 := [unit-resolution #26733 #26731 #26672]: #26204
-#26105 := (<= #26104 0::Int)
-#26116 := (+ f462 #17327)
-#26117 := (>= #26116 0::Int)
-#26277 := (not #26117)
-#26735 := (or #26277 #17329 #26663)
-#26736 := [th-lemma arith assign-bounds -1 -1]: #26735
-#26737 := [unit-resolution #26736 #26662 #26727]: #26277
-#21309 := (or #20750 #17018)
-#21311 := [def-axiom]: #21309
-#26738 := [unit-resolution #21311 #26726]: #17018
-#21312 := (or #20750 #17017)
-#21313 := [def-axiom]: #21312
-#26739 := [unit-resolution #21313 #26726]: #17017
-#26125 := (or #21676 #20734 #20735 #26105 #26117)
-#26095 := (+ #17025 #12901)
-#26096 := (<= #26095 0::Int)
-#26087 := (+ ?v0!14 #12447)
-#26088 := (>= #26087 0::Int)
-#26097 := (or #20734 #20735 #26088 #26096)
-#26126 := (or #21676 #26097)
-#26133 := (iff #26126 #26125)
-#26122 := (or #20734 #20735 #26105 #26117)
-#26128 := (or #21676 #26122)
-#26131 := (iff #26128 #26125)
-#26132 := [rewrite]: #26131
-#26129 := (iff #26126 #26128)
-#26123 := (iff #26097 #26122)
-#26120 := (iff #26096 #26117)
-#26110 := (+ #12901 #17025)
-#26113 := (<= #26110 0::Int)
-#26118 := (iff #26113 #26117)
-#26119 := [rewrite]: #26118
-#26114 := (iff #26096 #26113)
-#26111 := (= #26095 #26110)
-#26112 := [rewrite]: #26111
-#26115 := [monotonicity #26112]: #26114
-#26121 := [trans #26115 #26119]: #26120
-#26108 := (iff #26088 #26105)
-#26098 := (+ #12447 ?v0!14)
-#26101 := (>= #26098 0::Int)
-#26106 := (iff #26101 #26105)
-#26107 := [rewrite]: #26106
-#26102 := (iff #26088 #26101)
-#26099 := (= #26087 #26098)
-#26100 := [rewrite]: #26099
-#26103 := [monotonicity #26100]: #26102
-#26109 := [trans #26103 #26107]: #26108
-#26124 := [monotonicity #26109 #26121]: #26123
-#26130 := [monotonicity #26124]: #26129
-#26134 := [trans #26130 #26132]: #26133
-#26127 := [quant-inst #17016]: #26126
-#26135 := [mp #26127 #26134]: #26125
-#26740 := [unit-resolution #26135 #25443 #26739 #26738 #26737]: #26105
-#26323 := (not #26204)
-#26302 := (not #26105)
-#26310 := (or #26462 #26302 #26323)
-#26311 := [th-lemma arith triangle-eq]: #26310
-#26741 := [unit-resolution #26311 #26740 #26734 #26730]: false
-#26742 := [lemma #26741]: #12612
-#26324 := (or #21723 #12613)
-#26205 := [hypothesis]: #21726
-#26203 := [unit-resolution #21208 #26205]: #21717
-#26212 := [unit-resolution #26198 #26203]: #21714
-#26213 := [unit-resolution #21261 #26212]: #21705
-#26216 := [unit-resolution #21254 #26212]: #12523
-#26225 := [unit-resolution #26219 #26216]: #21308
-#26228 := [unit-resolution #26227 #26225 #25703]: #12460
-#26224 := [unit-resolution #21272 #26228 #26213]: #21702
-#26262 := [unit-resolution #21264 #26224]: #21693
-#26269 := (= f469 f470)
-#21221 := (or #21723 #4956)
-#21231 := [def-axiom]: #21221
-#26263 := [unit-resolution #21231 #26205]: #4956
-#26270 := [symm #26263]: #26269
-#26271 := (= #4983 f469)
-#26199 := (= #4940 f469)
-#21226 := (or #21723 #4943)
-#21227 := [def-axiom]: #21226
-#26229 := [unit-resolution #21227 #26205]: #4943
-#26268 := [symm #26229]: #26199
-#26266 := (= #4983 #4940)
-#26264 := (= #4982 #4930)
-#26253 := (= #4981 #4929)
-#21198 := (or #21723 #4958)
-#21200 := [def-axiom]: #21198
-#26252 := [unit-resolution #21200 #26205]: #4958
-#26254 := [monotonicity #26252]: #26253
-#26265 := [monotonicity #26254]: #26264
-#26267 := [monotonicity #26265]: #26266
-#26272 := [trans #26267 #26268]: #26271
-#26289 := [trans #26272 #26270]: #4984
-#25798 := (+ f464 #12502)
-#25800 := (>= #25798 0::Int)
-#25797 := (= f464 f471)
-#26215 := [symm #26252]: #25797
-#26255 := (not #25797)
-#26256 := (or #26255 #25800)
-#26251 := [th-lemma arith triangle-eq]: #26256
-#26257 := [unit-resolution #26251 #26215]: #25800
-#26202 := (not #25800)
-#26258 := (or #12505 #26202 #12706)
-#26259 := [th-lemma arith assign-bounds -1 -1]: #26258
-#26260 := [unit-resolution #26259 #26257 #25703]: #12505
-#26261 := [unit-resolution #21281 #26260 #26289]: #20771
-#26217 := [unit-resolution #21290 #26261]: #21687
-#26293 := [unit-resolution #21279 #26217 #26262]: #20755
-#26284 := [unit-resolution #21296 #26293]: #21314
-#25801 := (= #4940 f470)
-#26285 := [trans #26268 #26270]: #25801
-#26283 := (not #25801)
-#26286 := (or #26283 #25803)
-#26287 := [th-lemma arith triangle-eq]: #26286
-#26288 := [unit-resolution #26287 #26285]: #25803
-#26294 := [unit-resolution #21301 #26293]: #21300
-#26295 := [unit-resolution #26483 #26294 #26288]: #26480
-#26291 := [unit-resolution #26473 #26295]: #26470
-#26275 := [unit-resolution #26498 #26291]: #26496
-#26309 := (or #26462 #26323)
-#26278 := [hypothesis]: #12612
-#26279 := (or #26277 #17329 #26481 #12613)
-#26280 := [th-lemma arith assign-bounds 1 1 1]: #26279
-#26292 := [unit-resolution #26280 #26294 #26288 #26278]: #26277
-#26296 := (or #26105 #26117)
-#26307 := [unit-resolution #21311 #26293]: #17018
-#26308 := [unit-resolution #21313 #26293]: #17017
-#26315 := [unit-resolution #26135 #25443 #26308 #26307]: #26296
-#26316 := [unit-resolution #26315 #26292]: #26105
-#26312 := [unit-resolution #26311 #26316]: #26309
-#26313 := [unit-resolution #26312 #26275]: #26323
-#26314 := [th-lemma arith farkas 1 -1 1 #26225 #26313 #26284]: false
-#26325 := [lemma #26314]: #26324
-#26500 := [unit-resolution #26325 #26742]: #21723
-#26402 := [unit-resolution #26397 #26500 #26077]: #21732
-#21136 := (or #21753 #12613)
-#21139 := [def-axiom]: #21136
-#26445 := [unit-resolution #21139 #26742]: #21753
-#26377 := [unit-resolution #25792 #26445]: #21741
-#26222 := [unit-resolution #26375 #26377]: #21738
-[unit-resolution #21188 #26222 #26402]: false
-unsat
--- a/src/HOL/Boogie/Examples/VCC_Max.thy	Tue Jul 23 13:14:14 2013 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,61 +0,0 @@
-(*  Title:      HOL/Boogie/Examples/VCC_Max.thy
-    Author:     Sascha Boehme, TU Muenchen
-*)
-
-header {* Boogie example from VCC: get the greatest element of a list *}
-
-theory VCC_Max
-imports Boogie
-begin
-
-text {*
-We prove correct the verification condition generated from the following
-C code:
-
-\begin{verbatim}
-#include "vcc.h"
-
-typedef unsigned char byte;
-typedef unsigned long ulong;
-
-static byte maximum(byte arr[], ulong len)
-  requires(wrapped(as_array(arr, len)))
-  requires (0 < len && len < (1UI64 << 40))
-  ensures (forall(ulong i; i<len ==> arr[i] <= result))
-  ensures (exists(ulong i; i<len && arr[i] == result))
-{
-  byte max = arr[0];
-  ulong p;
-  spec(ulong witness = 0;)
-
-  for (p = 1; p < len; p++)
-    invariant(p <= len)
-    invariant(forall(ulong i; i < p ==> arr[i] <= max))
-    invariant(witness < len && arr[witness] == max)
-  {
-    if (arr[p] > max)
-    {
-      max = arr[p];
-      speconly(witness = p;)
-    }
-  }
-  return max;
-}
-\end{verbatim}
-*}
-
-boogie_open (quiet) "VCC_Max.b2i"
-
-declare [[smt_oracle = false]]
-declare [[z3_with_extensions = true]]
-declare [[smt_certificates = "VCC_Max.certs"]]
-declare [[smt_read_only_certificates = true]]
-
-boogie_status
-
-boogie_vc maximum
-  by boogie
-
-boogie_end
-
-end
--- a/src/HOL/Boogie/Tools/boogie_commands.ML	Tue Jul 23 13:14:14 2013 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,355 +0,0 @@
-(*  Title:      HOL/Boogie/Tools/boogie_commands.ML
-    Author:     Sascha Boehme, TU Muenchen
-
-Isar commands to create a Boogie environment simulation.
-*)
-
-signature BOOGIE_COMMANDS =
-sig
-  val setup: theory -> theory
-end
-
-structure Boogie_Commands: BOOGIE_COMMANDS =
-struct
-
-(* commands *)
-
-fun boogie_open ((quiet, files), offsets) thy =
-  let
-    val ([{src_path = path, text, ...}: Token.file], thy') = files thy
-
-    val ext = "b2i"
-    val _ = snd (Path.split_ext path) = ext orelse
-      error ("Bad file ending of file " ^ Path.print path ^ ", " ^
-        "expected file ending " ^ quote ext)
-
-    val _ = Boogie_VCs.is_closed thy orelse
-      error ("Found the beginning of a new Boogie environment, " ^
-        "but another Boogie environment is still open.")
-  in
-    thy'
-    |> Boogie_Loader.parse_b2i (not quiet) offsets text
-  end
-
-
-datatype vc_opts =
-  VC_Complete |
-  VC_Take of int list * (bool * string list) option |
-  VC_Only of string list |
-  VC_Without of string list |
-  VC_Examine of string list |
-  VC_Single of string
-
-fun get_vc thy vc_name =
-  (case Boogie_VCs.lookup thy vc_name of
-    SOME vc => vc
-  | NONE =>
-      (case AList.lookup (op =) (Boogie_VCs.state_of thy) vc_name of
-        SOME Boogie_VCs.Proved => error ("The verification condition " ^
-          quote vc_name ^ " has already been proved.")
-      | _ => error ("There is no verification condition " ^
-          quote vc_name ^ ".")))
-
-local
-  fun split_goal t =
-    (case Boogie_Tactics.split t of
-      [tp] => tp
-    | _ => error "Multiple goals.")
-
-  fun single_prep t =
-    let
-      val (us, u) = split_goal t
-      val assms = [((@{binding vc_trace}, []), map (rpair []) us)]
-    in
-      pair [u] o snd o Proof_Context.add_assms_i Assumption.assume_export assms
-    end
-
-  fun single_prove goal ctxt thm =
-    Goal.prove ctxt [] [] goal (fn {context, ...} => HEADGOAL (
-      Boogie_Tactics.split_tac
-      THEN' Boogie_Tactics.drop_assert_at_tac
-      THEN' SUBPROOF (fn _ => Tactic.rtac thm 1) context))
-in
-fun boogie_vc (vc_name, vc_opts) thy =
-  let
-    val vc = get_vc thy vc_name
-
-    fun extract vc l =
-      (case Boogie_VCs.extract vc l of
-        SOME vc' => vc'
-      | NONE => error ("There is no assertion to be proved with label " ^
-          quote l ^ "."))
-
-    val vcs =
-      (case vc_opts of
-        VC_Complete => [vc]
-      | VC_Take (ps, NONE) => [Boogie_VCs.paths_and ps [] vc]
-      | VC_Take (ps, SOME (true, ls)) => [Boogie_VCs.paths_and ps ls vc]
-      | VC_Take (ps, SOME (false, ls)) => [Boogie_VCs.paths_without ps ls vc]
-      | VC_Only ls => [Boogie_VCs.only ls vc]
-      | VC_Without ls => [Boogie_VCs.without ls vc]
-      | VC_Examine ls => map (extract vc) ls
-      | VC_Single l => [extract vc l])
-    val ts = map Boogie_VCs.prop_of_vc vcs
-
-    val (prepare, finish) =
-      (case vc_opts of
-         VC_Single _ => (single_prep (hd ts), single_prove (hd ts))
-      | _ => (pair ts, K I))
-
-    val discharge = fold (Boogie_VCs.discharge o pair vc_name)
-    fun after_qed [thms] = Proof_Context.background_theory (discharge (vcs ~~ thms))
-      | after_qed _ = I
-  in
-    Proof_Context.init_global thy
-    |> fold Variable.auto_fixes ts
-    |> (fn ctxt1 => ctxt1
-    |> prepare
-    |-> (fn us => fn ctxt2 => ctxt2
-    |> Proof.theorem NONE (fn thmss => fn ctxt =>
-         let val export = map (finish ctxt1) o Proof_Context.export ctxt ctxt2
-         in after_qed (map export thmss) ctxt end) [map (rpair []) us]))
-  end
-end
-
-fun write_list head =
-  map Pretty.str o sort (dict_ord string_ord o pairself raw_explode) #>
-  Pretty.writeln o Pretty.big_list head
-
-fun parens s = "(" ^ s ^ ")"
-
-fun boogie_status thy =
-  let
-    fun string_of_state Boogie_VCs.Proved = "proved"
-      | string_of_state Boogie_VCs.NotProved = "not proved"
-      | string_of_state Boogie_VCs.PartiallyProved = "partially proved"
-  in
-    Boogie_VCs.state_of thy
-    |> map (fn (name, proved) => name ^ " " ^ parens (string_of_state proved))
-    |> write_list "Boogie verification conditions:"
-  end
-
-fun boogie_status_vc full vc_name thy =
-  let
-    fun pretty tag s = s ^ " " ^ parens tag
-
-    val (not_proved, proved) = Boogie_VCs.state_of_vc thy vc_name
-  in
-    if full
-    then write_list ("Assertions of Boogie verification condition " ^
-      quote vc_name ^ ":")
-      (map (pretty "proved") proved @ map (pretty "not proved") not_proved)
-    else write_list ("Unproved assertions of Boogie verification condition " ^
-      quote vc_name ^ ":") not_proved
-  end
-
-fun boogie_status_vc_paths full vc_name thy =
-  let
-    fun labels ls = Pretty.blk (0, Pretty.breaks (map Pretty.str ls))
-
-    fun pp (i, ns) =
-      if full
-      then
-        [Pretty.big_list ("Path " ^ string_of_int (i+1) ^ ":")
-          [labels (map (fn (n, true) => n | (n, _) => parens n) ns)]]
-      else
-        let val ns' = map_filter (fn (n, true) => SOME n | _ => NONE) ns
-        in
-          if null ns' then []
-          else
-            [Pretty.big_list ("Unproved assertions of path " ^
-              string_of_int (i+1) ^ ":") [labels ns']]
-        end
-  in
-    Pretty.writeln
-      (Pretty.big_list
-        ("Paths of Boogie verification condition " ^ quote vc_name ^ ":")
-        (flat (map_index pp (Boogie_VCs.path_names_of (get_vc thy vc_name)))))
-  end
-
-
-local
-  fun trying s = tracing ("Trying " ^ s ^ " ...")
-  fun success_on s = tracing ("Succeeded on " ^ s ^ ".")
-  fun failure_on s c = tracing ("Failed on " ^ s ^ c)
-
-  fun string_of_asserts vc = space_implode ", " (fst (Boogie_VCs.names_of vc))
-
-  fun string_of_path (i, n) =
-    "path " ^ string_of_int i ^ " of " ^ string_of_int n
-
-  fun itemize_paths ps =
-    let val n = length ps
-    in fst (fold_map (fn p => fn i => (((i, n), p), i+1)) ps 1) end
-
-  fun par_map f = flat o Par_List.map f
-
-  fun divide f vc =
-    let val n = Boogie_VCs.size_of vc
-    in
-      if n <= 1 then fst (Boogie_VCs.names_of vc)
-      else
-        let val (vc1, vc2) = the (Boogie_VCs.split_path (n div 2) vc)
-        in par_map f [vc1, vc2] end
-    end
-
-  fun prove thy meth vc =
-    Proof_Context.init_global thy
-    |> Proof.theorem NONE (K I) [[(Boogie_VCs.prop_of_vc vc, [])]]
-    |> Proof.apply meth
-    |> Seq.hd
-    |> Proof.global_done_proof
-in
-fun boogie_narrow_vc (quick, timeout) vc_name meth thy =
-  let
-    fun tp t = TimeLimit.timeLimit (Time.fromSeconds t) (prove thy meth)
-
-    fun try_vc t (tag, split_tag) split vc = (trying tag;
-      (case try (tp t) vc of
-        SOME _ => (success_on tag; [])
-      | NONE => (failure_on tag split_tag; split vc)))
-
-    fun some_asserts vc =
-      let
-        val (t, sep) = if Boogie_VCs.size_of vc = 1 then (timeout, ".")
-          else (quick, ", further splitting ...")
-      in try_vc t (string_of_asserts vc, sep) (divide some_asserts) vc end
-
-    fun single_path p =
-      try_vc quick (string_of_path p, ", splitting into assertions ...")
-        (divide some_asserts)
-
-    val complete_vc =
-      try_vc quick ("full goal", ", splitting into paths ...")
-        (par_map (uncurry single_path) o itemize_paths o Boogie_VCs.paths_of)
-
-    val unsolved = complete_vc (get_vc thy vc_name)
-  in
-    if null unsolved
-    then writeln ("Completely solved Boogie verification condition " ^
-      quote vc_name ^ ".")
-    else write_list ("Unsolved assertions of Boogie verification condition " ^
-      quote vc_name ^ ":") unsolved
-  end
-
-fun boogie_scan_vc timeout vc_name meth thy =
-  let
-    val tp = TimeLimit.timeLimit (Time.fromSeconds timeout) (prove thy meth)
-
-    val vc = get_vc thy vc_name
-    fun prove_assert name =
-      (trying name; tp (the (Boogie_VCs.extract vc name)))
-    val find_first_failure = find_first (is_none o try prove_assert)
-  in
-    (case find_first_failure (fst (Boogie_VCs.names_of vc)) of
-      SOME name => writeln ("failed on " ^ quote name)
-    | NONE => writeln "succeeded")
-  end
-end
-
-
-
-fun boogie_end thy =
-  let
-    fun not_proved (_, Boogie_VCs.Proved) = NONE
-      | not_proved (name, _) = SOME name
-
-    val unproved = map_filter not_proved (Boogie_VCs.state_of thy)
-  in
-    if null unproved then Boogie_VCs.close thy
-    else error (Pretty.string_of (Pretty.big_list
-      "The following verification conditions have not been proved:"
-      (map Pretty.str unproved)))
-  end
-
-
-
-(* syntax and setup *)
-
-fun scan_val n f = Args.$$$ n -- Args.colon |-- f
-fun scan_arg f = Args.parens f
-fun scan_opt n = Scan.optional (scan_arg (Args.$$$ n >> K true)) false
-
-val vc_offsets = Scan.optional (Args.bracks (Parse.list1
-  (Parse.string --| Args.colon -- Parse.nat))) []
-
-val _ =
-  Outer_Syntax.command @{command_spec "boogie_open"}
-    "open a new Boogie environment and load a Boogie-generated .b2i file"
-    (scan_opt "quiet" -- Thy_Load.provide_parse_files "boogie_open" -- vc_offsets >>
-      (Toplevel.theory o boogie_open))
-
-
-val vc_name = Parse.name
-
-val vc_label = Parse.name
-val vc_labels = Scan.repeat1 vc_label
-
-val vc_paths =
-  Parse.nat -- (Args.$$$ "-" |-- Parse.nat) >> (op upto) ||
-  Parse.nat >> single
-
-val vc_opts =
-  scan_arg
-   (scan_val "assertion" vc_label >> VC_Single ||
-    scan_val "examine" vc_labels >> VC_Examine ||
-    scan_val "take" ((Parse.list vc_paths >> flat) -- Scan.option (
-      scan_val "without" vc_labels >> pair false ||
-      scan_val "and_also" vc_labels >> pair true) >> VC_Take) ||
-    scan_val "only" vc_labels >> VC_Only ||
-    scan_val "without" vc_labels >> VC_Without) ||
-  Scan.succeed VC_Complete
-
-val _ =
-  Outer_Syntax.command @{command_spec "boogie_vc"}
-    "enter into proof mode for a specific Boogie verification condition"
-    (vc_name -- vc_opts >> (fn args =>
-      (Toplevel.print o Toplevel.theory_to_proof (boogie_vc args))))
-
-
-val quick_timeout = 5
-val default_timeout = 20
-
-fun timeout name = Scan.optional (scan_val name Parse.nat)
-
-val status_test =
-  scan_arg (
-    Args.$$$ "scan" |-- timeout "timeout" quick_timeout >> boogie_scan_vc ||
-    Args.$$$ "narrow" |-- timeout "step_timeout" quick_timeout --
-      timeout "final_timeout" default_timeout >> boogie_narrow_vc) --
-  vc_name -- Method.parse >>
-  (fn ((f, vc_name), (meth, _)) => f vc_name meth)
-
-val status_vc =
-  (scan_arg
-    (Args.$$$ "full" |--
-      (Args.$$$ "paths" >> K (boogie_status_vc_paths true) ||
-       Scan.succeed (boogie_status_vc true)) ||
-     Args.$$$ "paths" >> K (boogie_status_vc_paths false)) ||
-   Scan.succeed (boogie_status_vc false)) --
-  vc_name >> (fn (f, vc_name) => f vc_name)
-
-fun status_cmd f = Toplevel.keep (f o Toplevel.theory_of)
-
-val _ =
-  Outer_Syntax.improper_command @{command_spec "boogie_status"}
-    "show the name and state of all loaded Boogie verification conditions"
-    (status_test >> status_cmd ||
-     status_vc >> status_cmd ||
-     Scan.succeed (status_cmd boogie_status))
-
-
-val _ =
-  Outer_Syntax.command @{command_spec "boogie_end"}
-    "close the current Boogie environment"
-    (Scan.succeed (Toplevel.theory boogie_end))
-
-
-val setup = Theory.at_end (fn thy =>
-  let
-    val _ = Boogie_VCs.is_closed thy
-      orelse error ("Found the end of the theory, " ^
-        "but the last Boogie environment is still open.")
-  in NONE end)
-
-end
--- a/src/HOL/Boogie/Tools/boogie_loader.ML	Tue Jul 23 13:14:14 2013 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,600 +0,0 @@
-(*  Title:      HOL/Boogie/Tools/boogie_loader.ML
-    Author:     Sascha Boehme, TU Muenchen
-
-Loading and interpreting Boogie-generated files.
-*)
-
-signature BOOGIE_LOADER =
-sig
-  val load_b2i: bool -> (string * int) list -> Path.T -> theory -> theory
-  val parse_b2i: bool -> (string * int) list -> string -> theory -> theory
-end
-
-structure Boogie_Loader: BOOGIE_LOADER =
-struct
-
-fun log verbose text args x =
-  if verbose andalso not (null args)
-  then (Pretty.writeln (Pretty.big_list text (map Pretty.str args)); x)
-  else x
-
-val isabelle_name =
-  let 
-    fun purge s = if Symbol.is_letter s orelse Symbol.is_digit s then s else
-      (case s of
-        "." => "_o_"
-      | "_" => "_n_"
-      | "$" => "_S_"
-      | "@" => "_G_"
-      | "#" => "_H_"
-      | "^" => "_T_"
-      | _   => ("_" ^ string_of_int (ord s) ^ "_"))
-  in prefix "b_" o translate_string purge end
-
-fun drop_underscore s =
-  try (unsuffix "_") s
-  |> Option.map drop_underscore
-  |> the_default s
-
-val short_name =
-  translate_string (fn s => if Symbol.is_letdig s then s else "") #>
-  drop_underscore
-
-(* these prefixes must be distinct: *)
-val var_prefix = "V_"
-val label_prefix = "L_"
-val position_prefix = "P_"
-
-val no_label_name = label_prefix ^ "unknown"
-fun label_name line col =
-  if line = 0 orelse col = 0 then no_label_name
-  else label_prefix ^ string_of_int line ^ "_" ^ string_of_int col
-
-fun mk_syntax name i =
-  let
-    val syn = Syntax_Ext.escape name
-    val args = space_implode ",/ " (replicate i "_")
-  in
-    if i = 0 then Mixfix (syn, [], 1000)
-    else Mixfix (syn ^ "()'(/" ^ args ^ "')", replicate i 0, 1000)
-  end
-
-
-datatype attribute_value = StringValue of string | TermValue of term
-
-
-fun mk_distinct [] = @{const HOL.True}
-  | mk_distinct [_] = @{const HOL.True}
-  | mk_distinct (t :: ts) =
-      let
-        fun mk_noteq u u' =
-          HOLogic.mk_conj (HOLogic.mk_not (HOLogic.mk_eq (t, u)), u')
-      in fold_rev mk_noteq ts (mk_distinct ts) end
-
-
-local
-  fun lookup_type_name thy name arity =
-    let val intern = Sign.intern_type thy name
-    in
-      if Sign.declared_tyname thy intern
-      then
-        if Sign.arity_number thy intern = arity then SOME intern
-        else error ("Boogie: type already declared with different arity: " ^
-          quote name)
-      else NONE
-    end
-
-  fun log_new bname name = bname ^ " (as " ^ name ^ ")"
-  fun log_ex bname name = "[" ^ bname ^ " has already been declared as " ^
-    name ^ "]"
-
-  fun declare (name, arity) thy =
-    let val isa_name = isabelle_name name
-    in
-      (case lookup_type_name thy isa_name arity of
-        SOME type_name => (((name, type_name), log_ex name type_name), thy)
-      | NONE =>
-          let
-            val args = map (rpair dummyS) (Name.invent Name.context "'a" arity)
-            val (T, thy') =
-              Typedecl.typedecl_global (Binding.name isa_name, args, mk_syntax name arity) thy
-            val type_name = fst (Term.dest_Type T)
-          in (((name, type_name), log_new name type_name), thy') end)
-    end
-in
-fun declare_types verbose tys =
-  fold_map declare tys #>> split_list #-> (fn (tds, logs) =>
-  log verbose "Declared types:" logs #>
-  rpair (Symtab.make tds))
-end
-
-
-
-local
-  fun maybe_builtin T =
-    let
-      fun const name = SOME (Const (name, T))
-      fun const2_abs name =
-        let val U = Term.domain_type T
-        in
-          SOME (Abs (Name.uu, U, Abs (Name.uu, U,
-            Const (name, T) $ Bound 0 $ Bound 1)))
-        end
-
-      fun choose builtin =
-        (case builtin of
-          "bvnot" => const @{const_name bitNOT}
-        | "bvand" => const @{const_name bitAND}
-        | "bvor" => const @{const_name bitOR}
-        | "bvxor" => const @{const_name bitXOR}
-        | "bvadd" => const @{const_name plus}
-        | "bvneg" => const @{const_name uminus}
-        | "bvsub" => const @{const_name minus}
-        | "bvmul" => const @{const_name times}
-(* FIXME:
-        | "bvudiv" => const @{const_name div}
-        | "bvurem" => const @{const_name mod}
-        | "bvsdiv" => const @{const_name sdiv}
-        | "bvsrem" => const @{const_name srem}
-        | "bvshl" => const @{const_name bv_shl}
-        | "bvlshr" => const @{const_name bv_lshr}
-        | "bvashr" => const @{const_name bv_ashr}
-*)
-        | "bvult" => const @{const_name less}
-        | "bvule" => const @{const_name less_eq}
-        | "bvugt" => const2_abs @{const_name less}
-        | "bvuge" => const2_abs @{const_name less_eq}
-        | "bvslt" => const @{const_name word_sless}
-        | "bvsle" => const @{const_name word_sle}
-        | "bvsgt" => const2_abs @{const_name word_sless}
-        | "bvsge" => const2_abs @{const_name word_sle}
-        | "zero_extend" => const @{const_name ucast}
-        | "sign_extend" => const @{const_name scast}
-        | _ => NONE)
-
-      fun is_builtin att =
-        (case att of
-          ("bvbuiltin", [StringValue builtin]) => choose builtin
-        | ("bvint", [StringValue "ITE"]) => const @{const_name If}
-        | _ => NONE)
-    in get_first is_builtin end
-
-  fun lookup_const thy name T =
-    let val intern = Sign.intern_const thy name
-    in
-      if Sign.declared_const thy intern
-      then
-        if Sign.typ_instance thy (T, Sign.the_const_type thy intern)
-        then SOME (Const (intern, T))
-        else error ("Boogie: function already declared with different type: " ^
-          quote name)
-      else NONE
-    end
-
-  fun log_term thy t = Syntax.string_of_term_global thy t
-  fun log_new thy name t = name ^ " (as " ^ log_term thy t ^ ")"
-  fun log_ex thy name t = "[" ^ name ^ " has already been declared as " ^
-    log_term thy t ^ "]"
-  fun log_builtin thy name t = "[" ^ name ^ " has been identified as " ^
-    log_term thy t ^ "]"
-
-  fun declare' name isa_name T arity atts thy =
-    (case lookup_const thy isa_name T of
-      SOME t => (((name, t), log_ex thy name t), thy)
-    | NONE =>
-        (case maybe_builtin T atts of
-          SOME t => (((name, t), log_builtin thy name t), thy)
-        | NONE =>
-            thy
-            |> Sign.declare_const_global ((Binding.name isa_name, T),
-                 mk_syntax name arity)
-            |> (fn (t, thy') => (((name, t), log_new thy' name t), thy'))))
-  fun declare (name, ((Ts, T), atts)) =
-    declare' name (isabelle_name name) (Ts ---> T) (length Ts) atts
-
-  fun uniques fns fds =
-    let
-      fun is_unique (name, (([], _), atts)) =
-            (case AList.lookup (op =) atts "unique" of
-              SOME _ => Symtab.lookup fds name
-            | NONE => NONE)
-        | is_unique _ = NONE
-    in
-      map_filter is_unique fns
-      |> map (swap o Term.dest_Const)
-      |> AList.group (op =)
-      |> map (fn (T, ns) => mk_distinct (map (Const o rpair T) ns))
-    end
-in
-fun declare_functions verbose fns =
-  fold_map declare fns #>> split_list #-> (fn (fds, logs) =>
-  log verbose "Loaded constants:" logs #>
-  rpair (` (uniques fns) (Symtab.make fds)))
-end
-
-
-
-local
-  fun name_axioms axs =
-    let fun mk_name idx = "axiom_" ^ string_of_int (idx + 1)
-    in map_index (fn (idx, t) => (mk_name idx, HOLogic.mk_Trueprop t)) axs end
-
-  datatype kind = Unused of thm | Used of thm | New of string
-
-  fun mark (name, t) axs =
-    (case Termtab.lookup axs t of
-      SOME (Unused thm) => Termtab.update (t, Used thm) axs
-    | NONE => Termtab.update (t, New name) axs
-    | SOME _ => axs)
-
-  val sort_fst_str = sort (prod_ord fast_string_ord (K EQUAL)) 
-  fun split_list_kind thy axs =
-    let
-      fun split (_, Used thm) (used, new) = (thm :: used, new)
-        | split (t, New name) (used, new) = (used, (name, t) :: new)
-        | split (_, Unused thm) (used, new) =
-           (warning (Pretty.str_of
-             (Pretty.big_list "This background axiom has not been loaded:"
-               [Display.pretty_thm_global thy thm]));
-            (used, new))
-    in apsnd sort_fst_str (fold split axs ([], [])) end
-
-  fun mark_axioms thy axs =
-    Boogie_Axioms.get (Proof_Context.init_global thy)
-    |> Termtab.make o map (fn thm => (Thm.prop_of thm, Unused thm))
-    |> fold mark axs
-    |> split_list_kind thy o Termtab.dest
-in
-fun add_axioms verbose axs thy =
-  let
-    val (used, new) = mark_axioms thy (name_axioms axs)
-  in
-    thy
-    |> fold_map (fn (n, t) => Specification.axiom ((Binding.name n, []), t)) new
-    |-> Context.theory_map o fold (Boogie_Axioms.add_thm o Drule.export_without_context)
-    |> log verbose "The following axioms were added:" (map fst new)
-    |> (fn thy' => log verbose "The following axioms already existed:"
-         (map (Display.string_of_thm_global thy') used) thy')
-  end
-end
-
-
-
-local
-  fun burrow_distinct eq f xs =
-    let
-      val ys = distinct eq xs
-      val tab = ys ~~ f ys
-    in map (the o AList.lookup eq tab) xs end
-
-  fun indexed names =
-    let
-      val dup = member (op =) (duplicates (op =) (map fst names))
-      fun make_name (n, i) = n ^ (if dup n then "_" ^ string_of_int i else "")
-    in map make_name names end
-
-  fun rename idx_names =
-    idx_names
-    |> burrow_fst (burrow_distinct (op =)
-         (map short_name #> ` (duplicates (op =)) #-> Name.variant_list))
-    |> indexed
-in
-fun add_vcs verbose vcs thy =
-  let val vcs' = burrow_fst rename vcs
-  in
-    thy
-    |> Boogie_VCs.set vcs'
-    |> log verbose "The following verification conditions were loaded:"
-         (map fst vcs')
-  end
-end
-
-
-
-local
-  fun mk_bitT i T =
-    if i = 0
-    then Type (@{type_name "Numeral_Type.bit0"}, [T])
-    else Type (@{type_name "Numeral_Type.bit1"}, [T])
-
-  fun mk_binT size = 
-    if size = 0 then @{typ "Numeral_Type.num0"}
-    else if size = 1 then @{typ "Numeral_Type.num1"}
-    else let val (q, r) = Integer.div_mod size 2 in mk_bitT r (mk_binT q) end
-in
-fun mk_wordT size =
-  if size >= 0 then Type (@{type_name "word"}, [mk_binT size])
-  else raise TYPE ("mk_wordT: " ^ quote (string_of_int size), [], [])
-end
-
-local
-  fun dest_binT T =
-    (case T of
-      Type (@{type_name "Numeral_Type.num0"}, _) => 0
-    | Type (@{type_name "Numeral_Type.num1"}, _) => 1
-    | Type (@{type_name "Numeral_Type.bit0"}, [T]) => 2 * dest_binT T
-    | Type (@{type_name "Numeral_Type.bit1"}, [T]) => 1 + 2 * dest_binT T
-    | _ => raise TYPE ("dest_binT", [T], []))
-in
-val dest_wordT = (fn
-    Type (@{type_name "word"}, [T]) => dest_binT T
-  | T => raise TYPE ("dest_wordT", [T], []))
-end
-
-fun mk_arrayT (Ts, T) = Type (@{type_name "fun"}, [HOLogic.mk_tupleT Ts, T])
-
-
-
-datatype token = Token of string | Newline | EOF
-
-fun tokenize fold_lines input =
-  let
-    fun blank c = (c = #" " orelse c = #"\t" orelse c = #"\n" orelse c = #"\r")
-    fun split line (i, tss) = (i + 1,
-      map (pair i) (map Token (String.tokens blank line) @ [Newline]) :: tss)
-  in apsnd (flat o rev) (fold_lines split input (1, [])) end
-
-fun stopper i = Scan.stopper (K (i, EOF)) (fn (_, EOF) => true | _ => false)
-
-fun scan_err msg [] = (fn () => "Boogie (at end of input): " ^ msg ())
-  | scan_err msg ((i, _) :: _) =
-      (fn () => "Boogie (at line " ^ string_of_int i ^ "): " ^ msg ())
-
-fun scan_fail' msg = Scan.fail_with (scan_err msg)
-fun scan_fail s = scan_fail' (fn () => s)
-
-fun finite scan fold_lines input =
-  let val (i, ts) = tokenize fold_lines input
-  in
-    (case Scan.error (Scan.finite (stopper i) scan) ts of
-      (x, []) => x
-    | (_, ts') => error ((scan_err (fn () => "unparsed input") ts') ()))
-  end
-
-fun read_int' s = (case read_int (raw_explode s) of (i, []) => SOME i | _ => NONE)
-
-fun $$$ s = Scan.one (fn (_, Token s') => s = s' | _ => false)
-fun str st = Scan.some (fn (_, Token s) => SOME s | _ => NONE) st
-fun num st = Scan.some (fn (_, Token s) => read_int' s | _ => NONE) st
-
-fun scan_line key scan =
-  $$$ key |-- scan --| Scan.one (fn (_, Newline) => true | _ => false)
-fun scan_line' key = scan_line key (Scan.succeed ())
-
-fun scan_count scan i =
-  if i > 0 then scan ::: scan_count scan (i - 1) else Scan.succeed []
-
-fun scan_lookup kind tab key =
-  (case Symtab.lookup tab key of
-    SOME value => Scan.succeed value
-  | NONE => scan_fail' (fn () => "undefined " ^ kind ^ ": " ^ quote key))
-
-fun typ tds =
-  let
-    fun tp st =
-     (scan_line' "bool" >> K @{typ bool} ||
-      scan_line' "int" >> K @{typ int} ||
-      scan_line "bv" num >> mk_wordT ||
-      scan_line "type-con" (str -- num) :|-- (fn (name, arity) =>
-        scan_lookup "type constructor" tds name -- scan_count tp arity >>
-        Type) ||
-      scan_line "array" num :|-- (fn arity =>
-        scan_count tp (arity - 1) -- tp >> mk_arrayT) ||
-      scan_fail "illegal type") st
-  in tp end
-
-local
-  fun mk_nary _ t [] = t
-    | mk_nary f _ ts = uncurry (fold_rev f) (split_last ts)
-
-  fun mk_list T = HOLogic.mk_list T
-
-  fun quant name f = scan_line name (num -- num -- num) >> pair f
-  val quants =
-    quant "forall" HOLogic.all_const ||
-    quant "exists" HOLogic.exists_const ||
-    scan_fail "illegal quantifier kind"
-  fun mk_quant q (x, T) t = q T $ absfree (x, T) t
-
-  val patternT = @{typ "SMT.pattern"}
-  fun mk_pattern _ [] = raise TERM ("mk_pattern", [])
-    | mk_pattern n ts =
-        let fun mk_pat t = Const (n, Term.fastype_of t --> patternT) $ t
-        in mk_list patternT (map mk_pat ts) end
-  fun patt n c scan =
-    scan_line n num :|-- scan_count scan >> (mk_pattern c)
-  fun pattern scan =
-    patt "pat" @{const_name "SMT.pat"} scan ||
-    patt "nopat" @{const_name "SMT.nopat"} scan ||
-    scan_fail "illegal pattern kind"
-  fun mk_trigger [] t = t
-    | mk_trigger ps t =
-        @{term "SMT.trigger"} $ mk_list @{typ "SMT.pattern list"} ps $ t
-
-  fun make_label (line, col) = Free (label_name line col, @{typ bool})
-  fun labelled_by kind pos t = kind $ make_label pos $ t
-  fun label offset =
-    $$$ "pos" |-- num -- num >> (fn (line, col) =>
-      if label_name line col = no_label_name then I
-      else labelled_by @{term block_at} (line - offset, col)) ||
-    $$$ "neg" |-- num -- num >> (fn (line, col) =>
-      labelled_by @{term assert_at} (line - offset, col)) ||
-    scan_fail "illegal label kind"
-
-  fun mk_store ((m, k), v) =
-    let
-      val mT = Term.fastype_of m and kT = Term.fastype_of k
-      val vT = Term.fastype_of v
-    in Const (@{const_name fun_upd}, mT --> kT --> vT --> mT) $ m $ k $ v end
-  
-  fun mk_extract ((msb, lsb), t) =
-    let
-      val dT = Term.fastype_of t and rT = mk_wordT (msb - lsb)
-      val nT = @{typ nat}
-      val mk_nat_num = HOLogic.mk_number @{typ nat}
-    in Const (@{const_name slice}, [nT, dT] ---> rT) $ mk_nat_num lsb $ t end
-
-  fun mk_concat (t1, t2) =
-    let
-      val T1 = Term.fastype_of t1 and T2 = Term.fastype_of t2
-      val U = mk_wordT (dest_wordT T1 + dest_wordT T2)
-    in Const (@{const_name word_cat}, [T1, T2] ---> U) $ t1 $ t2 end
-
-  fun unique_labels t =
-    let
-      fun names_of (@{term assert_at} $ Free (n, _) $ t) = cons n #> names_of t
-        | names_of (t $ u) = names_of t #> names_of u
-        | names_of (Abs (_, _, t)) = names_of t
-        | names_of _ = I
-      val nctxt = Name.make_context (duplicates (op =) (names_of t []))
-
-      fun fresh (i, nctxt) = (position_prefix ^ string_of_int i, (i+1, nctxt))
-      fun renamed n (i, nctxt) = Name.variant n nctxt ||> pair i
-      fun mk_label (name, t) = @{term assert_at} $ Free (name, @{typ bool}) $ t
-
-      fun unique t =
-        (case t of
-          @{term assert_at} $ Free (n, _) $ u =>
-            if n = no_label_name
-            then fresh ##>> unique u #>> mk_label
-            else renamed n ##>> unique u #>> mk_label
-        | u1 $ u2 => unique u1 ##>> unique u2 #>> (op $)
-        | Abs (n, T, u) => unique u #>> (fn u' => Abs (n, T, u'))
-        | _ => pair t)
-    in fst (unique t (1, nctxt)) end
-
-  val var_name = str >> prefix var_prefix
-  val dest_var_name = unprefix var_prefix
-  fun rename_variables t =
-    let
-      fun short_var_name n = short_name (dest_var_name n)
-
-      val all_names = Term.add_free_names t []
-      val (names, nctxt) =
-        all_names
-        |> map_filter (try (fn n => (n, short_var_name n)))
-        |> split_list
-        ||>> (fn names => fold_map Name.variant names (Name.make_context all_names))
-        |>> Symtab.make o (op ~~)
-
-      fun rename_free n = the_default n (Symtab.lookup names n)
-      fun rename_abs n = Name.variant (short_var_name n)
-
-      fun rename _ (Free (n, T)) = Free (rename_free n, T)
-        | rename nctxt (Abs (n, T, t)) =
-            let val (n', nctxt') = rename_abs n nctxt
-            in Abs (n', T, rename nctxt' t) end
-        | rename nctxt (t $ u) = rename nctxt t $ rename nctxt u
-        | rename _ t = t
-    in rename nctxt t end
-in
-fun expr offset tds fds =
-  let
-    fun binop t (u1, u2) = t $ u1 $ u2
-    fun binexp s f = scan_line' s |-- exp -- exp >> f
-
-    and exp st =
-     (scan_line' "true" >> K @{term True} ||
-      scan_line' "false" >> K @{term False} ||
-      scan_line' "not" |-- exp >> HOLogic.mk_not ||
-      scan_line "and" num :|-- scan_count exp >> 
-        mk_nary (curry HOLogic.mk_conj) @{term True} ||
-      scan_line "or" num :|-- scan_count exp >>
-        mk_nary (curry HOLogic.mk_disj) @{term False} ||
-      scan_line' "ite" |-- exp -- exp -- exp >> (fn ((c, t1), t2) =>
-        let val T = Term.fastype_of t1
-        in
-          Const (@{const_name If}, [@{typ bool}, T, T] ---> T) $ c $ t1 $ t2
-        end) ||
-      binexp "implies" (binop @{term HOL.implies}) ||
-      scan_line "distinct" num :|-- scan_count exp >> mk_distinct ||
-      binexp "=" HOLogic.mk_eq ||
-      scan_line "var" var_name -- typ tds >> Free ||
-      scan_line "fun" (str -- num) :|-- (fn (name, arity) =>
-        scan_lookup "constant" fds name -- scan_count exp arity >>
-        Term.list_comb) ||
-      quants :|-- (fn (q, ((n, k), i)) =>
-        scan_count (scan_line "var" var_name -- typ tds) n --
-        scan_count (pattern exp) k --
-        scan_count (attribute offset tds fds) i --
-        exp >> (fn (((vs, ps), _), t) =>
-          fold_rev (mk_quant q) vs (mk_trigger ps t))) ||
-      scan_line "label" (label offset) -- exp >> (fn (mk, t) => mk t) ||
-      scan_line "int-num" num >> HOLogic.mk_number @{typ int} ||
-      binexp "<" (binop @{term "op < :: int => _"}) ||
-      binexp "<=" (binop @{term "op <= :: int => _"}) ||
-      binexp ">" (binop @{term "op < :: int => _"} o swap) ||
-      binexp ">=" (binop @{term "op <= :: int => _"} o swap) ||
-      binexp "+" (binop @{term "op + :: int => _"}) ||
-      binexp "-" (binop @{term "op - :: int => _"}) ||
-      binexp "*" (binop @{term "op * :: int => _"}) ||
-      binexp "/" (binop @{term boogie_div}) ||
-      binexp "%" (binop @{term boogie_mod}) ||
-      scan_line "select" num :|-- (fn arity =>
-        exp -- (scan_count exp (arity - 1) >> HOLogic.mk_tuple) >> (op $)) ||
-      scan_line "store" num :|-- (fn arity =>
-        exp -- (scan_count exp (arity - 2) >> HOLogic.mk_tuple) -- exp >> 
-          mk_store) ||
-      scan_line "bv-num" (num -- num) >> (fn (size, i) =>
-        HOLogic.mk_number (mk_wordT size) i) ||
-      scan_line "bv-extract" (num -- num) -- exp >> mk_extract ||
-      binexp "bv-concat" mk_concat ||
-      scan_fail "illegal expression") st
-  in exp >> (rename_variables o unique_labels) end
-
-and attribute offset tds fds =
-  let
-    val attr_val = 
-      scan_line' "expr-attr" |-- expr offset tds fds >> TermValue ||
-      scan_line "string-attr" (Scan.repeat1 str) >>
-        (StringValue o space_implode " ") ||
-      scan_fail "illegal attribute value"
-  in
-    scan_line "attribute" (str -- num) :|-- (fn (name, i) =>
-      scan_count attr_val i >> pair name) ||
-    scan_fail "illegal attribute"
-  end
-end
-
-fun type_decls verbose = Scan.depend (fn thy => 
-  Scan.repeat (scan_line "type-decl" (str -- num -- num) :|-- (fn (ty, i) =>
-    scan_count (attribute 0 Symtab.empty Symtab.empty) i >> K ty)) >>
-    (fn tys => declare_types verbose tys thy))
-
-fun fun_decls verbose tds = Scan.depend (fn thy =>
-  Scan.repeat (scan_line "fun-decl" (str -- num -- num) :|--
-    (fn ((name, arity), i) =>
-      scan_count (typ tds) (arity - 1) -- typ tds --
-      scan_count (attribute 0 tds Symtab.empty) i >> pair name)) >>
-    (fn fns => declare_functions verbose fns thy))
-
-fun axioms verbose tds fds unique_axs = Scan.depend (fn thy =>
-  Scan.repeat (scan_line "axiom" num :|-- (fn i =>
-    expr 0 tds fds --| scan_count (attribute 0 tds fds) i)) >>
-    (fn axs => (add_axioms verbose (unique_axs @ axs) thy, ())))
-
-fun var_decls tds fds = Scan.depend (fn thy =>
-  Scan.repeat (scan_line "var-decl" (str -- num) :|-- (fn (_, i) =>
-    typ tds -- scan_count (attribute 0 tds fds) i >> K ())) >> K (thy, ()))
-
-fun local_vc_offset offsets vc_name =
-   Integer.add ~1 (the_default 1 (AList.lookup (op =) offsets vc_name))
-
-fun vcs verbose offsets tds fds = Scan.depend (fn thy =>
-  Scan.repeat (scan_line "vc" (str -- num) :-- (fn (name, _) =>
-    (expr (local_vc_offset offsets name) tds fds))) >> 
-    (fn vcs => ((), add_vcs verbose vcs thy)))
-
-fun parse verbose offsets thy = Scan.pass thy
- (type_decls verbose :|-- (fn tds =>
-  fun_decls verbose tds :|-- (fn (unique_axs, fds) =>
-  axioms verbose tds fds unique_axs |--
-  var_decls tds fds |--
-  vcs verbose offsets tds fds)))
-
-fun load_b2i verbose offsets path thy = finite (parse verbose offsets thy) File.fold_lines path
-
-fun parse_b2i verbose offsets text thy =
-  finite (parse verbose offsets thy) (fn f => fold f o String.tokens (fn c => c = #"\n")) text
-
-end
--- a/src/HOL/Boogie/Tools/boogie_tactics.ML	Tue Jul 23 13:14:14 2013 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,108 +0,0 @@
-(*  Title:      HOL/Boogie/Tools/boogie_tactics.ML
-    Author:     Sascha Boehme, TU Muenchen
-
-Boogie tactics and Boogie methods.
-*)
-
-signature BOOGIE_TACTICS =
-sig
-  val unfold_labels_tac: Proof.context -> int -> tactic
-  val boogie_tac: Proof.context -> thm list -> int -> tactic
-  val boogie_all_tac: Proof.context -> thm list -> tactic
-  val split: term -> (term list * term) list
-  val split_tac: int -> tactic
-  val drop_assert_at_tac: int -> tactic
-  val setup: theory -> theory
-end
-
-structure Boogie_Tactics: BOOGIE_TACTICS =
-struct
-
-fun as_meta_eq eq = eq RS @{thm eq_reflection}
-
-val assert_at_def = as_meta_eq @{thm assert_at_def}
-val block_at_def = as_meta_eq @{thm block_at_def}
-val label_eqs = [assert_at_def, block_at_def]
-
-fun unfold_labels_tac ctxt =
-  let val unfold = Conv.rewrs_conv label_eqs
-  in CONVERSION (Conv.top_sweep_conv (K unfold) ctxt) end
-
-val boogie_rules =
-  [@{thm fst_conv}, @{thm snd_conv}, @{thm pair_collapse}] @
-  [@{thm fun_upd_same}, @{thm fun_upd_apply}]
-
-fun boogie_tac ctxt rules =
-  unfold_labels_tac ctxt
-  THEN' SMT_Solver.smt_tac ctxt (Boogie_Axioms.get ctxt @ boogie_rules @ rules)
-
-fun boogie_all_tac ctxt rules =
-  PARALLEL_GOALS (ALLGOALS (boogie_tac ctxt rules))
-
-fun boogie_method all =
-  Scan.optional Attrib.thms [] >> (fn thms => fn ctxt => METHOD (fn facts =>
-    let val tac = if all then boogie_all_tac else HEADGOAL oo boogie_tac
-    in tac ctxt (thms @ facts) end))
-
-val setup_boogie = Method.setup @{binding boogie} (boogie_method false)
-  "apply an SMT solver to the current goal \
-  \using the current set of Boogie background axioms"
-
-val setup_boogie_all = Method.setup @{binding boogie_all} (boogie_method true)
-  "apply an SMT solver to all goals \
-  \using the current set of Boogie background axioms"
-
-
-local
-  fun explode_conj (@{term HOL.conj} $ t $ u) = explode_conj t @ explode_conj u
-    | explode_conj t = [t] 
-
-  fun splt (ts, @{term HOL.implies} $ t $ u) = splt (ts @ explode_conj t, u)
-    | splt (ts, @{term HOL.conj} $ t $ u) = splt (ts, t) @ splt (ts, u)
-    | splt (ts, @{term assert_at} $ _ $ t) = [(ts, t)]
-    | splt (_, @{term True}) = []
-    | splt tp = [tp]
-in
-fun split t =
-  splt ([], HOLogic.dest_Trueprop t)
-  |> map (fn (us, u) => (map HOLogic.mk_Trueprop us, HOLogic.mk_Trueprop u))
-end
-
-val split_tac = REPEAT_ALL_NEW (
-  Tactic.resolve_tac [@{thm impI}, @{thm conjI}, @{thm TrueI}]
-  ORELSE' Tactic.etac @{thm conjE})
-
-val drop_assert_at_tac = CONVERSION (Conv.concl_conv ~1 (Conv.try_conv (
-  Conv.arg_conv (Conv.rewr_conv assert_at_def))))
-
-local
-  fun case_name_of t =
-    (case HOLogic.dest_Trueprop (Logic.strip_imp_concl t) of
-      @{term assert_at} $ Free (n, _) $ _ => n
-    | _ => raise TERM ("case_name_of", [t]))
-
-  fun boogie_cases ctxt = METHOD_CASES (fn facts =>
-    ALLGOALS (Method.insert_tac facts THEN' split_tac) #>
-    Seq.maps (fn st =>
-      st
-      |> ALLGOALS drop_assert_at_tac
-      |> Seq.map (pair (map (rpair [] o case_name_of) (Thm.prems_of st)))) #>
-    Seq.maps (fn (names, st) =>
-      CASES
-        (Rule_Cases.make_common
-          (Proof_Context.theory_of ctxt,
-           Thm.prop_of (Rule_Cases.internalize_params st)) (map (rpair []) names))
-        all_tac st))
-in
-val setup_boogie_cases = Method.setup @{binding boogie_cases}
-  (Scan.succeed boogie_cases)
-  "prepare a set of Boogie assertions for case-based proofs"
-end
-
-
-val setup =
-  setup_boogie #>
-  setup_boogie_all #>
-  setup_boogie_cases
-
-end
--- a/src/HOL/Boogie/Tools/boogie_vcs.ML	Tue Jul 23 13:14:14 2013 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,370 +0,0 @@
-(*  Title:      HOL/Boogie/Tools/boogie_vcs.ML
-    Author:     Sascha Boehme, TU Muenchen
-
-Store for Boogie's verification conditions.
-*)
-
-signature BOOGIE_VCS =
-sig
-  type vc
-  val prop_of_vc: vc -> term
-  val size_of: vc -> int
-  val names_of: vc -> string list * string list
-  val path_names_of: vc -> (string * bool) list list
-  val paths_of: vc -> vc list
-  val split_path: int -> vc -> (vc * vc) option
-  val extract: vc -> string -> vc option
-  val only: string list -> vc -> vc
-  val without: string list -> vc -> vc
-  val paths_and: int list -> string list -> vc -> vc
-  val paths_without: int list -> string list -> vc -> vc
-
-  datatype state = Proved | NotProved | PartiallyProved
-  val set: (string * term) list -> theory -> theory
-  val lookup: theory -> string -> vc option
-  val discharge: string * (vc * thm) -> theory -> theory
-  val state_of: theory -> (string * state) list
-  val state_of_vc: theory -> string -> string list * string list
-  val close: theory -> theory
-  val is_closed: theory -> bool
-
-  val rewrite_vcs: (theory -> term -> term) -> (theory -> thm -> thm) ->
-    theory -> theory
-  val add_assertion_filter: (term -> bool) -> theory -> theory
-end
-
-structure Boogie_VCs: BOOGIE_VCS =
-struct
-
-fun app_both f g (x, y) = (f x, g y)
-fun app_hd_tl f g = (fn [] => [] | x :: xs => f x :: map g xs)
-
-
-(* abstract representation of verification conditions *)
-
-datatype vc =
-  Assume of term * vc |
-  Assert of (string * term) * vc |
-  Ignore of vc |
-  Proved of string * vc |
-  Choice of vc * vc |
-  True
-
-val assume = curry Assume and assert = curry Assert
-and proved = curry Proved and choice = curry Choice
-and choice' = curry (Choice o swap)
-
-val vc_of_term =
-  let
-    fun vc_of @{term True} = NONE
-      | vc_of (@{term assert_at} $ Free (n, _) $ t) =
-          SOME (Assert ((n, t), True))
-      | vc_of (@{term HOL.implies} $ @{term True} $ u) = vc_of u
-      | vc_of (@{term HOL.implies} $ t $ u) =
-          vc_of u |> Option.map (assume t)
-      | vc_of (@{term HOL.conj} $ (@{term assert_at} $ Free (n, _) $ t) $ u) =
-          SOME (vc_of u |> the_default True |> assert (n, t))
-      | vc_of (@{term HOL.conj} $ t $ u) =
-          (case (vc_of t, vc_of u) of
-            (NONE, r) => r
-          | (l, NONE) => l
-          | (SOME lv, SOME rv) => SOME (Choice (lv, rv)))
-      | vc_of t = raise TERM ("vc_of_term", [t])
-  in the_default True o vc_of end
-
-val prop_of_vc =
-  let
-    fun mk_conj t u = @{term HOL.conj} $ t $ u
-
-    fun term_of (Assume (t, v)) = @{term HOL.implies} $ t $ term_of v
-      | term_of (Assert ((n, t), v)) =
-          mk_conj (@{term assert_at} $ Free (n, @{typ bool}) $ t) (term_of v)
-      | term_of (Ignore v) = term_of v
-      | term_of (Proved (_, v)) = term_of v
-      | term_of (Choice (lv, rv)) = mk_conj (term_of lv) (term_of rv)
-      | term_of True = @{term True}
-  in HOLogic.mk_Trueprop o term_of end
-
-
-(* properties of verification conditions *)
-
-fun size_of (Assume (_, v)) = size_of v
-  | size_of (Assert (_, v)) = size_of v + 1
-  | size_of (Ignore v) = size_of v
-  | size_of (Proved (_, v)) = size_of v
-  | size_of (Choice (lv, rv)) = size_of lv + size_of rv
-  | size_of True = 0
-
-val names_of =
-  let
-    fun add (Assume (_, v)) = add v
-      | add (Assert ((n, _), v)) = apfst (cons n) #> add v
-      | add (Ignore v) = add v
-      | add (Proved (n, v)) = apsnd (cons n) #> add v
-      | add (Choice (lv, rv)) = add lv #> add rv
-      | add True = I
-  in (fn vc => pairself rev (add vc ([], []))) end
-
-fun path_names_of (Assume (_, v)) = path_names_of v
-  | path_names_of (Assert ((n, _), v)) =
-      path_names_of v
-      |> app_hd_tl (cons (n, true)) (cons (n, false))
-  | path_names_of (Ignore v) = path_names_of v
-  | path_names_of (Proved (n, v)) = map (cons (n, false)) (path_names_of v)
-  | path_names_of (Choice (lv, rv)) = path_names_of lv @ path_names_of rv
-  | path_names_of True = [[]]
-
-fun count_paths (Assume (_, v)) = count_paths v
-  | count_paths (Assert (_, v)) = count_paths v
-  | count_paths (Ignore v) = count_paths v
-  | count_paths (Proved (_, v)) = count_paths v
-  | count_paths (Choice (lv, rv)) = count_paths lv + count_paths rv
-  | count_paths True = 1
-
-
-(* extract parts of a verification condition *)
-
-fun paths_of (Assume (t, v)) = paths_of v |> map (assume t)
-  | paths_of (Assert (a, v)) = paths_of v |> app_hd_tl (assert a) Ignore
-  | paths_of (Ignore v) = paths_of v |> map Ignore
-  | paths_of (Proved (n, v)) = paths_of v |> app_hd_tl (proved n) Ignore
-  | paths_of (Choice (lv, rv)) =
-      map (choice' True) (paths_of lv) @ map (choice True) (paths_of rv)
-  | paths_of True = [True]
-
-fun prune f (Assume (t, v)) = Option.map (assume t) (prune f v)
-  | prune f (Assert (a, v)) = f a v
-  | prune f (Ignore v) = Option.map Ignore (prune f v)
-  | prune f (Proved (n, v)) = Option.map (proved n) (prune f v)
-  | prune f (Choice (lv, rv)) =
-      (case (prune f lv, prune f rv) of
-        (NONE, r) => r |> Option.map (choice True)
-      | (l, NONE) => l |> Option.map (choice' True)
-      | (SOME lv', SOME rv') => SOME (Choice (lv', rv')))
-  | prune _ True = NONE
-
-val split_path =
-  let
-    fun app f = Option.map (pairself f)
-
-    fun split i (Assume (t, v)) = app (assume t) (split i v)
-      | split i (Assert (a, v)) =
-          if i > 1
-          then Option.map (app_both (assert a) Ignore) (split (i-1) v)
-          else Option.map (pair (Assert (a, True)))
-            (prune (SOME o Assert oo pair) (Ignore v))
-      | split i (Ignore v) = app Ignore (split i v)
-      | split i (Proved (n, v)) = app (proved n) (split i v)
-      | split i (Choice (v, True)) = app (choice' True) (split i v)
-      | split i (Choice (True, v)) = app (choice True) (split i v)
-      | split _ _ = NONE
-  in split end
-
-fun select_labels P =
-  let
-    fun assert (a as (n, _)) v =
-      if P n then SOME (Assert (a, the_default True v))
-      else Option.map Ignore v
-    fun sel vc = prune (fn a => assert a o sel) vc
-  in sel end
-
-fun extract vc l = select_labels (equal l) vc
-fun only ls = the_default True o select_labels (member (op =) ls)
-fun without ls = the_default True o select_labels (not o member (op =) ls)
-
-fun select_paths ps sub_select =
-  let
-    fun disjoint pp = null (inter (op =) ps pp)
-
-    fun sel pp (Assume (t, v)) = Assume (t, sel pp v)
-      | sel pp (Assert (a, v)) =
-          if member (op =) ps (hd pp)
-          then Assert (a, sel pp v)
-          else Ignore (sel pp v)
-      | sel pp (Ignore v) = Ignore (sel pp v)
-      | sel pp (Proved (n, v)) = Proved (n, sel pp v)
-      | sel pp (Choice (lv, rv)) =
-          let val (lpp, rpp) = chop (count_paths lv) pp
-          in
-            if disjoint lpp then Choice (sub_select lv, sel rpp rv)
-            else if disjoint rpp then Choice (sel lpp lv, sub_select rv)
-            else Choice (sel lpp lv, sel rpp rv)
-          end
-      | sel _ True = True
-
-    fun sel0 vc =
-      let val pp = 1 upto count_paths vc
-      in if disjoint pp then True else sel pp vc end
-  in sel0 end
-
-fun paths_and ps ls = select_paths ps (only ls)
-fun paths_without ps ls = without ls o select_paths ps (K True)
-
-
-(* discharge parts of a verification condition *)
-
-local
-  fun cprop_of thy t = Thm.cterm_of thy (HOLogic.mk_Trueprop t)
-  fun imp_intr ct thm = Thm.implies_intr ct thm COMP_INCR @{thm impI}
-  fun imp_elim th thm = @{thm mp} OF [thm, th]
-  fun conj1 thm = @{thm conjunct1} OF [thm]
-  fun conj2 thm = @{thm conjunct2} OF [thm]
-  fun conj_intr lth rth = @{thm conjI} OF [lth, rth]
-in
-fun thm_of thy (Assume (t, v)) = imp_intr (cprop_of thy t) (thm_of thy v)
-  | thm_of thy (Assert (_, v)) = thm_of thy v
-  | thm_of thy (Ignore v) = thm_of thy v
-  | thm_of thy (Proved (_, v)) = thm_of thy v
-  | thm_of thy (Choice (lv, rv)) = conj_intr (thm_of thy lv) (thm_of thy rv)
-  | thm_of _ True = @{thm TrueI}
-
-fun join (Assume (_, pv), pthm) (Assume (t, v), thm) =
-      let
-        val mk_prop = Thm.apply @{cterm Trueprop}
-        val ct = Thm.cprop_of thm |> Thm.dest_arg |> Thm.dest_arg1 |> mk_prop
-        val th = Thm.assume ct
-        val (v', thm') = join (pv, imp_elim th pthm) (v, imp_elim th thm)
-      in (Assume (t, v'), imp_intr ct thm') end
-  | join (Assert ((pn, pt), pv), pthm) (Assert ((n, t), v), thm) =
-      let val pthm1 = conj1 pthm
-      in
-        if pn = n andalso pt aconv t
-        then
-          let val (v', thm') = join (pv, conj2 pthm) (v, thm)
-          in (Proved (n, v'), conj_intr pthm1 thm') end
-        else raise THM ("join: not matching", 1, [thm, pthm])
-      end
-  | join (Ignore pv, pthm) (Assert (a, v), thm) =
-      join (pv, pthm) (v, thm) |>> assert a
-  | join (Proved (_, pv), pthm) (Proved (n, v), thm) =
-      let val (v', thm') = join (pv, pthm) (v, conj2 thm)
-      in (Proved (n, v'), conj_intr (conj1 thm) thm') end
-  | join (Ignore pv, pthm) (Proved (n, v), thm) =
-      let val (v', thm') = join (pv, pthm) (v, conj2 thm)
-      in (Proved (n, v'), conj_intr (conj1 thm) thm') end
-  | join (Choice (plv, prv), pthm) (Choice (lv, rv), thm) =
-      let
-        val (lv', lthm) = join (plv, conj1 pthm) (lv, conj1 thm)
-        val (rv', rthm) = join (prv, conj2 pthm) (rv, conj2 thm)
-      in (Choice (lv', rv'), conj_intr lthm rthm) end
-  | join (True, pthm) (v, thm) =
-      if Thm.prop_of pthm aconv @{prop True} then (v, thm)
-      else raise THM ("join: not True", 1, [pthm])
-  | join (_, pthm) (_, thm) = raise THM ("join: not matching", 1, [thm, pthm])
-end
-
-
-fun err_unfinished () = error "An unfinished Boogie environment is still open."
-
-fun err_vcs names = error (Pretty.string_of
-  (Pretty.big_list "Undischarged Boogie verification conditions found:"
-    (map Pretty.str names)))
-
-type vcs_data = {
-  vcs: (vc * (term * thm)) Symtab.table option,
-  rewrite: theory -> thm -> thm,
-  filters: (serial * (term -> bool)) Ord_List.T }
-
-fun make_vcs_data (vcs, rewrite, filters) =
-  {vcs=vcs, rewrite=rewrite, filters=filters}
-
-fun map_vcs_data f ({vcs, rewrite, filters}) =
-  make_vcs_data (f (vcs, rewrite, filters))
-
-fun serial_ord ((i, _), (j, _)) = int_ord (i, j)
-
-structure VCs_Data = Theory_Data
-(
-  type T = vcs_data
-  val empty : T = make_vcs_data (NONE, K I, [])
-  val extend = I
-  fun merge ({vcs=vcs1, filters=fs1, ...} : T, {vcs=vcs2, filters=fs2, ...} : T) =
-    (case (vcs1, vcs2) of
-      (NONE, NONE) =>
-        make_vcs_data (NONE, K I, Ord_List.merge serial_ord (fs1, fs2))
-    | _ => err_unfinished ())
-)
-
-fun add_assertion_filter f =
-  VCs_Data.map (map_vcs_data (fn (vcs, rewrite, filters) =>
-    (vcs, rewrite, Ord_List.insert serial_ord (serial (), f) filters)))
-
-fun filter_assertions thy =
-  let
-    fun filt_assert [] a = assert a
-      | filt_assert ((_, f) :: fs) (a as (_, t)) =
-          if f t then filt_assert fs a else I
-
-    fun filt fs vc =
-      the_default True (prune (fn a => SOME o filt_assert fs a o filt fs) vc)
-
-  in filt (#filters (VCs_Data.get thy)) end
-
-fun prep thy =
-  vc_of_term #>
-  filter_assertions thy #>
-  (fn vc => (vc, (prop_of_vc vc, thm_of thy vc)))
-
-fun set new_vcs thy = VCs_Data.map (map_vcs_data (fn (vcs, rewrite, filters) =>
-  (case vcs of
-    NONE => (SOME (Symtab.make (map (apsnd (prep thy)) new_vcs)), K I, filters)
-  | SOME _ => err_unfinished ()))) thy
-
-fun lookup thy name =
-  (case #vcs (VCs_Data.get thy) of
-    SOME vcs => Option.map fst (Symtab.lookup vcs name)
-  | NONE => NONE)
-
-fun discharge (name, prf) =
-  let fun jn (vc, (t, thm)) = join prf (vc, thm) |> apsnd (pair t)
-  in
-    VCs_Data.map (map_vcs_data (fn (vcs, rewrite, filters) =>
-      (Option.map (Symtab.map_entry name jn) vcs, rewrite, filters)))
-  end
-
-datatype state = Proved | NotProved | PartiallyProved
-
-fun state_of_vc thy name =
-  (case lookup thy name of
-    SOME vc => names_of vc
-  | NONE => ([], []))
-
-fun state_of_vc' (vc, _) =
-  (case names_of vc of
-    ([], _) => Proved
-  | (_, []) => NotProved
-  | (_, _) => PartiallyProved)
-
-fun state_of thy =
-  (case #vcs (VCs_Data.get thy) of
-    SOME vcs => map (apsnd state_of_vc') (Symtab.dest vcs)
-  | NONE => [])
-
-fun finished g (_, (t, thm)) = Thm.prop_of (g thm) aconv t
-
-fun close thy = VCs_Data.map (map_vcs_data (fn (vcs, rewrite, filters) =>
-  (case vcs of
-    SOME raw_vcs =>
-      let
-        fun check vc =
-          state_of_vc' vc = Proved andalso finished (rewrite thy) vc
-
-        val _ =
-          Symtab.dest raw_vcs
-          |> map_filter (fn (n, vc) => if check vc then NONE else SOME n)
-          |> (fn names => if null names then () else err_vcs names)
-      in (NONE, rewrite, filters) end
-  | NONE => (NONE, rewrite, filters)))) thy
-
-val is_closed = is_none o #vcs o VCs_Data.get
-
-fun rewrite_vcs f g thy =
-  let
-    fun rewr (_, (t, _)) = vc_of_term (f thy t)
-      |> (fn vc => (vc, (t, thm_of thy vc)))
-  in
-    VCs_Data.map (map_vcs_data (fn (vcs, _, filters) =>
-      (Option.map (Symtab.map (K rewr)) vcs, g, filters))) thy
-  end
-
-end
--- a/src/HOL/ROOT	Tue Jul 23 13:14:14 2013 +0200
+++ b/src/HOL/ROOT	Tue Jul 23 18:36:23 2013 +0200
@@ -775,30 +775,18 @@
 session "HOL-Word-SMT_Examples" in SMT_Examples = "HOL-Word" +
   options [document = false, quick_and_dirty]
   theories
+    Boogie
     SMT_Examples
     SMT_Word_Examples
   theories [condition = ISABELLE_FULL_TEST]
     SMT_Tests
   files
-    "SMT_Examples.certs"
-    "SMT_Word_Examples.certs"
-
-session "HOL-Boogie" in "Boogie" = "HOL-Word" +
-  options [document = false]
-  theories Boogie
-
-session "HOL-Boogie-Examples" in "Boogie/Examples" = "HOL-Boogie" +
-  options [document = false]
-  theories
-    Boogie_Max_Stepwise
-    Boogie_Max
-    Boogie_Dijkstra
-    VCC_Max
-  files
     "Boogie_Dijkstra.b2i"
     "Boogie_Dijkstra.certs"
     "Boogie_Max.b2i"
     "Boogie_Max.certs"
+    "SMT_Examples.certs"
+    "SMT_Word_Examples.certs"
     "VCC_Max.b2i"
     "VCC_Max.certs"
 
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/SMT_Examples/Boogie.thy	Tue Jul 23 18:36:23 2013 +0200
@@ -0,0 +1,72 @@
+(*  Title:      HOL/SMT_Examples/Boogie.thy
+    Author:     Sascha Boehme, TU Muenchen
+*)
+
+header {* Proving Boogie-generated verification conditions *}
+
+theory Boogie
+imports Main
+begin
+
+text {*
+HOL-Boogie and its applications are described in:
+\begin{itemize}
+
+\item Sascha B"ohme, K. Rustan M. Leino, and Burkhart Wolff.
+HOL-Boogie --- An Interactive Prover for the Boogie Program-Verifier.
+Theorem Proving in Higher Order Logics, 2008.
+
+\item Sascha B"ohme, Micha{\l} Moskal, Wolfram Schulte, and Burkhart Wolff.
+HOL-Boogie --- An Interactive Prover-Backend for the Verifying C Compiler.
+Journal of Automated Reasoning, 2009.
+
+\end{itemize}
+*}
+
+
+
+section {* Built-in types and functions of Boogie *}
+
+subsection {* Integer arithmetics *}
+
+text {*
+The operations @{text div} and @{text mod} are built-in in Boogie, but
+without a particular semantics due to different interpretations in
+programming languages. We assume that each application comes with a
+proper axiomatization.
+*}
+
+axiomatization
+  boogie_div :: "int \<Rightarrow> int \<Rightarrow> int" (infixl "div'_b" 70) and
+  boogie_mod :: "int \<Rightarrow> int \<Rightarrow> int" (infixl "mod'_b" 70)
+
+
+
+section {* Setup *}
+
+ML_file "boogie.ML"
+
+
+
+section {* Verification condition proofs *}
+
+declare [[smt_oracle = false]]
+declare [[smt_read_only_certificates = true]]
+
+
+declare [[smt_certificates = "Boogie_Max.certs"]]
+
+setup {* Boogie.boogie_prove "Boogie_Max.b2i" *}
+
+
+declare [[smt_certificates = "Boogie_Dijkstra.certs"]]
+
+setup {* Boogie.boogie_prove "Boogie_Dijkstra.b2i" *}
+
+
+declare [[z3_with_extensions = true]]
+declare [[smt_certificates = "VCC_Max.certs"]]
+
+setup {* Boogie.boogie_prove "VCC_Max.b2i" *}
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/SMT_Examples/Boogie_Dijkstra.b2i	Tue Jul 23 18:36:23 2013 +0200
@@ -0,0 +1,1879 @@
+type-decl Vertex 0 0
+fun-decl G 1 0
+    array 3
+      type-con Vertex 0
+      type-con Vertex 0
+      int
+fun-decl Infinity 1 0
+    int
+fun-decl Source 1 0
+    type-con Vertex 0
+axiom 0
+    forall 2 0 3
+      var x
+        type-con Vertex 0
+      var y
+        type-con Vertex 0
+      attribute qid 1
+        string-attr BoogieDi.3:15
+      attribute uniqueId 1
+        string-attr 0
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    not
+    =
+    var x
+      type-con Vertex 0
+    var y
+      type-con Vertex 0
+    <
+    int-num 0
+    select 3
+    fun G 0
+    var x
+      type-con Vertex 0
+    var y
+      type-con Vertex 0
+axiom 0
+    forall 2 0 3
+      var x
+        type-con Vertex 0
+      var y
+        type-con Vertex 0
+      attribute qid 1
+        string-attr BoogieDi.4:15
+      attribute uniqueId 1
+        string-attr 1
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    =
+    var x
+      type-con Vertex 0
+    var y
+      type-con Vertex 0
+    =
+    select 3
+    fun G 0
+    var x
+      type-con Vertex 0
+    var y
+      type-con Vertex 0
+    int-num 0
+axiom 0
+    <
+    int-num 0
+    fun Infinity 0
+var-decl SP 0
+    array 2
+      type-con Vertex 0
+      int
+vc Dijkstra 1
+    implies
+    label pos 26 3
+    true
+    implies
+    true
+    implies
+    forall 1 0 3
+      var x
+        type-con Vertex 0
+      attribute qid 1
+        string-attr BoogieDi.27:18
+      attribute uniqueId 1
+        string-attr 5
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    =
+    var x
+      type-con Vertex 0
+    fun Source 0
+    =
+    select 2
+    var SP@0
+      array 2
+        type-con Vertex 0
+        int
+    var x
+      type-con Vertex 0
+    int-num 0
+    implies
+    forall 1 0 3
+      var x
+        type-con Vertex 0
+      attribute qid 1
+        string-attr BoogieDi.28:18
+      attribute uniqueId 1
+        string-attr 6
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    not
+    =
+    var x
+      type-con Vertex 0
+    fun Source 0
+    =
+    select 2
+    var SP@0
+      array 2
+        type-con Vertex 0
+        int
+    var x
+      type-con Vertex 0
+    fun Infinity 0
+    implies
+    forall 1 0 3
+      var x
+        type-con Vertex 0
+      attribute qid 1
+        string-attr BoogieDi.31:18
+      attribute uniqueId 1
+        string-attr 7
+      attribute bvZ3Native 1
+        string-attr False
+    not
+    select 2
+    var Visited@0
+      array 2
+        type-con Vertex 0
+        bool
+    var x
+      type-con Vertex 0
+    implies
+    true
+    and 2
+    label neg 34 5
+    =
+    select 2
+    var SP@0
+      array 2
+        type-con Vertex 0
+        int
+    fun Source 0
+    int-num 0
+    implies
+    =
+    select 2
+    var SP@0
+      array 2
+        type-con Vertex 0
+        int
+    fun Source 0
+    int-num 0
+    and 2
+    label neg 35 5
+    forall 1 0 3
+      var x
+        type-con Vertex 0
+      attribute qid 1
+        string-attr BoogieDi.35:23
+      attribute uniqueId 1
+        string-attr 9
+      attribute bvZ3Native 1
+        string-attr False
+    >=
+    select 2
+    var SP@0
+      array 2
+        type-con Vertex 0
+        int
+    var x
+      type-con Vertex 0
+    int-num 0
+    implies
+    forall 1 0 3
+      var x
+        type-con Vertex 0
+      attribute qid 1
+        string-attr BoogieDi.35:23
+      attribute uniqueId 1
+        string-attr 9
+      attribute bvZ3Native 1
+        string-attr False
+    >=
+    select 2
+    var SP@0
+      array 2
+        type-con Vertex 0
+        int
+    var x
+      type-con Vertex 0
+    int-num 0
+    and 2
+    label neg 36 5
+    forall 2 0 3
+      var y
+        type-con Vertex 0
+      var z
+        type-con Vertex 0
+      attribute qid 1
+        string-attr BoogieDi.36:23
+      attribute uniqueId 1
+        string-attr 10
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 2
+    not
+    select 2
+    var Visited@0
+      array 2
+        type-con Vertex 0
+        bool
+    var z
+      type-con Vertex 0
+    select 2
+    var Visited@0
+      array 2
+        type-con Vertex 0
+        bool
+    var y
+      type-con Vertex 0
+    <=
+    select 2
+    var SP@0
+      array 2
+        type-con Vertex 0
+        int
+    var y
+      type-con Vertex 0
+    select 2
+    var SP@0
+      array 2
+        type-con Vertex 0
+        int
+    var z
+      type-con Vertex 0
+    implies
+    forall 2 0 3
+      var y
+        type-con Vertex 0
+      var z
+        type-con Vertex 0
+      attribute qid 1
+        string-attr BoogieDi.36:23
+      attribute uniqueId 1
+        string-attr 10
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 2
+    not
+    select 2
+    var Visited@0
+      array 2
+        type-con Vertex 0
+        bool
+    var z
+      type-con Vertex 0
+    select 2
+    var Visited@0
+      array 2
+        type-con Vertex 0
+        bool
+    var y
+      type-con Vertex 0
+    <=
+    select 2
+    var SP@0
+      array 2
+        type-con Vertex 0
+        int
+    var y
+      type-con Vertex 0
+    select 2
+    var SP@0
+      array 2
+        type-con Vertex 0
+        int
+    var z
+      type-con Vertex 0
+    and 2
+    label neg 38 5
+    forall 2 0 3
+      var z
+        type-con Vertex 0
+      var y
+        type-con Vertex 0
+      attribute qid 1
+        string-attr BoogieDi.38:23
+      attribute uniqueId 1
+        string-attr 11
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 2
+    select 2
+    var Visited@0
+      array 2
+        type-con Vertex 0
+        bool
+    var y
+      type-con Vertex 0
+    <
+    select 3
+    fun G 0
+    var y
+      type-con Vertex 0
+    var z
+      type-con Vertex 0
+    fun Infinity 0
+    <=
+    select 2
+    var SP@0
+      array 2
+        type-con Vertex 0
+        int
+    var z
+      type-con Vertex 0
+    +
+    select 2
+    var SP@0
+      array 2
+        type-con Vertex 0
+        int
+    var y
+      type-con Vertex 0
+    select 3
+    fun G 0
+    var y
+      type-con Vertex 0
+    var z
+      type-con Vertex 0
+    implies
+    forall 2 0 3
+      var z
+        type-con Vertex 0
+      var y
+        type-con Vertex 0
+      attribute qid 1
+        string-attr BoogieDi.38:23
+      attribute uniqueId 1
+        string-attr 11
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 2
+    select 2
+    var Visited@0
+      array 2
+        type-con Vertex 0
+        bool
+    var y
+      type-con Vertex 0
+    <
+    select 3
+    fun G 0
+    var y
+      type-con Vertex 0
+    var z
+      type-con Vertex 0
+    fun Infinity 0
+    <=
+    select 2
+    var SP@0
+      array 2
+        type-con Vertex 0
+        int
+    var z
+      type-con Vertex 0
+    +
+    select 2
+    var SP@0
+      array 2
+        type-con Vertex 0
+        int
+    var y
+      type-con Vertex 0
+    select 3
+    fun G 0
+    var y
+      type-con Vertex 0
+    var z
+      type-con Vertex 0
+    and 2
+    label neg 40 5
+    forall 1 0 3
+      var z
+        type-con Vertex 0
+      attribute qid 1
+        string-attr BoogieDi.40:23
+      attribute uniqueId 1
+        string-attr 13
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 2
+    not
+    =
+    var z
+      type-con Vertex 0
+    fun Source 0
+    <
+    select 2
+    var SP@0
+      array 2
+        type-con Vertex 0
+        int
+    var z
+      type-con Vertex 0
+    fun Infinity 0
+    exists 1 0 3
+      var y
+        type-con Vertex 0
+      attribute qid 1
+        string-attr BoogieDi.41:15
+      attribute uniqueId 1
+        string-attr 12
+      attribute bvZ3Native 1
+        string-attr False
+    and 3
+    <
+    select 2
+    var SP@0
+      array 2
+        type-con Vertex 0
+        int
+    var y
+      type-con Vertex 0
+    select 2
+    var SP@0
+      array 2
+        type-con Vertex 0
+        int
+    var z
+      type-con Vertex 0
+    select 2
+    var Visited@0
+      array 2
+        type-con Vertex 0
+        bool
+    var y
+      type-con Vertex 0
+    =
+    select 2
+    var SP@0
+      array 2
+        type-con Vertex 0
+        int
+    var z
+      type-con Vertex 0
+    +
+    select 2
+    var SP@0
+      array 2
+        type-con Vertex 0
+        int
+    var y
+      type-con Vertex 0
+    select 3
+    fun G 0
+    var y
+      type-con Vertex 0
+    var z
+      type-con Vertex 0
+    implies
+    forall 1 0 3
+      var z
+        type-con Vertex 0
+      attribute qid 1
+        string-attr BoogieDi.40:23
+      attribute uniqueId 1
+        string-attr 13
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 2
+    not
+    =
+    var z
+      type-con Vertex 0
+    fun Source 0
+    <
+    select 2
+    var SP@0
+      array 2
+        type-con Vertex 0
+        int
+    var z
+      type-con Vertex 0
+    fun Infinity 0
+    exists 1 0 3
+      var y
+        type-con Vertex 0
+      attribute qid 1
+        string-attr BoogieDi.41:15
+      attribute uniqueId 1
+        string-attr 12
+      attribute bvZ3Native 1
+        string-attr False
+    and 3
+    <
+    select 2
+    var SP@0
+      array 2
+        type-con Vertex 0
+        int
+    var y
+      type-con Vertex 0
+    select 2
+    var SP@0
+      array 2
+        type-con Vertex 0
+        int
+    var z
+      type-con Vertex 0
+    select 2
+    var Visited@0
+      array 2
+        type-con Vertex 0
+        bool
+    var y
+      type-con Vertex 0
+    =
+    select 2
+    var SP@0
+      array 2
+        type-con Vertex 0
+        int
+    var z
+      type-con Vertex 0
+    +
+    select 2
+    var SP@0
+      array 2
+        type-con Vertex 0
+        int
+    var y
+      type-con Vertex 0
+    select 3
+    fun G 0
+    var y
+      type-con Vertex 0
+    var z
+      type-con Vertex 0
+    implies
+    label pos 33 3
+    true
+    implies
+    true
+    implies
+    =
+    select 2
+    var SP@1
+      array 2
+        type-con Vertex 0
+        int
+    fun Source 0
+    int-num 0
+    implies
+    forall 1 0 3
+      var x
+        type-con Vertex 0
+      attribute qid 1
+        string-attr BoogieDi.35:23
+      attribute uniqueId 1
+        string-attr 9
+      attribute bvZ3Native 1
+        string-attr False
+    >=
+    select 2
+    var SP@1
+      array 2
+        type-con Vertex 0
+        int
+    var x
+      type-con Vertex 0
+    int-num 0
+    implies
+    forall 2 0 3
+      var y
+        type-con Vertex 0
+      var z
+        type-con Vertex 0
+      attribute qid 1
+        string-attr BoogieDi.36:23
+      attribute uniqueId 1
+        string-attr 10
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 2
+    not
+    select 2
+    var Visited@1
+      array 2
+        type-con Vertex 0
+        bool
+    var z
+      type-con Vertex 0
+    select 2
+    var Visited@1
+      array 2
+        type-con Vertex 0
+        bool
+    var y
+      type-con Vertex 0
+    <=
+    select 2
+    var SP@1
+      array 2
+        type-con Vertex 0
+        int
+    var y
+      type-con Vertex 0
+    select 2
+    var SP@1
+      array 2
+        type-con Vertex 0
+        int
+    var z
+      type-con Vertex 0
+    implies
+    forall 2 0 3
+      var z
+        type-con Vertex 0
+      var y
+        type-con Vertex 0
+      attribute qid 1
+        string-attr BoogieDi.38:23
+      attribute uniqueId 1
+        string-attr 11
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 2
+    select 2
+    var Visited@1
+      array 2
+        type-con Vertex 0
+        bool
+    var y
+      type-con Vertex 0
+    <
+    select 3
+    fun G 0
+    var y
+      type-con Vertex 0
+    var z
+      type-con Vertex 0
+    fun Infinity 0
+    <=
+    select 2
+    var SP@1
+      array 2
+        type-con Vertex 0
+        int
+    var z
+      type-con Vertex 0
+    +
+    select 2
+    var SP@1
+      array 2
+        type-con Vertex 0
+        int
+    var y
+      type-con Vertex 0
+    select 3
+    fun G 0
+    var y
+      type-con Vertex 0
+    var z
+      type-con Vertex 0
+    implies
+    forall 1 0 3
+      var z
+        type-con Vertex 0
+      attribute qid 1
+        string-attr BoogieDi.40:23
+      attribute uniqueId 1
+        string-attr 13
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 2
+    not
+    =
+    var z
+      type-con Vertex 0
+    fun Source 0
+    <
+    select 2
+    var SP@1
+      array 2
+        type-con Vertex 0
+        int
+    var z
+      type-con Vertex 0
+    fun Infinity 0
+    exists 1 0 3
+      var y
+        type-con Vertex 0
+      attribute qid 1
+        string-attr BoogieDi.41:15
+      attribute uniqueId 1
+        string-attr 12
+      attribute bvZ3Native 1
+        string-attr False
+    and 3
+    <
+    select 2
+    var SP@1
+      array 2
+        type-con Vertex 0
+        int
+    var y
+      type-con Vertex 0
+    select 2
+    var SP@1
+      array 2
+        type-con Vertex 0
+        int
+    var z
+      type-con Vertex 0
+    select 2
+    var Visited@1
+      array 2
+        type-con Vertex 0
+        bool
+    var y
+      type-con Vertex 0
+    =
+    select 2
+    var SP@1
+      array 2
+        type-con Vertex 0
+        int
+    var z
+      type-con Vertex 0
+    +
+    select 2
+    var SP@1
+      array 2
+        type-con Vertex 0
+        int
+    var y
+      type-con Vertex 0
+    select 3
+    fun G 0
+    var y
+      type-con Vertex 0
+    var z
+      type-con Vertex 0
+    implies
+    true
+    and 2
+    implies
+    label pos 33 3
+    true
+    implies
+    true
+    implies
+    not
+    exists 1 0 3
+      var x
+        type-con Vertex 0
+      attribute qid 1
+        string-attr BoogieDi.33:18
+      attribute uniqueId 1
+        string-attr 8
+      attribute bvZ3Native 1
+        string-attr False
+    and 2
+    not
+    select 2
+    var Visited@1
+      array 2
+        type-con Vertex 0
+        bool
+    var x
+      type-con Vertex 0
+    <
+    select 2
+    var SP@1
+      array 2
+        type-con Vertex 0
+        int
+    var x
+      type-con Vertex 0
+    fun Infinity 0
+    implies
+    true
+    implies
+    label pos 0 0
+    true
+    implies
+    =
+    var Visited@3
+      array 2
+        type-con Vertex 0
+        bool
+    var Visited@1
+      array 2
+        type-con Vertex 0
+        bool
+    implies
+    =
+    var v@2
+      type-con Vertex 0
+    var v@0
+      type-con Vertex 0
+    implies
+    =
+    var SP@3
+      array 2
+        type-con Vertex 0
+        int
+    var SP@1
+      array 2
+        type-con Vertex 0
+        int
+    implies
+    =
+    var oldSP@1
+      array 2
+        type-con Vertex 0
+        int
+    var oldSP@0
+      array 2
+        type-con Vertex 0
+        int
+    implies
+    label pos 0 0
+    true
+    and 2
+    label neg 17 3
+    forall 1 0 3
+      var z
+        type-con Vertex 0
+      attribute qid 1
+        string-attr BoogieDi.17:19
+      attribute uniqueId 1
+        string-attr 4
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 2
+    not
+    =
+    var z
+      type-con Vertex 0
+    fun Source 0
+    <
+    select 2
+    var SP@3
+      array 2
+        type-con Vertex 0
+        int
+    var z
+      type-con Vertex 0
+    fun Infinity 0
+    exists 1 0 3
+      var y
+        type-con Vertex 0
+      attribute qid 1
+        string-attr BoogieDi.18:13
+      attribute uniqueId 1
+        string-attr 3
+      attribute bvZ3Native 1
+        string-attr False
+    and 2
+    <
+    select 2
+    var SP@3
+      array 2
+        type-con Vertex 0
+        int
+    var y
+      type-con Vertex 0
+    select 2
+    var SP@3
+      array 2
+        type-con Vertex 0
+        int
+    var z
+      type-con Vertex 0
+    =
+    select 2
+    var SP@3
+      array 2
+        type-con Vertex 0
+        int
+    var z
+      type-con Vertex 0
+    +
+    select 2
+    var SP@3
+      array 2
+        type-con Vertex 0
+        int
+    var y
+      type-con Vertex 0
+    select 3
+    fun G 0
+    var y
+      type-con Vertex 0
+    var z
+      type-con Vertex 0
+    implies
+    forall 1 0 3
+      var z
+        type-con Vertex 0
+      attribute qid 1
+        string-attr BoogieDi.17:19
+      attribute uniqueId 1
+        string-attr 4
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 2
+    not
+    =
+    var z
+      type-con Vertex 0
+    fun Source 0
+    <
+    select 2
+    var SP@3
+      array 2
+        type-con Vertex 0
+        int
+    var z
+      type-con Vertex 0
+    fun Infinity 0
+    exists 1 0 3
+      var y
+        type-con Vertex 0
+      attribute qid 1
+        string-attr BoogieDi.18:13
+      attribute uniqueId 1
+        string-attr 3
+      attribute bvZ3Native 1
+        string-attr False
+    and 2
+    <
+    select 2
+    var SP@3
+      array 2
+        type-con Vertex 0
+        int
+    var y
+      type-con Vertex 0
+    select 2
+    var SP@3
+      array 2
+        type-con Vertex 0
+        int
+    var z
+      type-con Vertex 0
+    =
+    select 2
+    var SP@3
+      array 2
+        type-con Vertex 0
+        int
+    var z
+      type-con Vertex 0
+    +
+    select 2
+    var SP@3
+      array 2
+        type-con Vertex 0
+        int
+    var y
+      type-con Vertex 0
+    select 3
+    fun G 0
+    var y
+      type-con Vertex 0
+    var z
+      type-con Vertex 0
+    and 2
+    label neg 15 3
+    forall 2 0 3
+      var z
+        type-con Vertex 0
+      var y
+        type-con Vertex 0
+      attribute qid 1
+        string-attr BoogieDi.15:19
+      attribute uniqueId 1
+        string-attr 2
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 2
+    <
+    select 2
+    var SP@3
+      array 2
+        type-con Vertex 0
+        int
+    var y
+      type-con Vertex 0
+    fun Infinity 0
+    <
+    select 3
+    fun G 0
+    var y
+      type-con Vertex 0
+    var z
+      type-con Vertex 0
+    fun Infinity 0
+    <=
+    select 2
+    var SP@3
+      array 2
+        type-con Vertex 0
+        int
+    var z
+      type-con Vertex 0
+    +
+    select 2
+    var SP@3
+      array 2
+        type-con Vertex 0
+        int
+    var y
+      type-con Vertex 0
+    select 3
+    fun G 0
+    var y
+      type-con Vertex 0
+    var z
+      type-con Vertex 0
+    implies
+    forall 2 0 3
+      var z
+        type-con Vertex 0
+      var y
+        type-con Vertex 0
+      attribute qid 1
+        string-attr BoogieDi.15:19
+      attribute uniqueId 1
+        string-attr 2
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 2
+    <
+    select 2
+    var SP@3
+      array 2
+        type-con Vertex 0
+        int
+    var y
+      type-con Vertex 0
+    fun Infinity 0
+    <
+    select 3
+    fun G 0
+    var y
+      type-con Vertex 0
+    var z
+      type-con Vertex 0
+    fun Infinity 0
+    <=
+    select 2
+    var SP@3
+      array 2
+        type-con Vertex 0
+        int
+    var z
+      type-con Vertex 0
+    +
+    select 2
+    var SP@3
+      array 2
+        type-con Vertex 0
+        int
+    var y
+      type-con Vertex 0
+    select 3
+    fun G 0
+    var y
+      type-con Vertex 0
+    var z
+      type-con Vertex 0
+    and 2
+    label neg 14 3
+    =
+    select 2
+    var SP@3
+      array 2
+        type-con Vertex 0
+        int
+    fun Source 0
+    int-num 0
+    implies
+    =
+    select 2
+    var SP@3
+      array 2
+        type-con Vertex 0
+        int
+    fun Source 0
+    int-num 0
+    true
+    implies
+    label pos 44 5
+    true
+    implies
+    true
+    implies
+    exists 1 0 3
+      var x
+        type-con Vertex 0
+      attribute qid 1
+        string-attr BoogieDi.33:18
+      attribute uniqueId 1
+        string-attr 8
+      attribute bvZ3Native 1
+        string-attr False
+    and 2
+    not
+    select 2
+    var Visited@1
+      array 2
+        type-con Vertex 0
+        bool
+    var x
+      type-con Vertex 0
+    <
+    select 2
+    var SP@1
+      array 2
+        type-con Vertex 0
+        int
+    var x
+      type-con Vertex 0
+    fun Infinity 0
+    implies
+    not
+    select 2
+    var Visited@1
+      array 2
+        type-con Vertex 0
+        bool
+    var v@1
+      type-con Vertex 0
+    implies
+    <
+    select 2
+    var SP@1
+      array 2
+        type-con Vertex 0
+        int
+    var v@1
+      type-con Vertex 0
+    fun Infinity 0
+    implies
+    forall 1 0 3
+      var x
+        type-con Vertex 0
+      attribute qid 1
+        string-attr BoogieDi.47:20
+      attribute uniqueId 1
+        string-attr 14
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    not
+    select 2
+    var Visited@1
+      array 2
+        type-con Vertex 0
+        bool
+    var x
+      type-con Vertex 0
+    <=
+    select 2
+    var SP@1
+      array 2
+        type-con Vertex 0
+        int
+    var v@1
+      type-con Vertex 0
+    select 2
+    var SP@1
+      array 2
+        type-con Vertex 0
+        int
+    var x
+      type-con Vertex 0
+    implies
+    =
+    var Visited@2
+      array 2
+        type-con Vertex 0
+        bool
+    store 3
+    var Visited@1
+      array 2
+        type-con Vertex 0
+        bool
+    var v@1
+      type-con Vertex 0
+    true
+    implies
+    forall 1 0 3
+      var u
+        type-con Vertex 0
+      attribute qid 1
+        string-attr BoogieDi.53:20
+      attribute uniqueId 1
+        string-attr 15
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 2
+    <
+    select 3
+    fun G 0
+    var v@1
+      type-con Vertex 0
+    var u
+      type-con Vertex 0
+    fun Infinity 0
+    <
+    +
+    select 2
+    var SP@1
+      array 2
+        type-con Vertex 0
+        int
+    var v@1
+      type-con Vertex 0
+    select 3
+    fun G 0
+    var v@1
+      type-con Vertex 0
+    var u
+      type-con Vertex 0
+    select 2
+    var SP@1
+      array 2
+        type-con Vertex 0
+        int
+    var u
+      type-con Vertex 0
+    =
+    select 2
+    var SP@2
+      array 2
+        type-con Vertex 0
+        int
+    var u
+      type-con Vertex 0
+    +
+    select 2
+    var SP@1
+      array 2
+        type-con Vertex 0
+        int
+    var v@1
+      type-con Vertex 0
+    select 3
+    fun G 0
+    var v@1
+      type-con Vertex 0
+    var u
+      type-con Vertex 0
+    implies
+    forall 1 0 3
+      var u
+        type-con Vertex 0
+      attribute qid 1
+        string-attr BoogieDi.56:20
+      attribute uniqueId 1
+        string-attr 16
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    not
+    and 2
+    <
+    select 3
+    fun G 0
+    var v@1
+      type-con Vertex 0
+    var u
+      type-con Vertex 0
+    fun Infinity 0
+    <
+    +
+    select 2
+    var SP@1
+      array 2
+        type-con Vertex 0
+        int
+    var v@1
+      type-con Vertex 0
+    select 3
+    fun G 0
+    var v@1
+      type-con Vertex 0
+    var u
+      type-con Vertex 0
+    select 2
+    var SP@1
+      array 2
+        type-con Vertex 0
+        int
+    var u
+      type-con Vertex 0
+    =
+    select 2
+    var SP@2
+      array 2
+        type-con Vertex 0
+        int
+    var u
+      type-con Vertex 0
+    select 2
+    var SP@1
+      array 2
+        type-con Vertex 0
+        int
+    var u
+      type-con Vertex 0
+    and 2
+    label neg 59 5
+    forall 1 0 3
+      var z
+        type-con Vertex 0
+      attribute qid 1
+        string-attr BoogieDi.59:20
+      attribute uniqueId 1
+        string-attr 17
+      attribute bvZ3Native 1
+        string-attr False
+    <=
+    select 2
+    var SP@2
+      array 2
+        type-con Vertex 0
+        int
+    var z
+      type-con Vertex 0
+    select 2
+    var SP@1
+      array 2
+        type-con Vertex 0
+        int
+    var z
+      type-con Vertex 0
+    implies
+    forall 1 0 3
+      var z
+        type-con Vertex 0
+      attribute qid 1
+        string-attr BoogieDi.59:20
+      attribute uniqueId 1
+        string-attr 17
+      attribute bvZ3Native 1
+        string-attr False
+    <=
+    select 2
+    var SP@2
+      array 2
+        type-con Vertex 0
+        int
+    var z
+      type-con Vertex 0
+    select 2
+    var SP@1
+      array 2
+        type-con Vertex 0
+        int
+    var z
+      type-con Vertex 0
+    and 2
+    label neg 60 5
+    forall 1 0 3
+      var y
+        type-con Vertex 0
+      attribute qid 1
+        string-attr BoogieDi.60:20
+      attribute uniqueId 1
+        string-attr 18
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    select 2
+    var Visited@2
+      array 2
+        type-con Vertex 0
+        bool
+    var y
+      type-con Vertex 0
+    =
+    select 2
+    var SP@2
+      array 2
+        type-con Vertex 0
+        int
+    var y
+      type-con Vertex 0
+    select 2
+    var SP@1
+      array 2
+        type-con Vertex 0
+        int
+    var y
+      type-con Vertex 0
+    implies
+    forall 1 0 3
+      var y
+        type-con Vertex 0
+      attribute qid 1
+        string-attr BoogieDi.60:20
+      attribute uniqueId 1
+        string-attr 18
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    select 2
+    var Visited@2
+      array 2
+        type-con Vertex 0
+        bool
+    var y
+      type-con Vertex 0
+    =
+    select 2
+    var SP@2
+      array 2
+        type-con Vertex 0
+        int
+    var y
+      type-con Vertex 0
+    select 2
+    var SP@1
+      array 2
+        type-con Vertex 0
+        int
+    var y
+      type-con Vertex 0
+    implies
+    true
+    implies
+    label pos 0 0
+    true
+    and 2
+    label neg 34 5
+    =
+    select 2
+    var SP@2
+      array 2
+        type-con Vertex 0
+        int
+    fun Source 0
+    int-num 0
+    implies
+    =
+    select 2
+    var SP@2
+      array 2
+        type-con Vertex 0
+        int
+    fun Source 0
+    int-num 0
+    and 2
+    label neg 35 5
+    forall 1 0 3
+      var x
+        type-con Vertex 0
+      attribute qid 1
+        string-attr BoogieDi.35:23
+      attribute uniqueId 1
+        string-attr 9
+      attribute bvZ3Native 1
+        string-attr False
+    >=
+    select 2
+    var SP@2
+      array 2
+        type-con Vertex 0
+        int
+    var x
+      type-con Vertex 0
+    int-num 0
+    implies
+    forall 1 0 3
+      var x
+        type-con Vertex 0
+      attribute qid 1
+        string-attr BoogieDi.35:23
+      attribute uniqueId 1
+        string-attr 9
+      attribute bvZ3Native 1
+        string-attr False
+    >=
+    select 2
+    var SP@2
+      array 2
+        type-con Vertex 0
+        int
+    var x
+      type-con Vertex 0
+    int-num 0
+    and 2
+    label neg 36 5
+    forall 2 0 3
+      var y
+        type-con Vertex 0
+      var z
+        type-con Vertex 0
+      attribute qid 1
+        string-attr BoogieDi.36:23
+      attribute uniqueId 1
+        string-attr 10
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 2
+    not
+    select 2
+    var Visited@2
+      array 2
+        type-con Vertex 0
+        bool
+    var z
+      type-con Vertex 0
+    select 2
+    var Visited@2
+      array 2
+        type-con Vertex 0
+        bool
+    var y
+      type-con Vertex 0
+    <=
+    select 2
+    var SP@2
+      array 2
+        type-con Vertex 0
+        int
+    var y
+      type-con Vertex 0
+    select 2
+    var SP@2
+      array 2
+        type-con Vertex 0
+        int
+    var z
+      type-con Vertex 0
+    implies
+    forall 2 0 3
+      var y
+        type-con Vertex 0
+      var z
+        type-con Vertex 0
+      attribute qid 1
+        string-attr BoogieDi.36:23
+      attribute uniqueId 1
+        string-attr 10
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 2
+    not
+    select 2
+    var Visited@2
+      array 2
+        type-con Vertex 0
+        bool
+    var z
+      type-con Vertex 0
+    select 2
+    var Visited@2
+      array 2
+        type-con Vertex 0
+        bool
+    var y
+      type-con Vertex 0
+    <=
+    select 2
+    var SP@2
+      array 2
+        type-con Vertex 0
+        int
+    var y
+      type-con Vertex 0
+    select 2
+    var SP@2
+      array 2
+        type-con Vertex 0
+        int
+    var z
+      type-con Vertex 0
+    and 2
+    label neg 38 5
+    forall 2 0 3
+      var z
+        type-con Vertex 0
+      var y
+        type-con Vertex 0
+      attribute qid 1
+        string-attr BoogieDi.38:23
+      attribute uniqueId 1
+        string-attr 11
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 2
+    select 2
+    var Visited@2
+      array 2
+        type-con Vertex 0
+        bool
+    var y
+      type-con Vertex 0
+    <
+    select 3
+    fun G 0
+    var y
+      type-con Vertex 0
+    var z
+      type-con Vertex 0
+    fun Infinity 0
+    <=
+    select 2
+    var SP@2
+      array 2
+        type-con Vertex 0
+        int
+    var z
+      type-con Vertex 0
+    +
+    select 2
+    var SP@2
+      array 2
+        type-con Vertex 0
+        int
+    var y
+      type-con Vertex 0
+    select 3
+    fun G 0
+    var y
+      type-con Vertex 0
+    var z
+      type-con Vertex 0
+    implies
+    forall 2 0 3
+      var z
+        type-con Vertex 0
+      var y
+        type-con Vertex 0
+      attribute qid 1
+        string-attr BoogieDi.38:23
+      attribute uniqueId 1
+        string-attr 11
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 2
+    select 2
+    var Visited@2
+      array 2
+        type-con Vertex 0
+        bool
+    var y
+      type-con Vertex 0
+    <
+    select 3
+    fun G 0
+    var y
+      type-con Vertex 0
+    var z
+      type-con Vertex 0
+    fun Infinity 0
+    <=
+    select 2
+    var SP@2
+      array 2
+        type-con Vertex 0
+        int
+    var z
+      type-con Vertex 0
+    +
+    select 2
+    var SP@2
+      array 2
+        type-con Vertex 0
+        int
+    var y
+      type-con Vertex 0
+    select 3
+    fun G 0
+    var y
+      type-con Vertex 0
+    var z
+      type-con Vertex 0
+    and 2
+    label neg 40 5
+    forall 1 0 3
+      var z
+        type-con Vertex 0
+      attribute qid 1
+        string-attr BoogieDi.40:23
+      attribute uniqueId 1
+        string-attr 13
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 2
+    not
+    =
+    var z
+      type-con Vertex 0
+    fun Source 0
+    <
+    select 2
+    var SP@2
+      array 2
+        type-con Vertex 0
+        int
+    var z
+      type-con Vertex 0
+    fun Infinity 0
+    exists 1 0 3
+      var y
+        type-con Vertex 0
+      attribute qid 1
+        string-attr BoogieDi.41:15
+      attribute uniqueId 1
+        string-attr 12
+      attribute bvZ3Native 1
+        string-attr False
+    and 3
+    <
+    select 2
+    var SP@2
+      array 2
+        type-con Vertex 0
+        int
+    var y
+      type-con Vertex 0
+    select 2
+    var SP@2
+      array 2
+        type-con Vertex 0
+        int
+    var z
+      type-con Vertex 0
+    select 2
+    var Visited@2
+      array 2
+        type-con Vertex 0
+        bool
+    var y
+      type-con Vertex 0
+    =
+    select 2
+    var SP@2
+      array 2
+        type-con Vertex 0
+        int
+    var z
+      type-con Vertex 0
+    +
+    select 2
+    var SP@2
+      array 2
+        type-con Vertex 0
+        int
+    var y
+      type-con Vertex 0
+    select 3
+    fun G 0
+    var y
+      type-con Vertex 0
+    var z
+      type-con Vertex 0
+    implies
+    forall 1 0 3
+      var z
+        type-con Vertex 0
+      attribute qid 1
+        string-attr BoogieDi.40:23
+      attribute uniqueId 1
+        string-attr 13
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 2
+    not
+    =
+    var z
+      type-con Vertex 0
+    fun Source 0
+    <
+    select 2
+    var SP@2
+      array 2
+        type-con Vertex 0
+        int
+    var z
+      type-con Vertex 0
+    fun Infinity 0
+    exists 1 0 3
+      var y
+        type-con Vertex 0
+      attribute qid 1
+        string-attr BoogieDi.41:15
+      attribute uniqueId 1
+        string-attr 12
+      attribute bvZ3Native 1
+        string-attr False
+    and 3
+    <
+    select 2
+    var SP@2
+      array 2
+        type-con Vertex 0
+        int
+    var y
+      type-con Vertex 0
+    select 2
+    var SP@2
+      array 2
+        type-con Vertex 0
+        int
+    var z
+      type-con Vertex 0
+    select 2
+    var Visited@2
+      array 2
+        type-con Vertex 0
+        bool
+    var y
+      type-con Vertex 0
+    =
+    select 2
+    var SP@2
+      array 2
+        type-con Vertex 0
+        int
+    var z
+      type-con Vertex 0
+    +
+    select 2
+    var SP@2
+      array 2
+        type-con Vertex 0
+        int
+    var y
+      type-con Vertex 0
+    select 3
+    fun G 0
+    var y
+      type-con Vertex 0
+    var z
+      type-con Vertex 0
+    implies
+    false
+    true
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/SMT_Examples/Boogie_Dijkstra.certs	Tue Jul 23 18:36:23 2013 +0200
@@ -0,0 +1,6137 @@
+5d68fa8702e4a020dc142c33743a5a5445fcba10 6136 0
+#2 := false
+#53 := 0::Int
+decl f30 :: (-> S2 Int)
+decl ?v1!7 :: (-> S2 S2)
+decl ?v0!20 :: S2
+#2123 := ?v0!20
+#5431 := (?v1!7 ?v0!20)
+#18254 := (f30 #5431)
+#1012 := -1::Int
+#17799 := (* -1::Int #18254)
+decl f15 :: (-> S4 Int)
+decl f5 :: (-> S5 S2 S4)
+decl f6 :: (-> S6 S2 S5)
+decl f7 :: S6
+#8 := f7
+#5439 := (f6 f7 #5431)
+#5440 := (f5 #5439 ?v0!20)
+#5441 := (f15 #5440)
+#5442 := (* -1::Int #5441)
+#18528 := (+ #5442 #17799)
+#2126 := (f30 ?v0!20)
+#18529 := (+ #2126 #18528)
+#15418 := (>= #18529 0::Int)
+decl f19 :: (-> S11 S2 Int)
+decl f20 :: S11
+#109 := f20
+#5432 := (f19 f20 #5431)
+#5433 := (* -1::Int #5432)
+#5443 := (+ #5433 #5442)
+#5169 := (f19 f20 ?v0!20)
+#5444 := (+ #5169 #5443)
+#11830 := (>= #5444 0::Int)
+#5445 := (= #5444 0::Int)
+#5446 := (not #5445)
+decl f1 :: S1
+#3 := f1
+decl f9 :: (-> S7 S2 S1)
+decl f21 :: S7
+#115 := f21
+#5436 := (f9 f21 #5431)
+#5437 := (= #5436 f1)
+#5438 := (not #5437)
+#5434 := (+ #5169 #5433)
+#5435 := (<= #5434 0::Int)
+#5447 := (or #5435 #5438 #5446)
+#5448 := (not #5447)
+#5194 := (* -1::Int #5169)
+decl f14 :: Int
+#54 := f14
+#5429 := (+ f14 #5194)
+#5430 := (<= #5429 0::Int)
+#17547 := (not #5430)
+#5195 := (+ #2126 #5194)
+#17503 := (>= #5195 0::Int)
+#5176 := (= #2126 #5169)
+decl f28 :: S2
+#186 := f28
+#19609 := (= f28 ?v0!20)
+#19613 := (not #19609)
+#14451 := (= ?v0!20 f28)
+#15274 := (not #14451)
+#16593 := (iff #15274 #19613)
+#15457 := (iff #14451 #19609)
+#14873 := (iff #19609 #14451)
+#7691 := [commutativity]: #14873
+#16696 := [symm #7691]: #15457
+#14829 := [monotonicity #16696]: #16593
+#5398 := (f9 f21 ?v0!20)
+#5399 := (= #5398 f1)
+#14460 := (or #14451 #5399)
+#15350 := (not #14460)
+decl f10 :: (-> S8 S1 S7)
+decl f11 :: (-> S9 S2 S8)
+decl f12 :: (-> S10 S7 S9)
+decl f13 :: S10
+#27 := f13
+#196 := (f12 f13 f21)
+#197 := (f11 #196 f28)
+#198 := (f10 #197 f1)
+#14446 := (f9 #198 ?v0!20)
+#14450 := (= #14446 f1)
+#14478 := (iff #14450 #14460)
+#11 := (:var 0 S2)
+#42 := (:var 1 S1)
+#40 := (:var 2 S2)
+#38 := (:var 3 S7)
+#39 := (f12 f13 #38)
+#41 := (f11 #39 #40)
+#43 := (f10 #41 #42)
+#44 := (f9 #43 #11)
+#3717 := (pattern #44)
+#48 := (f9 #38 #11)
+#49 := (= #48 f1)
+#47 := (= #42 f1)
+#46 := (= #11 #40)
+#50 := (if #46 #47 #49)
+#45 := (= #44 f1)
+#51 := (iff #45 #50)
+#3718 := (forall (vars (?v0 S7) (?v1 S2) (?v2 S1) (?v3 S2)) (:pat #3717) #51)
+#52 := (forall (vars (?v0 S7) (?v1 S2) (?v2 S1) (?v3 S2)) #51)
+#3721 := (iff #52 #3718)
+#3719 := (iff #51 #51)
+#3720 := [refl]: #3719
+#3722 := [quant-intro #3720]: #3721
+#1579 := (~ #52 #52)
+#1609 := (~ #51 #51)
+#1610 := [refl]: #1609
+#1580 := [nnf-pos #1610]: #1579
+#322 := [asserted]: #52
+#1611 := [mp~ #322 #1580]: #52
+#3723 := [mp #1611 #3722]: #3718
+#7628 := (not #3718)
+#15363 := (or #7628 #14478)
+#4146 := (= f1 f1)
+#14455 := (if #14451 #4146 #5399)
+#14456 := (iff #14450 #14455)
+#15337 := (or #7628 #14456)
+#15318 := (iff #15337 #15363)
+#15276 := (iff #15363 #15363)
+#15289 := [rewrite]: #15276
+#14479 := (iff #14456 #14478)
+#14476 := (iff #14455 #14460)
+#1 := true
+#14457 := (if #14451 true #5399)
+#14461 := (iff #14457 #14460)
+#14475 := [rewrite]: #14461
+#14458 := (iff #14455 #14457)
+#4148 := (iff #4146 true)
+#4149 := [rewrite]: #4148
+#14459 := [monotonicity #4149]: #14458
+#14477 := [trans #14459 #14475]: #14476
+#14480 := [monotonicity #14477]: #14479
+#15256 := [monotonicity #14480]: #15318
+#15235 := [trans #15256 #15289]: #15318
+#15310 := [quant-inst #115 #186 #3 #2123]: #15337
+#15352 := [mp #15310 #15235]: #15363
+#16371 := [unit-resolution #15352 #3723]: #14478
+#15284 := (not #14450)
+decl f29 :: S7
+#195 := f29
+#4622 := (f9 f29 ?v0!20)
+#4623 := (= #4622 f1)
+#4630 := (not #4623)
+#15122 := (iff #4630 #15284)
+#15124 := (iff #4623 #14450)
+#16582 := (iff #14450 #4623)
+#16482 := (= #14446 #4622)
+#9268 := (= #198 f29)
+#199 := (= f29 #198)
+#91 := (f6 f7 #11)
+#3782 := (pattern #91)
+#217 := (f9 f29 #11)
+#3943 := (pattern #217)
+#207 := (f30 #11)
+#3918 := (pattern #207)
+#2136 := (f5 #91 ?v0!20)
+#2137 := (f15 #2136)
+#2127 := (* -1::Int #2126)
+#2472 := (+ #2127 #2137)
+#2473 := (+ #207 #2472)
+#2476 := (= #2473 0::Int)
+#3030 := (not #2476)
+#218 := (= #217 f1)
+#225 := (not #218)
+#2133 := (+ #207 #2127)
+#2134 := (>= #2133 0::Int)
+#3031 := (or #2134 #225 #3030)
+#3977 := (forall (vars (?v1 S2)) (:pat #3918 #3943 #3782) #3031)
+#3982 := (not #3977)
+#2128 := (+ f14 #2127)
+#2129 := (<= #2128 0::Int)
+decl f16 :: S2
+#65 := f16
+#2124 := (= ?v0!20 f16)
+#9 := (:var 1 S2)
+#92 := (f5 #91 #9)
+#3773 := (pattern #92)
+#229 := (f30 #9)
+#1275 := (* -1::Int #229)
+#1276 := (+ #207 #1275)
+#93 := (f15 #92)
+#1296 := (+ #93 #1276)
+#1294 := (>= #1296 0::Int)
+#1027 := (* -1::Int #93)
+#1028 := (+ f14 #1027)
+#1029 := (<= #1028 0::Int)
+#3022 := (or #225 #1029 #1294)
+#3969 := (forall (vars (?v0 S2) (?v1 S2)) (:pat #3773) #3022)
+#3974 := (not #3969)
+#3985 := (or #3974 #2124 #2129 #3982)
+#3988 := (not #3985)
+decl ?v0!19 :: S2
+#2092 := ?v0!19
+#2105 := (f30 ?v0!19)
+#2106 := (* -1::Int #2105)
+decl ?v1!18 :: S2
+#2091 := ?v1!18
+#2104 := (f30 ?v1!18)
+#2107 := (+ #2104 #2106)
+#2095 := (f6 f7 ?v1!18)
+#2096 := (f5 #2095 ?v0!19)
+#2097 := (f15 #2096)
+#2108 := (+ #2097 #2107)
+#2109 := (>= #2108 0::Int)
+#2098 := (* -1::Int #2097)
+#2099 := (+ f14 #2098)
+#2100 := (<= #2099 0::Int)
+#2093 := (f9 f29 ?v1!18)
+#2094 := (= #2093 f1)
+#2985 := (not #2094)
+#3000 := (or #2985 #2100 #2109)
+#3005 := (not #3000)
+#3991 := (or #3005 #3988)
+#3994 := (not #3991)
+#3960 := (pattern #207 #229)
+#1274 := (>= #1276 0::Int)
+#226 := (f9 f29 #9)
+#227 := (= #226 f1)
+#2962 := (not #227)
+#2977 := (or #218 #2962 #1274)
+#3961 := (forall (vars (?v0 S2) (?v1 S2)) (:pat #3960) #2977)
+#3966 := (not #3961)
+#3997 := (or #3966 #3994)
+#4000 := (not #3997)
+decl ?v0!17 :: S2
+#2065 := ?v0!17
+#2074 := (f30 ?v0!17)
+#2075 := (* -1::Int #2074)
+decl ?v1!16 :: S2
+#2064 := ?v1!16
+#2073 := (f30 ?v1!16)
+#2076 := (+ #2073 #2075)
+#2077 := (>= #2076 0::Int)
+#2069 := (f9 f29 ?v0!17)
+#2070 := (= #2069 f1)
+#2939 := (not #2070)
+#2066 := (f9 f29 ?v1!16)
+#2067 := (= #2066 f1)
+#2954 := (or #2067 #2939 #2077)
+#2959 := (not #2954)
+#4003 := (or #2959 #4000)
+#4006 := (not #4003)
+#1265 := (>= #207 0::Int)
+#3952 := (forall (vars (?v0 S2)) (:pat #3918) #1265)
+#3957 := (not #3952)
+#4009 := (or #3957 #4006)
+#4012 := (not #4009)
+decl ?v0!15 :: S2
+#2049 := ?v0!15
+#2050 := (f30 ?v0!15)
+#2051 := (>= #2050 0::Int)
+#2052 := (not #2051)
+#4015 := (or #2052 #4012)
+#4018 := (not #4015)
+#221 := (f30 f16)
+#222 := (= #221 0::Int)
+#713 := (not #222)
+#4021 := (or #713 #4018)
+#4024 := (not #4021)
+#4027 := (or #713 #4024)
+#4030 := (not #4027)
+#112 := (f19 f20 #11)
+#3805 := (pattern #112)
+#212 := (= #207 #112)
+#603 := (or #225 #212)
+#3944 := (forall (vars (?v0 S2)) (:pat #3943 #3918 #3805) #603)
+#3949 := (not #3944)
+#4033 := (or #3949 #4030)
+#4036 := (not #4033)
+decl ?v0!14 :: S2
+#2024 := ?v0!14
+#2029 := (f19 f20 ?v0!14)
+#2028 := (f30 ?v0!14)
+#2030 := (= #2028 #2029)
+#2025 := (f9 f29 ?v0!14)
+#2026 := (= #2025 f1)
+#2027 := (not #2026)
+#2031 := (or #2027 #2030)
+#2032 := (not #2031)
+#4039 := (or #2032 #4036)
+#4042 := (not #4039)
+#1255 := (* -1::Int #207)
+#1256 := (+ #112 #1255)
+#1254 := (>= #1256 0::Int)
+#3935 := (forall (vars (?v0 S2)) (:pat #3805 #3918) #1254)
+#3940 := (not #3935)
+#4045 := (or #3940 #4042)
+#4048 := (not #4045)
+decl ?v0!13 :: S2
+#2006 := ?v0!13
+#2008 := (f30 ?v0!13)
+#2009 := (* -1::Int #2008)
+#2007 := (f19 f20 ?v0!13)
+#2010 := (+ #2007 #2009)
+#2011 := (>= #2010 0::Int)
+#2012 := (not #2011)
+#4051 := (or #2012 #4048)
+#4054 := (not #4051)
+#200 := (f6 f7 f28)
+#201 := (f5 #200 #11)
+#3917 := (pattern #201)
+#202 := (f15 #201)
+#1229 := (* -1::Int #202)
+#190 := (f19 f20 f28)
+#1235 := (* -1::Int #190)
+#1236 := (+ #1235 #1229)
+#1237 := (+ #112 #1236)
+#1238 := (<= #1237 0::Int)
+#1230 := (+ f14 #1229)
+#1231 := (<= #1230 0::Int)
+#2911 := (or #1231 #1238)
+#2912 := (not #2911)
+#2933 := (or #2912 #212)
+#3927 := (forall (vars (?v0 S2)) (:pat #3917 #3805 #3918) #2933)
+#3932 := (not #3927)
+#1385 := (+ #202 #1255)
+#1386 := (+ #190 #1385)
+#1383 := (= #1386 0::Int)
+#2925 := (or #1231 #1238 #1383)
+#3919 := (forall (vars (?v0 S2)) (:pat #3917 #3805 #3918) #2925)
+#3924 := (not #3919)
+#778 := (not #199)
+#116 := (f9 f21 #11)
+#3839 := (pattern #116)
+#1398 := (+ #112 #1235)
+#1397 := (>= #1398 0::Int)
+#117 := (= #116 f1)
+#1401 := (or #117 #1397)
+#3909 := (forall (vars (?v0 S2)) (:pat #3839 #3805) #1401)
+#3914 := (not #3909)
+#1410 := (+ f14 #1235)
+#1411 := (<= #1410 0::Int)
+#187 := (f9 f21 f28)
+#188 := (= #187 f1)
+decl ?v0!12 :: S2
+#1961 := ?v0!12
+#1965 := (f19 f20 ?v0!12)
+#1966 := (* -1::Int #1965)
+#1967 := (+ f14 #1966)
+#1968 := (<= #1967 0::Int)
+#1962 := (f9 f21 ?v0!12)
+#1963 := (= #1962 f1)
+#4057 := (or #1963 #1968 #188 #1411 #3914 #778 #3924 #3932 #4054)
+#4060 := (not #4057)
+decl f25 :: S11
+#148 := f25
+#168 := (f19 f25 f16)
+#169 := (= #168 0::Int)
+#156 := (f19 f25 #9)
+#1149 := (* -1::Int #156)
+#153 := (f19 f25 #11)
+#1150 := (+ #153 #1149)
+#1156 := (+ #93 #1150)
+#1179 := (>= #1156 0::Int)
+#1136 := (* -1::Int #153)
+#1137 := (+ f14 #1136)
+#1138 := (<= #1137 0::Int)
+#2865 := (or #1138 #1029 #1179)
+#3871 := (forall (vars (?v0 S2) (?v1 S2)) (:pat #3773) #2865)
+#3876 := (not #3871)
+#3879 := (or #3876 #169)
+#3882 := (not #3879)
+decl ?v0!11 :: S2
+#1905 := ?v0!11
+#1920 := (f19 f25 ?v0!11)
+#1921 := (* -1::Int #1920)
+decl ?v1!10 :: S2
+#1904 := ?v1!10
+#1911 := (f6 f7 ?v1!10)
+#1912 := (f5 #1911 ?v0!11)
+#1913 := (f15 #1912)
+#2441 := (+ #1913 #1921)
+#1906 := (f19 f25 ?v1!10)
+#2442 := (+ #1906 #2441)
+#2445 := (>= #2442 0::Int)
+#1914 := (* -1::Int #1913)
+#1915 := (+ f14 #1914)
+#1916 := (<= #1915 0::Int)
+#1907 := (* -1::Int #1906)
+#1908 := (+ f14 #1907)
+#1909 := (<= #1908 0::Int)
+#2843 := (or #1909 #1916 #2445)
+#2848 := (not #2843)
+#3885 := (or #2848 #3882)
+#3888 := (not #3885)
+#3848 := (pattern #153)
+decl ?v1!9 :: (-> S2 S2)
+#1880 := (?v1!9 #11)
+#1885 := (f6 f7 #1880)
+#1886 := (f5 #1885 #11)
+#1887 := (f15 #1886)
+#2424 := (* -1::Int #1887)
+#1881 := (f19 f25 #1880)
+#2407 := (* -1::Int #1881)
+#2425 := (+ #2407 #2424)
+#2426 := (+ #153 #2425)
+#2427 := (= #2426 0::Int)
+#2813 := (not #2427)
+#2408 := (+ #153 #2407)
+#2409 := (<= #2408 0::Int)
+#2814 := (or #2409 #2813)
+#2815 := (not #2814)
+#66 := (= #11 f16)
+#2821 := (or #66 #1138 #2815)
+#3863 := (forall (vars (?v0 S2)) (:pat #3848) #2821)
+#3868 := (not #3863)
+#3891 := (or #3868 #3888)
+#3894 := (not #3891)
+decl ?v0!8 :: S2
+#1840 := ?v0!8
+#1853 := (f5 #91 ?v0!8)
+#1854 := (f15 #1853)
+#1843 := (f19 f25 ?v0!8)
+#1844 := (* -1::Int #1843)
+#2377 := (+ #1844 #1854)
+#2378 := (+ #153 #2377)
+#2381 := (= #2378 0::Int)
+#2777 := (not #2381)
+#1850 := (+ #153 #1844)
+#1851 := (>= #1850 0::Int)
+#2778 := (or #1851 #2777)
+#3849 := (forall (vars (?v1 S2)) (:pat #3848 #3782) #2778)
+#3854 := (not #3849)
+#1845 := (+ f14 #1844)
+#1846 := (<= #1845 0::Int)
+#1841 := (= ?v0!8 f16)
+#3857 := (or #1841 #1846 #3854)
+#3860 := (not #3857)
+#3897 := (or #3860 #3894)
+#3900 := (not #3897)
+decl f27 :: S11
+#151 := f27
+decl f26 :: S11
+#150 := f26
+#152 := (= f26 f27)
+#522 := (not #152)
+#149 := (= f25 f20)
+#531 := (not #149)
+decl f24 :: S2
+#146 := f24
+decl f23 :: S2
+#145 := f23
+#147 := (= f23 f24)
+#540 := (not #147)
+decl f22 :: S7
+#143 := f22
+#144 := (= f22 f21)
+#549 := (not #144)
+#1091 := (* -1::Int #112)
+#1092 := (+ f14 #1091)
+#1093 := (<= #1092 0::Int)
+#2763 := (or #117 #1093)
+#3840 := (forall (vars (?v0 S2)) (:pat #3839 #3805) #2763)
+#3845 := (not #3840)
+#3903 := (or #3845 #549 #540 #531 #522 #3900)
+#3906 := (not #3903)
+#4063 := (or #3906 #4060)
+#4066 := (not #4063)
+#1796 := (?v1!7 #11)
+#1803 := (f6 f7 #1796)
+#1804 := (f5 #1803 #11)
+#1805 := (f15 #1804)
+#2350 := (* -1::Int #1805)
+#1797 := (f19 f20 #1796)
+#2333 := (* -1::Int #1797)
+#2351 := (+ #2333 #2350)
+#2352 := (+ #112 #2351)
+#2353 := (= #2352 0::Int)
+#2747 := (not #2353)
+#1801 := (f9 f21 #1796)
+#1802 := (= #1801 f1)
+#2746 := (not #1802)
+#2334 := (+ #112 #2333)
+#2335 := (<= #2334 0::Int)
+#2748 := (or #2335 #2746 #2747)
+#2749 := (not #2748)
+#2755 := (or #66 #1093 #2749)
+#3831 := (forall (vars (?v0 S2)) (:pat #3805) #2755)
+#3836 := (not #3831)
+#122 := (f19 f20 #9)
+#1105 := (* -1::Int #122)
+#1106 := (+ #112 #1105)
+#1107 := (+ #93 #1106)
+#1460 := (>= #1107 0::Int)
+#118 := (not #117)
+#2727 := (or #118 #1029 #1460)
+#3823 := (forall (vars (?v0 S2) (?v1 S2)) (:pat #3773) #2727)
+#3828 := (not #3823)
+#119 := (f9 f21 #9)
+#3814 := (pattern #116 #119)
+#1109 := (>= #1106 0::Int)
+#120 := (= #119 f1)
+#2690 := (not #120)
+#2705 := (or #117 #2690 #1109)
+#3815 := (forall (vars (?v0 S2) (?v1 S2)) (:pat #3814) #2705)
+#3820 := (not #3815)
+#1483 := (>= #112 0::Int)
+#3806 := (forall (vars (?v0 S2)) (:pat #3805) #1483)
+#3811 := (not #3806)
+#110 := (f19 f20 f16)
+#111 := (= #110 0::Int)
+#878 := (not #111)
+decl f17 :: (-> S2 Int)
+#67 := (f17 #11)
+#3736 := (pattern #67)
+decl ?v1!6 :: (-> S2 S2)
+#1743 := (?v1!6 #11)
+#1750 := (f6 f7 #1743)
+#1751 := (f5 #1750 #11)
+#1752 := (f15 #1751)
+#2308 := (* -1::Int #1752)
+#1744 := (f17 #1743)
+#2291 := (* -1::Int #1744)
+#2309 := (+ #2291 #2308)
+#2310 := (+ #67 #2309)
+#2311 := (= #2310 0::Int)
+#2674 := (not #2311)
+decl f18 :: S7
+#75 := f18
+#1748 := (f9 f18 #1743)
+#1749 := (= #1748 f1)
+#2673 := (not #1749)
+#2292 := (+ #67 #2291)
+#2293 := (<= #2292 0::Int)
+#2675 := (or #2293 #2673 #2674)
+#2676 := (not #2675)
+#1053 := (* -1::Int #67)
+#1054 := (+ f14 #1053)
+#1055 := (<= #1054 0::Int)
+#2682 := (or #66 #1055 #2676)
+#3797 := (forall (vars (?v0 S2)) (:pat #3736) #2682)
+#3802 := (not #3797)
+#4069 := (or #3802 #878 #3811 #3820 #3828 #3836 #4066)
+#4072 := (not #4069)
+#76 := (f9 f18 #11)
+#3749 := (pattern #76)
+decl ?v0!5 :: S2
+#1702 := ?v0!5
+#1715 := (f5 #91 ?v0!5)
+#1716 := (f15 #1715)
+#1705 := (f17 ?v0!5)
+#1706 := (* -1::Int #1705)
+#1717 := (+ #1706 #1716)
+#1718 := (+ #67 #1717)
+#1719 := (= #1718 0::Int)
+#2637 := (not #1719)
+#77 := (= #76 f1)
+#78 := (not #77)
+#1712 := (+ #67 #1706)
+#1713 := (>= #1712 0::Int)
+#2638 := (or #1713 #78 #2637)
+#3783 := (forall (vars (?v1 S2)) (:pat #3736 #3749 #3782) #2638)
+#3788 := (not #3783)
+#1707 := (+ f14 #1706)
+#1708 := (<= #1707 0::Int)
+#1703 := (= ?v0!5 f16)
+#3791 := (or #1703 #1708 #3788)
+#6181 := (= f14 #1705)
+#6178 := (= #1705 f14)
+#6207 := (iff #6178 #6181)
+#6203 := (iff #6181 #6178)
+#6186 := [commutativity]: #6203
+#6185 := [symm #6186]: #6207
+#1704 := (not #1703)
+#3794 := (not #3791)
+#6196 := [hypothesis]: #3794
+#3351 := (or #3791 #1704)
+#3352 := [def-axiom]: #3351
+#6201 := [unit-resolution #3352 #6196]: #1704
+#72 := (= #67 f14)
+#364 := (or #66 #72)
+#3743 := (forall (vars (?v0 S2)) (:pat #3736) #364)
+#367 := (forall (vars (?v0 S2)) #364)
+#3746 := (iff #367 #3743)
+#3744 := (iff #364 #364)
+#3745 := [refl]: #3744
+#3747 := [quant-intro #3745]: #3746
+#1589 := (~ #367 #367)
+#1619 := (~ #364 #364)
+#1620 := [refl]: #1619
+#1590 := [nnf-pos #1620]: #1589
+#1318 := (= #1296 0::Int)
+#1321 := (not #1274)
+#1330 := (and #1321 #218 #1318)
+#1335 := (exists (vars (?v1 S2)) #1330)
+#1307 := (+ f14 #1255)
+#1308 := (<= #1307 0::Int)
+#1309 := (not #1308)
+#71 := (not #66)
+#1312 := (and #71 #1309)
+#1315 := (not #1312)
+#1338 := (or #1315 #1335)
+#1341 := (forall (vars (?v0 S2)) #1338)
+#1030 := (not #1029)
+#1288 := (and #218 #1030)
+#1291 := (not #1288)
+#1298 := (or #1291 #1294)
+#1301 := (forall (vars (?v0 S2) (?v1 S2)) #1298)
+#1304 := (not #1301)
+#1344 := (or #1304 #1341)
+#1347 := (and #1301 #1344)
+#228 := (and #225 #227)
+#609 := (not #228)
+#1279 := (or #609 #1274)
+#1282 := (forall (vars (?v0 S2) (?v1 S2)) #1279)
+#1285 := (not #1282)
+#1350 := (or #1285 #1347)
+#1353 := (and #1282 #1350)
+#1268 := (forall (vars (?v0 S2)) #1265)
+#1271 := (not #1268)
+#1356 := (or #1271 #1353)
+#1359 := (and #1268 #1356)
+#1362 := (or #713 #1359)
+#1365 := (and #222 #1362)
+#606 := (forall (vars (?v0 S2)) #603)
+#736 := (not #606)
+#1368 := (or #736 #1365)
+#1371 := (and #606 #1368)
+#1259 := (forall (vars (?v0 S2)) #1254)
+#1262 := (not #1259)
+#1374 := (or #1262 #1371)
+#1377 := (and #1259 #1374)
+#1239 := (not #1238)
+#1232 := (not #1231)
+#1242 := (and #1232 #1239)
+#1245 := (or #1242 #212)
+#1248 := (forall (vars (?v0 S2)) #1245)
+#1251 := (not #1248)
+#1380 := (not #1242)
+#1388 := (or #1380 #1383)
+#1391 := (forall (vars (?v0 S2)) #1388)
+#1394 := (not #1391)
+#1404 := (forall (vars (?v0 S2)) #1401)
+#1407 := (not #1404)
+#1094 := (not #1093)
+#1203 := (and #118 #1094)
+#1206 := (exists (vars (?v0 S2)) #1203)
+#1422 := (not #1206)
+#1446 := (or #1422 #188 #1411 #1407 #778 #1394 #1251 #1377)
+#1139 := (not #1138)
+#1173 := (and #1139 #1030)
+#1176 := (not #1173)
+#1182 := (or #1176 #1179)
+#1185 := (forall (vars (?v0 S2) (?v1 S2)) #1182)
+#1188 := (not #1185)
+#1191 := (or #1188 #169)
+#1194 := (and #1185 #1191)
+#1154 := (= #1156 0::Int)
+#1148 := (>= #1150 0::Int)
+#1151 := (not #1148)
+#1158 := (and #1151 #1154)
+#1161 := (exists (vars (?v1 S2)) #1158)
+#1142 := (and #71 #1139)
+#1145 := (not #1142)
+#1164 := (or #1145 #1161)
+#1167 := (forall (vars (?v0 S2)) #1164)
+#1170 := (not #1167)
+#1197 := (or #1170 #1194)
+#1200 := (and #1167 #1197)
+#1224 := (or #1206 #549 #540 #531 #522 #1200)
+#1451 := (and #1224 #1446)
+#1103 := (= #1107 0::Int)
+#1110 := (not #1109)
+#1119 := (and #1110 #117 #1103)
+#1124 := (exists (vars (?v1 S2)) #1119)
+#1097 := (and #71 #1094)
+#1100 := (not #1097)
+#1127 := (or #1100 #1124)
+#1130 := (forall (vars (?v0 S2)) #1127)
+#1133 := (not #1130)
+#1454 := (and #117 #1030)
+#1457 := (not #1454)
+#1463 := (or #1457 #1460)
+#1466 := (forall (vars (?v0 S2) (?v1 S2)) #1463)
+#1469 := (not #1466)
+#121 := (and #118 #120)
+#391 := (not #121)
+#1474 := (or #391 #1109)
+#1477 := (forall (vars (?v0 S2) (?v1 S2)) #1474)
+#1480 := (not #1477)
+#1486 := (forall (vars (?v0 S2)) #1483)
+#1489 := (not #1486)
+#87 := (f17 #9)
+#1015 := (* -1::Int #87)
+#1042 := (+ #1015 #93)
+#1043 := (+ #67 #1042)
+#1065 := (= #1043 0::Int)
+#1016 := (+ #67 #1015)
+#1014 := (>= #1016 0::Int)
+#1068 := (not #1014)
+#1077 := (and #1068 #77 #1065)
+#1082 := (exists (vars (?v1 S2)) #1077)
+#1056 := (not #1055)
+#1059 := (and #71 #1056)
+#1062 := (not #1059)
+#1085 := (or #1062 #1082)
+#1088 := (forall (vars (?v0 S2)) #1085)
+#1492 := (not #1088)
+#1513 := (or #1492 #878 #1489 #1480 #1469 #1133 #1451)
+#1518 := (and #1088 #1513)
+#1040 := (>= #1043 0::Int)
+#1033 := (and #77 #1030)
+#1036 := (not #1033)
+#1044 := (or #1036 #1040)
+#1047 := (forall (vars (?v0 S2) (?v1 S2)) #1044)
+#1050 := (not #1047)
+#1521 := (or #1050 #1518)
+#1524 := (and #1047 #1521)
+#84 := (f9 f18 #9)
+#85 := (= #84 f1)
+#86 := (and #78 #85)
+#370 := (not #86)
+#1018 := (or #370 #1014)
+#1021 := (forall (vars (?v0 S2) (?v1 S2)) #1018)
+#1024 := (not #1021)
+#1527 := (or #1024 #1524)
+#1530 := (and #1021 #1527)
+#1005 := (>= #67 0::Int)
+#1006 := (forall (vars (?v0 S2)) #1005)
+#1009 := (not #1006)
+#1533 := (or #1009 #1530)
+#1536 := (and #1006 #1533)
+#80 := (f17 f16)
+#81 := (= #80 0::Int)
+#946 := (not #81)
+#1539 := (or #946 #1536)
+#1542 := (and #81 #1539)
+#79 := (forall (vars (?v0 S2)) #78)
+#965 := (not #79)
+#974 := (not #367)
+#68 := (= #67 0::Int)
+#358 := (or #71 #68)
+#361 := (forall (vars (?v0 S2)) #358)
+#983 := (not #361)
+#1554 := (or #983 #974 #965 #1542)
+#1559 := (not #1554)
+#247 := (implies false true)
+#234 := (+ #207 #93)
+#241 := (= #229 #234)
+#242 := (and #218 #241)
+#240 := (< #207 #229)
+#243 := (and #240 #242)
+#244 := (exists (vars (?v1 S2)) #243)
+#238 := (< #207 f14)
+#239 := (and #71 #238)
+#245 := (implies #239 #244)
+#246 := (forall (vars (?v0 S2)) #245)
+#248 := (implies #246 #247)
+#249 := (and #246 #248)
+#235 := (<= #229 #234)
+#94 := (< #93 f14)
+#233 := (and #218 #94)
+#236 := (implies #233 #235)
+#237 := (forall (vars (?v0 S2) (?v1 S2)) #236)
+#250 := (implies #237 #249)
+#251 := (and #237 #250)
+#230 := (<= #229 #207)
+#231 := (implies #228 #230)
+#232 := (forall (vars (?v0 S2) (?v1 S2)) #231)
+#252 := (implies #232 #251)
+#253 := (and #232 #252)
+#223 := (<= 0::Int #207)
+#224 := (forall (vars (?v0 S2)) #223)
+#254 := (implies #224 #253)
+#255 := (and #224 #254)
+#256 := (implies #222 #255)
+#257 := (and #222 #256)
+#258 := (implies true #257)
+#259 := (implies true #258)
+#219 := (implies #218 #212)
+#220 := (forall (vars (?v0 S2)) #219)
+#260 := (implies #220 #259)
+#261 := (and #220 #260)
+#215 := (<= #207 #112)
+#216 := (forall (vars (?v0 S2)) #215)
+#262 := (implies #216 #261)
+#263 := (and #216 #262)
+#204 := (+ #190 #202)
+#205 := (< #204 #112)
+#203 := (< #202 f14)
+#206 := (and #203 #205)
+#211 := (not #206)
+#213 := (implies #211 #212)
+#214 := (forall (vars (?v0 S2)) #213)
+#264 := (implies #214 #263)
+#208 := (= #207 #204)
+#209 := (implies #206 #208)
+#210 := (forall (vars (?v0 S2)) #209)
+#265 := (implies #210 #264)
+#266 := (implies #199 #265)
+#192 := (<= #190 #112)
+#193 := (implies #118 #192)
+#194 := (forall (vars (?v0 S2)) #193)
+#267 := (implies #194 #266)
+#191 := (< #190 f14)
+#268 := (implies #191 #267)
+#189 := (not #188)
+#269 := (implies #189 #268)
+#131 := (< #112 f14)
+#140 := (and #118 #131)
+#141 := (exists (vars (?v0 S2)) #140)
+#270 := (implies #141 #269)
+#271 := (implies true #270)
+#272 := (implies true #271)
+#170 := (implies #169 true)
+#171 := (and #169 #170)
+#158 := (+ #153 #93)
+#165 := (<= #156 #158)
+#154 := (< #153 f14)
+#164 := (and #154 #94)
+#166 := (implies #164 #165)
+#167 := (forall (vars (?v0 S2) (?v1 S2)) #166)
+#172 := (implies #167 #171)
+#173 := (and #167 #172)
+#159 := (= #156 #158)
+#157 := (< #153 #156)
+#160 := (and #157 #159)
+#161 := (exists (vars (?v1 S2)) #160)
+#155 := (and #71 #154)
+#162 := (implies #155 #161)
+#163 := (forall (vars (?v0 S2)) #162)
+#174 := (implies #163 #173)
+#175 := (and #163 #174)
+#176 := (implies true #175)
+#177 := (implies #152 #176)
+#178 := (implies #149 #177)
+#179 := (implies #147 #178)
+#180 := (implies #144 #179)
+#181 := (implies true #180)
+#182 := (implies true #181)
+#142 := (not #141)
+#183 := (implies #142 #182)
+#184 := (implies true #183)
+#185 := (implies true #184)
+#273 := (and #185 #272)
+#274 := (implies true #273)
+#127 := (+ #112 #93)
+#134 := (= #122 #127)
+#135 := (and #117 #134)
+#133 := (< #112 #122)
+#136 := (and #133 #135)
+#137 := (exists (vars (?v1 S2)) #136)
+#132 := (and #71 #131)
+#138 := (implies #132 #137)
+#139 := (forall (vars (?v0 S2)) #138)
+#275 := (implies #139 #274)
+#128 := (<= #122 #127)
+#126 := (and #117 #94)
+#129 := (implies #126 #128)
+#130 := (forall (vars (?v0 S2) (?v1 S2)) #129)
+#276 := (implies #130 #275)
+#123 := (<= #122 #112)
+#124 := (implies #121 #123)
+#125 := (forall (vars (?v0 S2) (?v1 S2)) #124)
+#277 := (implies #125 #276)
+#113 := (<= 0::Int #112)
+#114 := (forall (vars (?v0 S2)) #113)
+#278 := (implies #114 #277)
+#279 := (implies #111 #278)
+#280 := (implies true #279)
+#281 := (implies true #280)
+#96 := (+ #67 #93)
+#103 := (= #87 #96)
+#104 := (and #77 #103)
+#102 := (< #67 #87)
+#105 := (and #102 #104)
+#106 := (exists (vars (?v1 S2)) #105)
+#100 := (< #67 f14)
+#101 := (and #71 #100)
+#107 := (implies #101 #106)
+#108 := (forall (vars (?v0 S2)) #107)
+#282 := (implies #108 #281)
+#283 := (and #108 #282)
+#97 := (<= #87 #96)
+#95 := (and #77 #94)
+#98 := (implies #95 #97)
+#99 := (forall (vars (?v0 S2) (?v1 S2)) #98)
+#284 := (implies #99 #283)
+#285 := (and #99 #284)
+#88 := (<= #87 #67)
+#89 := (implies #86 #88)
+#90 := (forall (vars (?v0 S2) (?v1 S2)) #89)
+#286 := (implies #90 #285)
+#287 := (and #90 #286)
+#82 := (<= 0::Int #67)
+#83 := (forall (vars (?v0 S2)) #82)
+#288 := (implies #83 #287)
+#289 := (and #83 #288)
+#290 := (implies #81 #289)
+#291 := (and #81 #290)
+#292 := (implies true #291)
+#293 := (implies #79 #292)
+#73 := (implies #71 #72)
+#74 := (forall (vars (?v0 S2)) #73)
+#294 := (implies #74 #293)
+#69 := (implies #66 #68)
+#70 := (forall (vars (?v0 S2)) #69)
+#295 := (implies #70 #294)
+#296 := (implies true #295)
+#297 := (implies true #296)
+#298 := (not #297)
+#1562 := (iff #298 #1559)
+#616 := (+ #93 #207)
+#634 := (= #229 #616)
+#637 := (and #218 #634)
+#640 := (and #240 #637)
+#643 := (exists (vars (?v1 S2)) #640)
+#649 := (not #239)
+#650 := (or #649 #643)
+#655 := (forall (vars (?v0 S2)) #650)
+#619 := (<= #229 #616)
+#625 := (not #233)
+#626 := (or #625 #619)
+#631 := (forall (vars (?v0 S2) (?v1 S2)) #626)
+#677 := (not #631)
+#678 := (or #677 #655)
+#683 := (and #631 #678)
+#610 := (or #609 #230)
+#613 := (forall (vars (?v0 S2) (?v1 S2)) #610)
+#689 := (not #613)
+#690 := (or #689 #683)
+#695 := (and #613 #690)
+#701 := (not #224)
+#702 := (or #701 #695)
+#707 := (and #224 #702)
+#714 := (or #713 #707)
+#719 := (and #222 #714)
+#737 := (or #736 #719)
+#742 := (and #606 #737)
+#748 := (not #216)
+#749 := (or #748 #742)
+#754 := (and #216 #749)
+#597 := (or #206 #212)
+#600 := (forall (vars (?v0 S2)) #597)
+#760 := (not #600)
+#761 := (or #760 #754)
+#591 := (or #211 #208)
+#594 := (forall (vars (?v0 S2)) #591)
+#769 := (not #594)
+#770 := (or #769 #761)
+#779 := (or #778 #770)
+#585 := (or #117 #192)
+#588 := (forall (vars (?v0 S2)) #585)
+#787 := (not #588)
+#788 := (or #787 #779)
+#796 := (not #191)
+#797 := (or #796 #788)
+#805 := (or #188 #797)
+#813 := (or #142 #805)
+#440 := (+ #93 #153)
+#464 := (<= #156 #440)
+#470 := (not #164)
+#471 := (or #470 #464)
+#476 := (forall (vars (?v0 S2) (?v1 S2)) #471)
+#491 := (not #476)
+#492 := (or #491 #169)
+#497 := (and #476 #492)
+#443 := (= #156 #440)
+#446 := (and #157 #443)
+#449 := (exists (vars (?v1 S2)) #446)
+#455 := (not #155)
+#456 := (or #455 #449)
+#461 := (forall (vars (?v0 S2)) #456)
+#503 := (not #461)
+#504 := (or #503 #497)
+#509 := (and #461 #504)
+#523 := (or #522 #509)
+#532 := (or #531 #523)
+#541 := (or #540 #532)
+#550 := (or #549 #541)
+#569 := (or #141 #550)
+#829 := (and #569 #813)
+#398 := (+ #93 #112)
+#416 := (= #122 #398)
+#419 := (and #117 #416)
+#422 := (and #133 #419)
+#425 := (exists (vars (?v1 S2)) #422)
+#431 := (not #132)
+#432 := (or #431 #425)
+#437 := (forall (vars (?v0 S2)) #432)
+#842 := (not #437)
+#843 := (or #842 #829)
+#401 := (<= #122 #398)
+#407 := (not #126)
+#408 := (or #407 #401)
+#413 := (forall (vars (?v0 S2) (?v1 S2)) #408)
+#851 := (not #413)
+#852 := (or #851 #843)
+#392 := (or #391 #123)
+#395 := (forall (vars (?v0 S2) (?v1 S2)) #392)
+#860 := (not #395)
+#861 := (or #860 #852)
+#869 := (not #114)
+#870 := (or #869 #861)
+#879 := (or #878 #870)
+#384 := (not #101)
+#385 := (or #384 #106)
+#388 := (forall (vars (?v0 S2)) #385)
+#898 := (not #388)
+#899 := (or #898 #879)
+#904 := (and #388 #899)
+#377 := (not #95)
+#378 := (or #377 #97)
+#381 := (forall (vars (?v0 S2) (?v1 S2)) #378)
+#910 := (not #381)
+#911 := (or #910 #904)
+#916 := (and #381 #911)
+#371 := (or #370 #88)
+#374 := (forall (vars (?v0 S2) (?v1 S2)) #371)
+#922 := (not #374)
+#923 := (or #922 #916)
+#928 := (and #374 #923)
+#934 := (not #83)
+#935 := (or #934 #928)
+#940 := (and #83 #935)
+#947 := (or #946 #940)
+#952 := (and #81 #947)
+#966 := (or #965 #952)
+#975 := (or #974 #966)
+#984 := (or #983 #975)
+#1000 := (not #984)
+#1560 := (iff #1000 #1559)
+#1557 := (iff #984 #1554)
+#1545 := (or #965 #1542)
+#1548 := (or #974 #1545)
+#1551 := (or #983 #1548)
+#1555 := (iff #1551 #1554)
+#1556 := [rewrite]: #1555
+#1552 := (iff #984 #1551)
+#1549 := (iff #975 #1548)
+#1546 := (iff #966 #1545)
+#1543 := (iff #952 #1542)
+#1540 := (iff #947 #1539)
+#1537 := (iff #940 #1536)
+#1534 := (iff #935 #1533)
+#1531 := (iff #928 #1530)
+#1528 := (iff #923 #1527)
+#1525 := (iff #916 #1524)
+#1522 := (iff #911 #1521)
+#1519 := (iff #904 #1518)
+#1516 := (iff #899 #1513)
+#1495 := (or #1133 #1451)
+#1498 := (or #1469 #1495)
+#1501 := (or #1480 #1498)
+#1504 := (or #1489 #1501)
+#1507 := (or #878 #1504)
+#1510 := (or #1492 #1507)
+#1514 := (iff #1510 #1513)
+#1515 := [rewrite]: #1514
+#1511 := (iff #899 #1510)
+#1508 := (iff #879 #1507)
+#1505 := (iff #870 #1504)
+#1502 := (iff #861 #1501)
+#1499 := (iff #852 #1498)
+#1496 := (iff #843 #1495)
+#1452 := (iff #829 #1451)
+#1449 := (iff #813 #1446)
+#1425 := (or #1251 #1377)
+#1428 := (or #1394 #1425)
+#1431 := (or #778 #1428)
+#1434 := (or #1407 #1431)
+#1437 := (or #1411 #1434)
+#1440 := (or #188 #1437)
+#1443 := (or #1422 #1440)
+#1447 := (iff #1443 #1446)
+#1448 := [rewrite]: #1447
+#1444 := (iff #813 #1443)
+#1441 := (iff #805 #1440)
+#1438 := (iff #797 #1437)
+#1435 := (iff #788 #1434)
+#1432 := (iff #779 #1431)
+#1429 := (iff #770 #1428)
+#1426 := (iff #761 #1425)
+#1378 := (iff #754 #1377)
+#1375 := (iff #749 #1374)
+#1372 := (iff #742 #1371)
+#1369 := (iff #737 #1368)
+#1366 := (iff #719 #1365)
+#1363 := (iff #714 #1362)
+#1360 := (iff #707 #1359)
+#1357 := (iff #702 #1356)
+#1354 := (iff #695 #1353)
+#1351 := (iff #690 #1350)
+#1348 := (iff #683 #1347)
+#1345 := (iff #678 #1344)
+#1342 := (iff #655 #1341)
+#1339 := (iff #650 #1338)
+#1336 := (iff #643 #1335)
+#1333 := (iff #640 #1330)
+#1324 := (and #218 #1318)
+#1327 := (and #1321 #1324)
+#1331 := (iff #1327 #1330)
+#1332 := [rewrite]: #1331
+#1328 := (iff #640 #1327)
+#1325 := (iff #637 #1324)
+#1319 := (iff #634 #1318)
+#1320 := [rewrite]: #1319
+#1326 := [monotonicity #1320]: #1325
+#1322 := (iff #240 #1321)
+#1323 := [rewrite]: #1322
+#1329 := [monotonicity #1323 #1326]: #1328
+#1334 := [trans #1329 #1332]: #1333
+#1337 := [quant-intro #1334]: #1336
+#1316 := (iff #649 #1315)
+#1313 := (iff #239 #1312)
+#1310 := (iff #238 #1309)
+#1311 := [rewrite]: #1310
+#1314 := [monotonicity #1311]: #1313
+#1317 := [monotonicity #1314]: #1316
+#1340 := [monotonicity #1317 #1337]: #1339
+#1343 := [quant-intro #1340]: #1342
+#1305 := (iff #677 #1304)
+#1302 := (iff #631 #1301)
+#1299 := (iff #626 #1298)
+#1295 := (iff #619 #1294)
+#1297 := [rewrite]: #1295
+#1292 := (iff #625 #1291)
+#1289 := (iff #233 #1288)
+#1031 := (iff #94 #1030)
+#1032 := [rewrite]: #1031
+#1290 := [monotonicity #1032]: #1289
+#1293 := [monotonicity #1290]: #1292
+#1300 := [monotonicity #1293 #1297]: #1299
+#1303 := [quant-intro #1300]: #1302
+#1306 := [monotonicity #1303]: #1305
+#1346 := [monotonicity #1306 #1343]: #1345
+#1349 := [monotonicity #1303 #1346]: #1348
+#1286 := (iff #689 #1285)
+#1283 := (iff #613 #1282)
+#1280 := (iff #610 #1279)
+#1277 := (iff #230 #1274)
+#1278 := [rewrite]: #1277
+#1281 := [monotonicity #1278]: #1280
+#1284 := [quant-intro #1281]: #1283
+#1287 := [monotonicity #1284]: #1286
+#1352 := [monotonicity #1287 #1349]: #1351
+#1355 := [monotonicity #1284 #1352]: #1354
+#1272 := (iff #701 #1271)
+#1269 := (iff #224 #1268)
+#1266 := (iff #223 #1265)
+#1267 := [rewrite]: #1266
+#1270 := [quant-intro #1267]: #1269
+#1273 := [monotonicity #1270]: #1272
+#1358 := [monotonicity #1273 #1355]: #1357
+#1361 := [monotonicity #1270 #1358]: #1360
+#1364 := [monotonicity #1361]: #1363
+#1367 := [monotonicity #1364]: #1366
+#1370 := [monotonicity #1367]: #1369
+#1373 := [monotonicity #1370]: #1372
+#1263 := (iff #748 #1262)
+#1260 := (iff #216 #1259)
+#1257 := (iff #215 #1254)
+#1258 := [rewrite]: #1257
+#1261 := [quant-intro #1258]: #1260
+#1264 := [monotonicity #1261]: #1263
+#1376 := [monotonicity #1264 #1373]: #1375
+#1379 := [monotonicity #1261 #1376]: #1378
+#1252 := (iff #760 #1251)
+#1249 := (iff #600 #1248)
+#1246 := (iff #597 #1245)
+#1243 := (iff #206 #1242)
+#1240 := (iff #205 #1239)
+#1241 := [rewrite]: #1240
+#1233 := (iff #203 #1232)
+#1234 := [rewrite]: #1233
+#1244 := [monotonicity #1234 #1241]: #1243
+#1247 := [monotonicity #1244]: #1246
+#1250 := [quant-intro #1247]: #1249
+#1253 := [monotonicity #1250]: #1252
+#1427 := [monotonicity #1253 #1379]: #1426
+#1395 := (iff #769 #1394)
+#1392 := (iff #594 #1391)
+#1389 := (iff #591 #1388)
+#1384 := (iff #208 #1383)
+#1387 := [rewrite]: #1384
+#1381 := (iff #211 #1380)
+#1382 := [monotonicity #1244]: #1381
+#1390 := [monotonicity #1382 #1387]: #1389
+#1393 := [quant-intro #1390]: #1392
+#1396 := [monotonicity #1393]: #1395
+#1430 := [monotonicity #1396 #1427]: #1429
+#1433 := [monotonicity #1430]: #1432
+#1408 := (iff #787 #1407)
+#1405 := (iff #588 #1404)
+#1402 := (iff #585 #1401)
+#1399 := (iff #192 #1397)
+#1400 := [rewrite]: #1399
+#1403 := [monotonicity #1400]: #1402
+#1406 := [quant-intro #1403]: #1405
+#1409 := [monotonicity #1406]: #1408
+#1436 := [monotonicity #1409 #1433]: #1435
+#1420 := (iff #796 #1411)
+#1412 := (not #1411)
+#1415 := (not #1412)
+#1418 := (iff #1415 #1411)
+#1419 := [rewrite]: #1418
+#1416 := (iff #796 #1415)
+#1413 := (iff #191 #1412)
+#1414 := [rewrite]: #1413
+#1417 := [monotonicity #1414]: #1416
+#1421 := [trans #1417 #1419]: #1420
+#1439 := [monotonicity #1421 #1436]: #1438
+#1442 := [monotonicity #1439]: #1441
+#1423 := (iff #142 #1422)
+#1207 := (iff #141 #1206)
+#1204 := (iff #140 #1203)
+#1095 := (iff #131 #1094)
+#1096 := [rewrite]: #1095
+#1205 := [monotonicity #1096]: #1204
+#1208 := [quant-intro #1205]: #1207
+#1424 := [monotonicity #1208]: #1423
+#1445 := [monotonicity #1424 #1442]: #1444
+#1450 := [trans #1445 #1448]: #1449
+#1227 := (iff #569 #1224)
+#1209 := (or #522 #1200)
+#1212 := (or #531 #1209)
+#1215 := (or #540 #1212)
+#1218 := (or #549 #1215)
+#1221 := (or #1206 #1218)
+#1225 := (iff #1221 #1224)
+#1226 := [rewrite]: #1225
+#1222 := (iff #569 #1221)
+#1219 := (iff #550 #1218)
+#1216 := (iff #541 #1215)
+#1213 := (iff #532 #1212)
+#1210 := (iff #523 #1209)
+#1201 := (iff #509 #1200)
+#1198 := (iff #504 #1197)
+#1195 := (iff #497 #1194)
+#1192 := (iff #492 #1191)
+#1189 := (iff #491 #1188)
+#1186 := (iff #476 #1185)
+#1183 := (iff #471 #1182)
+#1180 := (iff #464 #1179)
+#1181 := [rewrite]: #1180
+#1177 := (iff #470 #1176)
+#1174 := (iff #164 #1173)
+#1140 := (iff #154 #1139)
+#1141 := [rewrite]: #1140
+#1175 := [monotonicity #1141 #1032]: #1174
+#1178 := [monotonicity #1175]: #1177
+#1184 := [monotonicity #1178 #1181]: #1183
+#1187 := [quant-intro #1184]: #1186
+#1190 := [monotonicity #1187]: #1189
+#1193 := [monotonicity #1190]: #1192
+#1196 := [monotonicity #1187 #1193]: #1195
+#1171 := (iff #503 #1170)
+#1168 := (iff #461 #1167)
+#1165 := (iff #456 #1164)
+#1162 := (iff #449 #1161)
+#1159 := (iff #446 #1158)
+#1155 := (iff #443 #1154)
+#1157 := [rewrite]: #1155
+#1152 := (iff #157 #1151)
+#1153 := [rewrite]: #1152
+#1160 := [monotonicity #1153 #1157]: #1159
+#1163 := [quant-intro #1160]: #1162
+#1146 := (iff #455 #1145)
+#1143 := (iff #155 #1142)
+#1144 := [monotonicity #1141]: #1143
+#1147 := [monotonicity #1144]: #1146
+#1166 := [monotonicity #1147 #1163]: #1165
+#1169 := [quant-intro #1166]: #1168
+#1172 := [monotonicity #1169]: #1171
+#1199 := [monotonicity #1172 #1196]: #1198
+#1202 := [monotonicity #1169 #1199]: #1201
+#1211 := [monotonicity #1202]: #1210
+#1214 := [monotonicity #1211]: #1213
+#1217 := [monotonicity #1214]: #1216
+#1220 := [monotonicity #1217]: #1219
+#1223 := [monotonicity #1208 #1220]: #1222
+#1228 := [trans #1223 #1226]: #1227
+#1453 := [monotonicity #1228 #1450]: #1452
+#1134 := (iff #842 #1133)
+#1131 := (iff #437 #1130)
+#1128 := (iff #432 #1127)
+#1125 := (iff #425 #1124)
+#1122 := (iff #422 #1119)
+#1113 := (and #117 #1103)
+#1116 := (and #1110 #1113)
+#1120 := (iff #1116 #1119)
+#1121 := [rewrite]: #1120
+#1117 := (iff #422 #1116)
+#1114 := (iff #419 #1113)
+#1104 := (iff #416 #1103)
+#1108 := [rewrite]: #1104
+#1115 := [monotonicity #1108]: #1114
+#1111 := (iff #133 #1110)
+#1112 := [rewrite]: #1111
+#1118 := [monotonicity #1112 #1115]: #1117
+#1123 := [trans #1118 #1121]: #1122
+#1126 := [quant-intro #1123]: #1125
+#1101 := (iff #431 #1100)
+#1098 := (iff #132 #1097)
+#1099 := [monotonicity #1096]: #1098
+#1102 := [monotonicity #1099]: #1101
+#1129 := [monotonicity #1102 #1126]: #1128
+#1132 := [quant-intro #1129]: #1131
+#1135 := [monotonicity #1132]: #1134
+#1497 := [monotonicity #1135 #1453]: #1496
+#1470 := (iff #851 #1469)
+#1467 := (iff #413 #1466)
+#1464 := (iff #408 #1463)
+#1461 := (iff #401 #1460)
+#1462 := [rewrite]: #1461
+#1458 := (iff #407 #1457)
+#1455 := (iff #126 #1454)
+#1456 := [monotonicity #1032]: #1455
+#1459 := [monotonicity #1456]: #1458
+#1465 := [monotonicity #1459 #1462]: #1464
+#1468 := [quant-intro #1465]: #1467
+#1471 := [monotonicity #1468]: #1470
+#1500 := [monotonicity #1471 #1497]: #1499
+#1481 := (iff #860 #1480)
+#1478 := (iff #395 #1477)
+#1475 := (iff #392 #1474)
+#1472 := (iff #123 #1109)
+#1473 := [rewrite]: #1472
+#1476 := [monotonicity #1473]: #1475
+#1479 := [quant-intro #1476]: #1478
+#1482 := [monotonicity #1479]: #1481
+#1503 := [monotonicity #1482 #1500]: #1502
+#1490 := (iff #869 #1489)
+#1487 := (iff #114 #1486)
+#1484 := (iff #113 #1483)
+#1485 := [rewrite]: #1484
+#1488 := [quant-intro #1485]: #1487
+#1491 := [monotonicity #1488]: #1490
+#1506 := [monotonicity #1491 #1503]: #1505
+#1509 := [monotonicity #1506]: #1508
+#1493 := (iff #898 #1492)
+#1089 := (iff #388 #1088)
+#1086 := (iff #385 #1085)
+#1083 := (iff #106 #1082)
+#1080 := (iff #105 #1077)
+#1071 := (and #77 #1065)
+#1074 := (and #1068 #1071)
+#1078 := (iff #1074 #1077)
+#1079 := [rewrite]: #1078
+#1075 := (iff #105 #1074)
+#1072 := (iff #104 #1071)
+#1066 := (iff #103 #1065)
+#1067 := [rewrite]: #1066
+#1073 := [monotonicity #1067]: #1072
+#1069 := (iff #102 #1068)
+#1070 := [rewrite]: #1069
+#1076 := [monotonicity #1070 #1073]: #1075
+#1081 := [trans #1076 #1079]: #1080
+#1084 := [quant-intro #1081]: #1083
+#1063 := (iff #384 #1062)
+#1060 := (iff #101 #1059)
+#1057 := (iff #100 #1056)
+#1058 := [rewrite]: #1057
+#1061 := [monotonicity #1058]: #1060
+#1064 := [monotonicity #1061]: #1063
+#1087 := [monotonicity #1064 #1084]: #1086
+#1090 := [quant-intro #1087]: #1089
+#1494 := [monotonicity #1090]: #1493
+#1512 := [monotonicity #1494 #1509]: #1511
+#1517 := [trans #1512 #1515]: #1516
+#1520 := [monotonicity #1090 #1517]: #1519
+#1051 := (iff #910 #1050)
+#1048 := (iff #381 #1047)
+#1045 := (iff #378 #1044)
+#1039 := (iff #97 #1040)
+#1041 := [rewrite]: #1039
+#1037 := (iff #377 #1036)
+#1034 := (iff #95 #1033)
+#1035 := [monotonicity #1032]: #1034
+#1038 := [monotonicity #1035]: #1037
+#1046 := [monotonicity #1038 #1041]: #1045
+#1049 := [quant-intro #1046]: #1048
+#1052 := [monotonicity #1049]: #1051
+#1523 := [monotonicity #1052 #1520]: #1522
+#1526 := [monotonicity #1049 #1523]: #1525
+#1025 := (iff #922 #1024)
+#1022 := (iff #374 #1021)
+#1019 := (iff #371 #1018)
+#1013 := (iff #88 #1014)
+#1017 := [rewrite]: #1013
+#1020 := [monotonicity #1017]: #1019
+#1023 := [quant-intro #1020]: #1022
+#1026 := [monotonicity #1023]: #1025
+#1529 := [monotonicity #1026 #1526]: #1528
+#1532 := [monotonicity #1023 #1529]: #1531
+#1010 := (iff #934 #1009)
+#1007 := (iff #83 #1006)
+#1003 := (iff #82 #1005)
+#1004 := [rewrite]: #1003
+#1008 := [quant-intro #1004]: #1007
+#1011 := [monotonicity #1008]: #1010
+#1535 := [monotonicity #1011 #1532]: #1534
+#1538 := [monotonicity #1008 #1535]: #1537
+#1541 := [monotonicity #1538]: #1540
+#1544 := [monotonicity #1541]: #1543
+#1547 := [monotonicity #1544]: #1546
+#1550 := [monotonicity #1547]: #1549
+#1553 := [monotonicity #1550]: #1552
+#1558 := [trans #1553 #1556]: #1557
+#1561 := [monotonicity #1558]: #1560
+#1001 := (iff #298 #1000)
+#998 := (iff #297 #984)
+#989 := (implies true #984)
+#992 := (iff #989 #984)
+#993 := [rewrite]: #992
+#996 := (iff #297 #989)
+#994 := (iff #296 #984)
+#990 := (iff #296 #989)
+#987 := (iff #295 #984)
+#980 := (implies #361 #975)
+#985 := (iff #980 #984)
+#986 := [rewrite]: #985
+#981 := (iff #295 #980)
+#978 := (iff #294 #975)
+#971 := (implies #367 #966)
+#976 := (iff #971 #975)
+#977 := [rewrite]: #976
+#972 := (iff #294 #971)
+#969 := (iff #293 #966)
+#962 := (implies #79 #952)
+#967 := (iff #962 #966)
+#968 := [rewrite]: #967
+#963 := (iff #293 #962)
+#960 := (iff #292 #952)
+#955 := (implies true #952)
+#958 := (iff #955 #952)
+#959 := [rewrite]: #958
+#956 := (iff #292 #955)
+#953 := (iff #291 #952)
+#950 := (iff #290 #947)
+#943 := (implies #81 #940)
+#948 := (iff #943 #947)
+#949 := [rewrite]: #948
+#944 := (iff #290 #943)
+#941 := (iff #289 #940)
+#938 := (iff #288 #935)
+#931 := (implies #83 #928)
+#936 := (iff #931 #935)
+#937 := [rewrite]: #936
+#932 := (iff #288 #931)
+#929 := (iff #287 #928)
+#926 := (iff #286 #923)
+#919 := (implies #374 #916)
+#924 := (iff #919 #923)
+#925 := [rewrite]: #924
+#920 := (iff #286 #919)
+#917 := (iff #285 #916)
+#914 := (iff #284 #911)
+#907 := (implies #381 #904)
+#912 := (iff #907 #911)
+#913 := [rewrite]: #912
+#908 := (iff #284 #907)
+#905 := (iff #283 #904)
+#902 := (iff #282 #899)
+#895 := (implies #388 #879)
+#900 := (iff #895 #899)
+#901 := [rewrite]: #900
+#896 := (iff #282 #895)
+#893 := (iff #281 #879)
+#884 := (implies true #879)
+#887 := (iff #884 #879)
+#888 := [rewrite]: #887
+#891 := (iff #281 #884)
+#889 := (iff #280 #879)
+#885 := (iff #280 #884)
+#882 := (iff #279 #879)
+#875 := (implies #111 #870)
+#880 := (iff #875 #879)
+#881 := [rewrite]: #880
+#876 := (iff #279 #875)
+#873 := (iff #278 #870)
+#866 := (implies #114 #861)
+#871 := (iff #866 #870)
+#872 := [rewrite]: #871
+#867 := (iff #278 #866)
+#864 := (iff #277 #861)
+#857 := (implies #395 #852)
+#862 := (iff #857 #861)
+#863 := [rewrite]: #862
+#858 := (iff #277 #857)
+#855 := (iff #276 #852)
+#848 := (implies #413 #843)
+#853 := (iff #848 #852)
+#854 := [rewrite]: #853
+#849 := (iff #276 #848)
+#846 := (iff #275 #843)
+#839 := (implies #437 #829)
+#844 := (iff #839 #843)
+#845 := [rewrite]: #844
+#840 := (iff #275 #839)
+#837 := (iff #274 #829)
+#832 := (implies true #829)
+#835 := (iff #832 #829)
+#836 := [rewrite]: #835
+#833 := (iff #274 #832)
+#830 := (iff #273 #829)
+#827 := (iff #272 #813)
+#818 := (implies true #813)
+#821 := (iff #818 #813)
+#822 := [rewrite]: #821
+#825 := (iff #272 #818)
+#823 := (iff #271 #813)
+#819 := (iff #271 #818)
+#816 := (iff #270 #813)
+#810 := (implies #141 #805)
+#814 := (iff #810 #813)
+#815 := [rewrite]: #814
+#811 := (iff #270 #810)
+#808 := (iff #269 #805)
+#802 := (implies #189 #797)
+#806 := (iff #802 #805)
+#807 := [rewrite]: #806
+#803 := (iff #269 #802)
+#800 := (iff #268 #797)
+#793 := (implies #191 #788)
+#798 := (iff #793 #797)
+#799 := [rewrite]: #798
+#794 := (iff #268 #793)
+#791 := (iff #267 #788)
+#784 := (implies #588 #779)
+#789 := (iff #784 #788)
+#790 := [rewrite]: #789
+#785 := (iff #267 #784)
+#782 := (iff #266 #779)
+#775 := (implies #199 #770)
+#780 := (iff #775 #779)
+#781 := [rewrite]: #780
+#776 := (iff #266 #775)
+#773 := (iff #265 #770)
+#766 := (implies #594 #761)
+#771 := (iff #766 #770)
+#772 := [rewrite]: #771
+#767 := (iff #265 #766)
+#764 := (iff #264 #761)
+#757 := (implies #600 #754)
+#762 := (iff #757 #761)
+#763 := [rewrite]: #762
+#758 := (iff #264 #757)
+#755 := (iff #263 #754)
+#752 := (iff #262 #749)
+#745 := (implies #216 #742)
+#750 := (iff #745 #749)
+#751 := [rewrite]: #750
+#746 := (iff #262 #745)
+#743 := (iff #261 #742)
+#740 := (iff #260 #737)
+#733 := (implies #606 #719)
+#738 := (iff #733 #737)
+#739 := [rewrite]: #738
+#734 := (iff #260 #733)
+#731 := (iff #259 #719)
+#722 := (implies true #719)
+#725 := (iff #722 #719)
+#726 := [rewrite]: #725
+#729 := (iff #259 #722)
+#727 := (iff #258 #719)
+#723 := (iff #258 #722)
+#720 := (iff #257 #719)
+#717 := (iff #256 #714)
+#710 := (implies #222 #707)
+#715 := (iff #710 #714)
+#716 := [rewrite]: #715
+#711 := (iff #256 #710)
+#708 := (iff #255 #707)
+#705 := (iff #254 #702)
+#698 := (implies #224 #695)
+#703 := (iff #698 #702)
+#704 := [rewrite]: #703
+#699 := (iff #254 #698)
+#696 := (iff #253 #695)
+#693 := (iff #252 #690)
+#686 := (implies #613 #683)
+#691 := (iff #686 #690)
+#692 := [rewrite]: #691
+#687 := (iff #252 #686)
+#684 := (iff #251 #683)
+#681 := (iff #250 #678)
+#674 := (implies #631 #655)
+#679 := (iff #674 #678)
+#680 := [rewrite]: #679
+#675 := (iff #250 #674)
+#672 := (iff #249 #655)
+#667 := (and #655 true)
+#670 := (iff #667 #655)
+#671 := [rewrite]: #670
+#668 := (iff #249 #667)
+#665 := (iff #248 true)
+#660 := (implies #655 true)
+#663 := (iff #660 true)
+#664 := [rewrite]: #663
+#661 := (iff #248 #660)
+#658 := (iff #247 true)
+#659 := [rewrite]: #658
+#656 := (iff #246 #655)
+#653 := (iff #245 #650)
+#646 := (implies #239 #643)
+#651 := (iff #646 #650)
+#652 := [rewrite]: #651
+#647 := (iff #245 #646)
+#644 := (iff #244 #643)
+#641 := (iff #243 #640)
+#638 := (iff #242 #637)
+#635 := (iff #241 #634)
+#617 := (= #234 #616)
+#618 := [rewrite]: #617
+#636 := [monotonicity #618]: #635
+#639 := [monotonicity #636]: #638
+#642 := [monotonicity #639]: #641
+#645 := [quant-intro #642]: #644
+#648 := [monotonicity #645]: #647
+#654 := [trans #648 #652]: #653
+#657 := [quant-intro #654]: #656
+#662 := [monotonicity #657 #659]: #661
+#666 := [trans #662 #664]: #665
+#669 := [monotonicity #657 #666]: #668
+#673 := [trans #669 #671]: #672
+#632 := (iff #237 #631)
+#629 := (iff #236 #626)
+#622 := (implies #233 #619)
+#627 := (iff #622 #626)
+#628 := [rewrite]: #627
+#623 := (iff #236 #622)
+#620 := (iff #235 #619)
+#621 := [monotonicity #618]: #620
+#624 := [monotonicity #621]: #623
+#630 := [trans #624 #628]: #629
+#633 := [quant-intro #630]: #632
+#676 := [monotonicity #633 #673]: #675
+#682 := [trans #676 #680]: #681
+#685 := [monotonicity #633 #682]: #684
+#614 := (iff #232 #613)
+#611 := (iff #231 #610)
+#612 := [rewrite]: #611
+#615 := [quant-intro #612]: #614
+#688 := [monotonicity #615 #685]: #687
+#694 := [trans #688 #692]: #693
+#697 := [monotonicity #615 #694]: #696
+#700 := [monotonicity #697]: #699
+#706 := [trans #700 #704]: #705
+#709 := [monotonicity #706]: #708
+#712 := [monotonicity #709]: #711
+#718 := [trans #712 #716]: #717
+#721 := [monotonicity #718]: #720
+#724 := [monotonicity #721]: #723
+#728 := [trans #724 #726]: #727
+#730 := [monotonicity #728]: #729
+#732 := [trans #730 #726]: #731
+#607 := (iff #220 #606)
+#604 := (iff #219 #603)
+#605 := [rewrite]: #604
+#608 := [quant-intro #605]: #607
+#735 := [monotonicity #608 #732]: #734
+#741 := [trans #735 #739]: #740
+#744 := [monotonicity #608 #741]: #743
+#747 := [monotonicity #744]: #746
+#753 := [trans #747 #751]: #752
+#756 := [monotonicity #753]: #755
+#601 := (iff #214 #600)
+#598 := (iff #213 #597)
+#599 := [rewrite]: #598
+#602 := [quant-intro #599]: #601
+#759 := [monotonicity #602 #756]: #758
+#765 := [trans #759 #763]: #764
+#595 := (iff #210 #594)
+#592 := (iff #209 #591)
+#593 := [rewrite]: #592
+#596 := [quant-intro #593]: #595
+#768 := [monotonicity #596 #765]: #767
+#774 := [trans #768 #772]: #773
+#777 := [monotonicity #774]: #776
+#783 := [trans #777 #781]: #782
+#589 := (iff #194 #588)
+#586 := (iff #193 #585)
+#587 := [rewrite]: #586
+#590 := [quant-intro #587]: #589
+#786 := [monotonicity #590 #783]: #785
+#792 := [trans #786 #790]: #791
+#795 := [monotonicity #792]: #794
+#801 := [trans #795 #799]: #800
+#804 := [monotonicity #801]: #803
+#809 := [trans #804 #807]: #808
+#812 := [monotonicity #809]: #811
+#817 := [trans #812 #815]: #816
+#820 := [monotonicity #817]: #819
+#824 := [trans #820 #822]: #823
+#826 := [monotonicity #824]: #825
+#828 := [trans #826 #822]: #827
+#583 := (iff #185 #569)
+#574 := (implies true #569)
+#577 := (iff #574 #569)
+#578 := [rewrite]: #577
+#581 := (iff #185 #574)
+#579 := (iff #184 #569)
+#575 := (iff #184 #574)
+#572 := (iff #183 #569)
+#566 := (implies #142 #550)
+#570 := (iff #566 #569)
+#571 := [rewrite]: #570
+#567 := (iff #183 #566)
+#564 := (iff #182 #550)
+#555 := (implies true #550)
+#558 := (iff #555 #550)
+#559 := [rewrite]: #558
+#562 := (iff #182 #555)
+#560 := (iff #181 #550)
+#556 := (iff #181 #555)
+#553 := (iff #180 #550)
+#546 := (implies #144 #541)
+#551 := (iff #546 #550)
+#552 := [rewrite]: #551
+#547 := (iff #180 #546)
+#544 := (iff #179 #541)
+#537 := (implies #147 #532)
+#542 := (iff #537 #541)
+#543 := [rewrite]: #542
+#538 := (iff #179 #537)
+#535 := (iff #178 #532)
+#528 := (implies #149 #523)
+#533 := (iff #528 #532)
+#534 := [rewrite]: #533
+#529 := (iff #178 #528)
+#526 := (iff #177 #523)
+#519 := (implies #152 #509)
+#524 := (iff #519 #523)
+#525 := [rewrite]: #524
+#520 := (iff #177 #519)
+#517 := (iff #176 #509)
+#512 := (implies true #509)
+#515 := (iff #512 #509)
+#516 := [rewrite]: #515
+#513 := (iff #176 #512)
+#510 := (iff #175 #509)
+#507 := (iff #174 #504)
+#500 := (implies #461 #497)
+#505 := (iff #500 #504)
+#506 := [rewrite]: #505
+#501 := (iff #174 #500)
+#498 := (iff #173 #497)
+#495 := (iff #172 #492)
+#488 := (implies #476 #169)
+#493 := (iff #488 #492)
+#494 := [rewrite]: #493
+#489 := (iff #172 #488)
+#486 := (iff #171 #169)
+#481 := (and #169 true)
+#484 := (iff #481 #169)
+#485 := [rewrite]: #484
+#482 := (iff #171 #481)
+#479 := (iff #170 true)
+#480 := [rewrite]: #479
+#483 := [monotonicity #480]: #482
+#487 := [trans #483 #485]: #486
+#477 := (iff #167 #476)
+#474 := (iff #166 #471)
+#467 := (implies #164 #464)
+#472 := (iff #467 #471)
+#473 := [rewrite]: #472
+#468 := (iff #166 #467)
+#465 := (iff #165 #464)
+#441 := (= #158 #440)
+#442 := [rewrite]: #441
+#466 := [monotonicity #442]: #465
+#469 := [monotonicity #466]: #468
+#475 := [trans #469 #473]: #474
+#478 := [quant-intro #475]: #477
+#490 := [monotonicity #478 #487]: #489
+#496 := [trans #490 #494]: #495
+#499 := [monotonicity #478 #496]: #498
+#462 := (iff #163 #461)
+#459 := (iff #162 #456)
+#452 := (implies #155 #449)
+#457 := (iff #452 #456)
+#458 := [rewrite]: #457
+#453 := (iff #162 #452)
+#450 := (iff #161 #449)
+#447 := (iff #160 #446)
+#444 := (iff #159 #443)
+#445 := [monotonicity #442]: #444
+#448 := [monotonicity #445]: #447
+#451 := [quant-intro #448]: #450
+#454 := [monotonicity #451]: #453
+#460 := [trans #454 #458]: #459
+#463 := [quant-intro #460]: #462
+#502 := [monotonicity #463 #499]: #501
+#508 := [trans #502 #506]: #507
+#511 := [monotonicity #463 #508]: #510
+#514 := [monotonicity #511]: #513
+#518 := [trans #514 #516]: #517
+#521 := [monotonicity #518]: #520
+#527 := [trans #521 #525]: #526
+#530 := [monotonicity #527]: #529
+#536 := [trans #530 #534]: #535
+#539 := [monotonicity #536]: #538
+#545 := [trans #539 #543]: #544
+#548 := [monotonicity #545]: #547
+#554 := [trans #548 #552]: #553
+#557 := [monotonicity #554]: #556
+#561 := [trans #557 #559]: #560
+#563 := [monotonicity #561]: #562
+#565 := [trans #563 #559]: #564
+#568 := [monotonicity #565]: #567
+#573 := [trans #568 #571]: #572
+#576 := [monotonicity #573]: #575
+#580 := [trans #576 #578]: #579
+#582 := [monotonicity #580]: #581
+#584 := [trans #582 #578]: #583
+#831 := [monotonicity #584 #828]: #830
+#834 := [monotonicity #831]: #833
+#838 := [trans #834 #836]: #837
+#438 := (iff #139 #437)
+#435 := (iff #138 #432)
+#428 := (implies #132 #425)
+#433 := (iff #428 #432)
+#434 := [rewrite]: #433
+#429 := (iff #138 #428)
+#426 := (iff #137 #425)
+#423 := (iff #136 #422)
+#420 := (iff #135 #419)
+#417 := (iff #134 #416)
+#399 := (= #127 #398)
+#400 := [rewrite]: #399
+#418 := [monotonicity #400]: #417
+#421 := [monotonicity #418]: #420
+#424 := [monotonicity #421]: #423
+#427 := [quant-intro #424]: #426
+#430 := [monotonicity #427]: #429
+#436 := [trans #430 #434]: #435
+#439 := [quant-intro #436]: #438
+#841 := [monotonicity #439 #838]: #840
+#847 := [trans #841 #845]: #846
+#414 := (iff #130 #413)
+#411 := (iff #129 #408)
+#404 := (implies #126 #401)
+#409 := (iff #404 #408)
+#410 := [rewrite]: #409
+#405 := (iff #129 #404)
+#402 := (iff #128 #401)
+#403 := [monotonicity #400]: #402
+#406 := [monotonicity #403]: #405
+#412 := [trans #406 #410]: #411
+#415 := [quant-intro #412]: #414
+#850 := [monotonicity #415 #847]: #849
+#856 := [trans #850 #854]: #855
+#396 := (iff #125 #395)
+#393 := (iff #124 #392)
+#394 := [rewrite]: #393
+#397 := [quant-intro #394]: #396
+#859 := [monotonicity #397 #856]: #858
+#865 := [trans #859 #863]: #864
+#868 := [monotonicity #865]: #867
+#874 := [trans #868 #872]: #873
+#877 := [monotonicity #874]: #876
+#883 := [trans #877 #881]: #882
+#886 := [monotonicity #883]: #885
+#890 := [trans #886 #888]: #889
+#892 := [monotonicity #890]: #891
+#894 := [trans #892 #888]: #893
+#389 := (iff #108 #388)
+#386 := (iff #107 #385)
+#387 := [rewrite]: #386
+#390 := [quant-intro #387]: #389
+#897 := [monotonicity #390 #894]: #896
+#903 := [trans #897 #901]: #902
+#906 := [monotonicity #390 #903]: #905
+#382 := (iff #99 #381)
+#379 := (iff #98 #378)
+#380 := [rewrite]: #379
+#383 := [quant-intro #380]: #382
+#909 := [monotonicity #383 #906]: #908
+#915 := [trans #909 #913]: #914
+#918 := [monotonicity #383 #915]: #917
+#375 := (iff #90 #374)
+#372 := (iff #89 #371)
+#373 := [rewrite]: #372
+#376 := [quant-intro #373]: #375
+#921 := [monotonicity #376 #918]: #920
+#927 := [trans #921 #925]: #926
+#930 := [monotonicity #376 #927]: #929
+#933 := [monotonicity #930]: #932
+#939 := [trans #933 #937]: #938
+#942 := [monotonicity #939]: #941
+#945 := [monotonicity #942]: #944
+#951 := [trans #945 #949]: #950
+#954 := [monotonicity #951]: #953
+#957 := [monotonicity #954]: #956
+#961 := [trans #957 #959]: #960
+#964 := [monotonicity #961]: #963
+#970 := [trans #964 #968]: #969
+#368 := (iff #74 #367)
+#365 := (iff #73 #364)
+#366 := [rewrite]: #365
+#369 := [quant-intro #366]: #368
+#973 := [monotonicity #369 #970]: #972
+#979 := [trans #973 #977]: #978
+#362 := (iff #70 #361)
+#359 := (iff #69 #358)
+#360 := [rewrite]: #359
+#363 := [quant-intro #360]: #362
+#982 := [monotonicity #363 #979]: #981
+#988 := [trans #982 #986]: #987
+#991 := [monotonicity #988]: #990
+#995 := [trans #991 #993]: #994
+#997 := [monotonicity #995]: #996
+#999 := [trans #997 #993]: #998
+#1002 := [monotonicity #999]: #1001
+#1563 := [trans #1002 #1561]: #1562
+#357 := [asserted]: #298
+#1564 := [mp #357 #1563]: #1559
+#1566 := [not-or-elim #1564]: #367
+#1621 := [mp~ #1566 #1590]: #367
+#3748 := [mp #1621 #3747]: #3743
+#3348 := (not #3743)
+#6198 := (or #3348 #1703 #6178)
+#6194 := (or #1703 #6178)
+#6199 := (or #3348 #6194)
+#6202 := (iff #6199 #6198)
+#6174 := [rewrite]: #6202
+#6200 := [quant-inst #1702]: #6199
+#6180 := [mp #6200 #6174]: #6198
+#6206 := [unit-resolution #6180 #3748 #6201]: #6178
+#6208 := [mp #6206 #6185]: #6181
+#6219 := (not #6181)
+#1709 := (not #1708)
+#3684 := (or #3791 #1709)
+#3685 := [def-axiom]: #3684
+#6197 := [unit-resolution #3685 #6196]: #1709
+#6214 := (or #6219 #1708)
+#6220 := [th-lemma arith triangle-eq]: #6214
+#6221 := [unit-resolution #6220 #6197]: #6219
+#6209 := [unit-resolution #6221 #6208]: false
+#6210 := [lemma #6209]: #3791
+#4075 := (or #3794 #4072)
+#4078 := (not #4075)
+#2629 := (or #78 #1029 #1040)
+#3774 := (forall (vars (?v0 S2) (?v1 S2)) (:pat #3773) #2629)
+#3779 := (not #3774)
+#4081 := (or #3779 #4078)
+#4084 := (not #4081)
+decl ?v0!4 :: S2
+#1671 := ?v0!4
+#1684 := (f17 ?v0!4)
+#1685 := (* -1::Int #1684)
+decl ?v1!3 :: S2
+#1670 := ?v1!3
+#1683 := (f17 ?v1!3)
+#2262 := (+ #1683 #1685)
+#1674 := (f6 f7 ?v1!3)
+#1675 := (f5 #1674 ?v0!4)
+#1676 := (f15 #1675)
+#2263 := (+ #1676 #2262)
+#2266 := (>= #2263 0::Int)
+#1677 := (* -1::Int #1676)
+#1678 := (+ f14 #1677)
+#1679 := (<= #1678 0::Int)
+#1672 := (f9 f18 ?v1!3)
+#1673 := (= #1672 f1)
+#2592 := (not #1673)
+#2607 := (or #2592 #1679 #2266)
+#2612 := (not #2607)
+#4087 := (or #2612 #4084)
+#4090 := (not #4087)
+#3764 := (pattern #67 #87)
+#1760 := (not #85)
+#2584 := (or #77 #1760 #1014)
+#3765 := (forall (vars (?v0 S2) (?v1 S2)) (:pat #3764) #2584)
+#3770 := (not #3765)
+#4093 := (or #3770 #4090)
+#4096 := (not #4093)
+decl ?v0!2 :: S2
+#1644 := ?v0!2
+#1653 := (f17 ?v0!2)
+#1654 := (* -1::Int #1653)
+decl ?v1!1 :: S2
+#1643 := ?v1!1
+#1652 := (f17 ?v1!1)
+#1655 := (+ #1652 #1654)
+#1656 := (>= #1655 0::Int)
+#1648 := (f9 f18 ?v0!2)
+#1649 := (= #1648 f1)
+#2035 := (not #1649)
+#1645 := (f9 f18 ?v1!1)
+#1646 := (= #1645 f1)
+#2083 := (or #1646 #2035 #1656)
+#5518 := [hypothesis]: #1649
+#3750 := (forall (vars (?v0 S2)) (:pat #3749) #78)
+#3753 := (iff #79 #3750)
+#3751 := (iff #78 #78)
+#3752 := [refl]: #3751
+#3754 := [quant-intro #3752]: #3753
+#1591 := (~ #79 #79)
+#1622 := (~ #78 #78)
+#1623 := [refl]: #1622
+#1592 := [nnf-pos #1623]: #1591
+#1567 := [not-or-elim #1564]: #79
+#1624 := [mp~ #1567 #1592]: #79
+#3755 := [mp #1624 #3754]: #3750
+#7029 := (not #3750)
+#4457 := (or #7029 #2035)
+#4481 := [quant-inst #1644]: #4457
+#5558 := [unit-resolution #4481 #3755 #5518]: false
+#6078 := [lemma #5558]: #2035
+#3275 := (or #2083 #1649)
+#3361 := [def-axiom]: #3275
+#7086 := [unit-resolution #3361 #6078]: #2083
+#1634 := (not #2083)
+#4099 := (or #1634 #4096)
+#4102 := (not #4099)
+#3756 := (forall (vars (?v0 S2)) (:pat #3736) #1005)
+#3761 := (not #3756)
+#4105 := (or #3761 #4102)
+#4108 := (not #4105)
+decl ?v0!0 :: S2
+#1628 := ?v0!0
+#1629 := (f17 ?v0!0)
+#1630 := (>= #1629 0::Int)
+#1631 := (not #1630)
+#3358 := [hypothesis]: #1631
+#3357 := (<= #1629 0::Int)
+#4162 := (or #3357 #1630)
+#4163 := [th-lemma arith farkas 1 1]: #4162
+#4164 := [unit-resolution #4163 #3358]: #3357
+#4139 := (not #3357)
+#4158 := (or #4139 #1630)
+#3325 := (= f14 #1629)
+#3384 := (= #1629 f14)
+#4135 := (iff #3384 #3325)
+#3302 := (iff #3325 #3384)
+#4134 := [commutativity]: #3302
+#4136 := [symm #4134]: #4135
+#3398 := (= ?v0!0 f16)
+#3392 := (not #3398)
+#3393 := (= #1629 0::Int)
+#3356 := (not #3393)
+#3312 := (or #3356 #1630)
+#3311 := [th-lemma arith triangle-eq]: #3312
+#3317 := [unit-resolution #3311 #3358]: #3356
+#3737 := (forall (vars (?v0 S2)) (:pat #3736) #358)
+#3740 := (iff #361 #3737)
+#3738 := (iff #358 #358)
+#3739 := [refl]: #3738
+#3741 := [quant-intro #3739]: #3740
+#1587 := (~ #361 #361)
+#1616 := (~ #358 #358)
+#1617 := [refl]: #1616
+#1588 := [nnf-pos #1617]: #1587
+#1565 := [not-or-elim #1564]: #361
+#1618 := [mp~ #1565 #1588]: #361
+#3742 := [mp #1618 #3741]: #3737
+#3375 := (not #3737)
+#3379 := (or #3375 #3392 #3393)
+#3383 := (or #3392 #3393)
+#3370 := (or #3375 #3383)
+#3380 := (iff #3370 #3379)
+#3347 := [rewrite]: #3380
+#3378 := [quant-inst #1628]: #3370
+#3349 := [mp #3378 #3347]: #3379
+#3292 := [unit-resolution #3349 #3742 #3317]: #3392
+#3359 := (or #3348 #3398 #3384)
+#3394 := (or #3398 #3384)
+#3342 := (or #3348 #3394)
+#3335 := (iff #3342 #3359)
+#3333 := [rewrite]: #3335
+#3334 := [quant-inst #1628]: #3342
+#3336 := [mp #3334 #3333]: #3359
+#3297 := [unit-resolution #3336 #3748 #3292]: #3384
+#4137 := [mp #3297 #4136]: #3325
+#3405 := (* -1::Int #1629)
+#3337 := (+ f14 #3405)
+#3313 := (<= #3337 0::Int)
+#4133 := (not #3313)
+#326 := (<= f14 0::Int)
+#327 := (not #326)
+#55 := (< 0::Int f14)
+#328 := (iff #55 #327)
+#329 := [rewrite]: #328
+#323 := [asserted]: #55
+#330 := [mp #323 #329]: #327
+#4138 := [hypothesis]: #3357
+#4140 := (or #4133 #326 #4139)
+#4141 := [th-lemma arith assign-bounds 1 1]: #4140
+#4142 := [unit-resolution #4141 #4138 #330]: #4133
+#4143 := (not #3325)
+#4157 := (or #4143 #3313)
+#4159 := [th-lemma arith triangle-eq]: #4157
+#4160 := [unit-resolution #4159 #4142 #4137]: false
+#4161 := [lemma #4160]: #4158
+#4165 := [unit-resolution #4161 #4164 #3358]: false
+#4166 := [lemma #4165]: #1630
+#4111 := (or #1631 #4108)
+#4114 := (not #4111)
+#4117 := (or #946 #4114)
+#4120 := (not #4117)
+#4240 := [hypothesis]: #946
+#4232 := (or #3375 #81)
+#6915 := (= f16 f16)
+#6988 := (not #6915)
+#4242 := (or #6988 #81)
+#4233 := (or #3375 #4242)
+#4235 := (iff #4233 #4232)
+#4234 := (iff #4232 #4232)
+#4237 := [rewrite]: #4234
+#4248 := (iff #4242 #81)
+#4214 := (or false #81)
+#4243 := (iff #4214 #81)
+#4247 := [rewrite]: #4243
+#4244 := (iff #4242 #4214)
+#6996 := (iff #6988 false)
+#6991 := (not true)
+#6994 := (iff #6991 false)
+#6995 := [rewrite]: #6994
+#6992 := (iff #6988 #6991)
+#6918 := (iff #6915 true)
+#6919 := [rewrite]: #6918
+#6993 := [monotonicity #6919]: #6992
+#6997 := [trans #6993 #6995]: #6996
+#4245 := [monotonicity #6997]: #4244
+#4246 := [trans #4245 #4247]: #4248
+#4236 := [monotonicity #4246]: #4235
+#4238 := [trans #4236 #4237]: #4235
+#4249 := [quant-inst #65]: #4233
+#4231 := [mp #4249 #4238]: #4232
+#4241 := [unit-resolution #4231 #3742 #4240]: false
+#4239 := [lemma #4241]: #81
+#4123 := (or #946 #4120)
+#3042 := (forall (vars (?v1 S2)) #3031)
+#3049 := (not #3042)
+#3027 := (forall (vars (?v0 S2) (?v1 S2)) #3022)
+#3048 := (not #3027)
+#3050 := (or #3048 #2124 #2129 #3049)
+#3051 := (not #3050)
+#3056 := (or #3005 #3051)
+#3063 := (not #3056)
+#2982 := (forall (vars (?v0 S2) (?v1 S2)) #2977)
+#3062 := (not #2982)
+#3064 := (or #3062 #3063)
+#3065 := (not #3064)
+#3070 := (or #2959 #3065)
+#3076 := (not #3070)
+#3077 := (or #1271 #3076)
+#3078 := (not #3077)
+#3083 := (or #2052 #3078)
+#3089 := (not #3083)
+#3090 := (or #713 #3089)
+#3091 := (not #3090)
+#3096 := (or #713 #3091)
+#3102 := (not #3096)
+#3103 := (or #736 #3102)
+#3104 := (not #3103)
+#3109 := (or #2032 #3104)
+#3115 := (not #3109)
+#3116 := (or #1262 #3115)
+#3117 := (not #3116)
+#3122 := (or #2012 #3117)
+#3130 := (not #3122)
+#2936 := (forall (vars (?v0 S2)) #2933)
+#3129 := (not #2936)
+#2930 := (forall (vars (?v0 S2)) #2925)
+#3128 := (not #2930)
+#3131 := (or #1963 #1968 #188 #1411 #1407 #778 #3128 #3129 #3130)
+#3132 := (not #3131)
+#2870 := (forall (vars (?v0 S2) (?v1 S2)) #2865)
+#2876 := (not #2870)
+#2877 := (or #2876 #169)
+#2878 := (not #2877)
+#2883 := (or #2848 #2878)
+#2890 := (not #2883)
+#2826 := (forall (vars (?v0 S2)) #2821)
+#2889 := (not #2826)
+#2891 := (or #2889 #2890)
+#2892 := (not #2891)
+#2789 := (forall (vars (?v1 S2)) #2778)
+#2795 := (not #2789)
+#2796 := (or #1841 #1846 #2795)
+#2797 := (not #2796)
+#2897 := (or #2797 #2892)
+#2904 := (not #2897)
+#2774 := (forall (vars (?v0 S2)) #2763)
+#2903 := (not #2774)
+#2905 := (or #2903 #549 #540 #531 #522 #2904)
+#2906 := (not #2905)
+#3137 := (or #2906 #3132)
+#3147 := (not #3137)
+#2760 := (forall (vars (?v0 S2)) #2755)
+#3146 := (not #2760)
+#2732 := (forall (vars (?v0 S2) (?v1 S2)) #2727)
+#3145 := (not #2732)
+#2710 := (forall (vars (?v0 S2) (?v1 S2)) #2705)
+#3144 := (not #2710)
+#2687 := (forall (vars (?v0 S2)) #2682)
+#3143 := (not #2687)
+#3148 := (or #3143 #878 #1489 #3144 #3145 #3146 #3147)
+#3149 := (not #3148)
+#2649 := (forall (vars (?v1 S2)) #2638)
+#2655 := (not #2649)
+#2656 := (or #1703 #1708 #2655)
+#2657 := (not #2656)
+#3154 := (or #2657 #3149)
+#3161 := (not #3154)
+#2634 := (forall (vars (?v0 S2) (?v1 S2)) #2629)
+#3160 := (not #2634)
+#3162 := (or #3160 #3161)
+#3163 := (not #3162)
+#3168 := (or #2612 #3163)
+#3175 := (not #3168)
+#2589 := (forall (vars (?v0 S2) (?v1 S2)) #2584)
+#3174 := (not #2589)
+#3176 := (or #3174 #3175)
+#3177 := (not #3176)
+#3182 := (or #1634 #3177)
+#3188 := (not #3182)
+#3189 := (or #1009 #3188)
+#3190 := (not #3189)
+#3195 := (or #1631 #3190)
+#3201 := (not #3195)
+#3202 := (or #946 #3201)
+#3203 := (not #3202)
+#3208 := (or #946 #3203)
+#4124 := (iff #3208 #4123)
+#4121 := (iff #3203 #4120)
+#4118 := (iff #3202 #4117)
+#4115 := (iff #3201 #4114)
+#4112 := (iff #3195 #4111)
+#4109 := (iff #3190 #4108)
+#4106 := (iff #3189 #4105)
+#4103 := (iff #3188 #4102)
+#4100 := (iff #3182 #4099)
+#4097 := (iff #3177 #4096)
+#4094 := (iff #3176 #4093)
+#4091 := (iff #3175 #4090)
+#4088 := (iff #3168 #4087)
+#4085 := (iff #3163 #4084)
+#4082 := (iff #3162 #4081)
+#4079 := (iff #3161 #4078)
+#4076 := (iff #3154 #4075)
+#4073 := (iff #3149 #4072)
+#4070 := (iff #3148 #4069)
+#4067 := (iff #3147 #4066)
+#4064 := (iff #3137 #4063)
+#4061 := (iff #3132 #4060)
+#4058 := (iff #3131 #4057)
+#4055 := (iff #3130 #4054)
+#4052 := (iff #3122 #4051)
+#4049 := (iff #3117 #4048)
+#4046 := (iff #3116 #4045)
+#4043 := (iff #3115 #4042)
+#4040 := (iff #3109 #4039)
+#4037 := (iff #3104 #4036)
+#4034 := (iff #3103 #4033)
+#4031 := (iff #3102 #4030)
+#4028 := (iff #3096 #4027)
+#4025 := (iff #3091 #4024)
+#4022 := (iff #3090 #4021)
+#4019 := (iff #3089 #4018)
+#4016 := (iff #3083 #4015)
+#4013 := (iff #3078 #4012)
+#4010 := (iff #3077 #4009)
+#4007 := (iff #3076 #4006)
+#4004 := (iff #3070 #4003)
+#4001 := (iff #3065 #4000)
+#3998 := (iff #3064 #3997)
+#3995 := (iff #3063 #3994)
+#3992 := (iff #3056 #3991)
+#3989 := (iff #3051 #3988)
+#3986 := (iff #3050 #3985)
+#3983 := (iff #3049 #3982)
+#3980 := (iff #3042 #3977)
+#3978 := (iff #3031 #3031)
+#3979 := [refl]: #3978
+#3981 := [quant-intro #3979]: #3980
+#3984 := [monotonicity #3981]: #3983
+#3975 := (iff #3048 #3974)
+#3972 := (iff #3027 #3969)
+#3970 := (iff #3022 #3022)
+#3971 := [refl]: #3970
+#3973 := [quant-intro #3971]: #3972
+#3976 := [monotonicity #3973]: #3975
+#3987 := [monotonicity #3976 #3984]: #3986
+#3990 := [monotonicity #3987]: #3989
+#3993 := [monotonicity #3990]: #3992
+#3996 := [monotonicity #3993]: #3995
+#3967 := (iff #3062 #3966)
+#3964 := (iff #2982 #3961)
+#3962 := (iff #2977 #2977)
+#3963 := [refl]: #3962
+#3965 := [quant-intro #3963]: #3964
+#3968 := [monotonicity #3965]: #3967
+#3999 := [monotonicity #3968 #3996]: #3998
+#4002 := [monotonicity #3999]: #4001
+#4005 := [monotonicity #4002]: #4004
+#4008 := [monotonicity #4005]: #4007
+#3958 := (iff #1271 #3957)
+#3955 := (iff #1268 #3952)
+#3953 := (iff #1265 #1265)
+#3954 := [refl]: #3953
+#3956 := [quant-intro #3954]: #3955
+#3959 := [monotonicity #3956]: #3958
+#4011 := [monotonicity #3959 #4008]: #4010
+#4014 := [monotonicity #4011]: #4013
+#4017 := [monotonicity #4014]: #4016
+#4020 := [monotonicity #4017]: #4019
+#4023 := [monotonicity #4020]: #4022
+#4026 := [monotonicity #4023]: #4025
+#4029 := [monotonicity #4026]: #4028
+#4032 := [monotonicity #4029]: #4031
+#3950 := (iff #736 #3949)
+#3947 := (iff #606 #3944)
+#3945 := (iff #603 #603)
+#3946 := [refl]: #3945
+#3948 := [quant-intro #3946]: #3947
+#3951 := [monotonicity #3948]: #3950
+#4035 := [monotonicity #3951 #4032]: #4034
+#4038 := [monotonicity #4035]: #4037
+#4041 := [monotonicity #4038]: #4040
+#4044 := [monotonicity #4041]: #4043
+#3941 := (iff #1262 #3940)
+#3938 := (iff #1259 #3935)
+#3936 := (iff #1254 #1254)
+#3937 := [refl]: #3936
+#3939 := [quant-intro #3937]: #3938
+#3942 := [monotonicity #3939]: #3941
+#4047 := [monotonicity #3942 #4044]: #4046
+#4050 := [monotonicity #4047]: #4049
+#4053 := [monotonicity #4050]: #4052
+#4056 := [monotonicity #4053]: #4055
+#3933 := (iff #3129 #3932)
+#3930 := (iff #2936 #3927)
+#3928 := (iff #2933 #2933)
+#3929 := [refl]: #3928
+#3931 := [quant-intro #3929]: #3930
+#3934 := [monotonicity #3931]: #3933
+#3925 := (iff #3128 #3924)
+#3922 := (iff #2930 #3919)
+#3920 := (iff #2925 #2925)
+#3921 := [refl]: #3920
+#3923 := [quant-intro #3921]: #3922
+#3926 := [monotonicity #3923]: #3925
+#3915 := (iff #1407 #3914)
+#3912 := (iff #1404 #3909)
+#3910 := (iff #1401 #1401)
+#3911 := [refl]: #3910
+#3913 := [quant-intro #3911]: #3912
+#3916 := [monotonicity #3913]: #3915
+#4059 := [monotonicity #3916 #3926 #3934 #4056]: #4058
+#4062 := [monotonicity #4059]: #4061
+#3907 := (iff #2906 #3906)
+#3904 := (iff #2905 #3903)
+#3901 := (iff #2904 #3900)
+#3898 := (iff #2897 #3897)
+#3895 := (iff #2892 #3894)
+#3892 := (iff #2891 #3891)
+#3889 := (iff #2890 #3888)
+#3886 := (iff #2883 #3885)
+#3883 := (iff #2878 #3882)
+#3880 := (iff #2877 #3879)
+#3877 := (iff #2876 #3876)
+#3874 := (iff #2870 #3871)
+#3872 := (iff #2865 #2865)
+#3873 := [refl]: #3872
+#3875 := [quant-intro #3873]: #3874
+#3878 := [monotonicity #3875]: #3877
+#3881 := [monotonicity #3878]: #3880
+#3884 := [monotonicity #3881]: #3883
+#3887 := [monotonicity #3884]: #3886
+#3890 := [monotonicity #3887]: #3889
+#3869 := (iff #2889 #3868)
+#3866 := (iff #2826 #3863)
+#3864 := (iff #2821 #2821)
+#3865 := [refl]: #3864
+#3867 := [quant-intro #3865]: #3866
+#3870 := [monotonicity #3867]: #3869
+#3893 := [monotonicity #3870 #3890]: #3892
+#3896 := [monotonicity #3893]: #3895
+#3861 := (iff #2797 #3860)
+#3858 := (iff #2796 #3857)
+#3855 := (iff #2795 #3854)
+#3852 := (iff #2789 #3849)
+#3850 := (iff #2778 #2778)
+#3851 := [refl]: #3850
+#3853 := [quant-intro #3851]: #3852
+#3856 := [monotonicity #3853]: #3855
+#3859 := [monotonicity #3856]: #3858
+#3862 := [monotonicity #3859]: #3861
+#3899 := [monotonicity #3862 #3896]: #3898
+#3902 := [monotonicity #3899]: #3901
+#3846 := (iff #2903 #3845)
+#3843 := (iff #2774 #3840)
+#3841 := (iff #2763 #2763)
+#3842 := [refl]: #3841
+#3844 := [quant-intro #3842]: #3843
+#3847 := [monotonicity #3844]: #3846
+#3905 := [monotonicity #3847 #3902]: #3904
+#3908 := [monotonicity #3905]: #3907
+#4065 := [monotonicity #3908 #4062]: #4064
+#4068 := [monotonicity #4065]: #4067
+#3837 := (iff #3146 #3836)
+#3834 := (iff #2760 #3831)
+#3832 := (iff #2755 #2755)
+#3833 := [refl]: #3832
+#3835 := [quant-intro #3833]: #3834
+#3838 := [monotonicity #3835]: #3837
+#3829 := (iff #3145 #3828)
+#3826 := (iff #2732 #3823)
+#3824 := (iff #2727 #2727)
+#3825 := [refl]: #3824
+#3827 := [quant-intro #3825]: #3826
+#3830 := [monotonicity #3827]: #3829
+#3821 := (iff #3144 #3820)
+#3818 := (iff #2710 #3815)
+#3816 := (iff #2705 #2705)
+#3817 := [refl]: #3816
+#3819 := [quant-intro #3817]: #3818
+#3822 := [monotonicity #3819]: #3821
+#3812 := (iff #1489 #3811)
+#3809 := (iff #1486 #3806)
+#3807 := (iff #1483 #1483)
+#3808 := [refl]: #3807
+#3810 := [quant-intro #3808]: #3809
+#3813 := [monotonicity #3810]: #3812
+#3803 := (iff #3143 #3802)
+#3800 := (iff #2687 #3797)
+#3798 := (iff #2682 #2682)
+#3799 := [refl]: #3798
+#3801 := [quant-intro #3799]: #3800
+#3804 := [monotonicity #3801]: #3803
+#4071 := [monotonicity #3804 #3813 #3822 #3830 #3838 #4068]: #4070
+#4074 := [monotonicity #4071]: #4073
+#3795 := (iff #2657 #3794)
+#3792 := (iff #2656 #3791)
+#3789 := (iff #2655 #3788)
+#3786 := (iff #2649 #3783)
+#3784 := (iff #2638 #2638)
+#3785 := [refl]: #3784
+#3787 := [quant-intro #3785]: #3786
+#3790 := [monotonicity #3787]: #3789
+#3793 := [monotonicity #3790]: #3792
+#3796 := [monotonicity #3793]: #3795
+#4077 := [monotonicity #3796 #4074]: #4076
+#4080 := [monotonicity #4077]: #4079
+#3780 := (iff #3160 #3779)
+#3777 := (iff #2634 #3774)
+#3775 := (iff #2629 #2629)
+#3776 := [refl]: #3775
+#3778 := [quant-intro #3776]: #3777
+#3781 := [monotonicity #3778]: #3780
+#4083 := [monotonicity #3781 #4080]: #4082
+#4086 := [monotonicity #4083]: #4085
+#4089 := [monotonicity #4086]: #4088
+#4092 := [monotonicity #4089]: #4091
+#3771 := (iff #3174 #3770)
+#3768 := (iff #2589 #3765)
+#3766 := (iff #2584 #2584)
+#3767 := [refl]: #3766
+#3769 := [quant-intro #3767]: #3768
+#3772 := [monotonicity #3769]: #3771
+#4095 := [monotonicity #3772 #4092]: #4094
+#4098 := [monotonicity #4095]: #4097
+#4101 := [monotonicity #4098]: #4100
+#4104 := [monotonicity #4101]: #4103
+#3762 := (iff #1009 #3761)
+#3759 := (iff #1006 #3756)
+#3757 := (iff #1005 #1005)
+#3758 := [refl]: #3757
+#3760 := [quant-intro #3758]: #3759
+#3763 := [monotonicity #3760]: #3762
+#4107 := [monotonicity #3763 #4104]: #4106
+#4110 := [monotonicity #4107]: #4109
+#4113 := [monotonicity #4110]: #4112
+#4116 := [monotonicity #4113]: #4115
+#4119 := [monotonicity #4116]: #4118
+#4122 := [monotonicity #4119]: #4121
+#4125 := [monotonicity #4122]: #4124
+#2135 := (not #2134)
+#2479 := (and #2135 #218 #2476)
+#2482 := (not #2479)
+#2485 := (forall (vars (?v1 S2)) #2482)
+#2130 := (not #2129)
+#2125 := (not #2124)
+#2494 := (and #1301 #2125 #2130 #2485)
+#2101 := (not #2100)
+#2102 := (and #2094 #2101)
+#2103 := (not #2102)
+#2110 := (or #2103 #2109)
+#2111 := (not #2110)
+#2499 := (or #2111 #2494)
+#2502 := (and #1282 #2499)
+#2068 := (not #2067)
+#2071 := (and #2068 #2070)
+#2072 := (not #2071)
+#2078 := (or #2072 #2077)
+#2079 := (not #2078)
+#2505 := (or #2079 #2502)
+#2508 := (and #1268 #2505)
+#2511 := (or #2052 #2508)
+#2514 := (and #222 #2511)
+#2517 := (or #713 #2514)
+#2520 := (and #606 #2517)
+#2523 := (or #2032 #2520)
+#2526 := (and #1259 #2523)
+#2529 := (or #2012 #2526)
+#1969 := (not #1968)
+#1964 := (not #1963)
+#2535 := (and #1964 #1969 #189 #1412 #1404 #199 #1391 #1248 #2529)
+#1938 := (not #169)
+#1941 := (and #1185 #1938)
+#1917 := (not #1916)
+#1910 := (not #1909)
+#1918 := (and #1910 #1917)
+#1919 := (not #1918)
+#2448 := (or #1919 #2445)
+#2451 := (not #2448)
+#2454 := (or #2451 #1941)
+#2414 := (not #2409)
+#2432 := (and #2414 #2427)
+#2435 := (or #1145 #2432)
+#2438 := (forall (vars (?v0 S2)) #2435)
+#2457 := (and #2438 #2454)
+#1852 := (not #1851)
+#2384 := (and #1852 #2381)
+#2387 := (not #2384)
+#2390 := (forall (vars (?v1 S2)) #2387)
+#1847 := (not #1846)
+#1842 := (not #1841)
+#2396 := (and #1842 #1847 #2390)
+#2460 := (or #2396 #2457)
+#1822 := (not #1203)
+#1825 := (forall (vars (?v0 S2)) #1822)
+#2463 := (and #1825 #144 #147 #149 #152 #2460)
+#2540 := (or #2463 #2535)
+#2340 := (not #2335)
+#2358 := (and #2340 #1802 #2353)
+#2361 := (or #1100 #2358)
+#2364 := (forall (vars (?v0 S2)) #2361)
+#2298 := (not #2293)
+#2316 := (and #2298 #1749 #2311)
+#2319 := (or #1062 #2316)
+#2322 := (forall (vars (?v0 S2)) #2319)
+#2543 := (and #2322 #111 #1486 #1477 #1466 #2364 #2540)
+#1714 := (not #1713)
+#1720 := (and #1714 #77 #1719)
+#1729 := (not #1720)
+#1732 := (forall (vars (?v1 S2)) #1729)
+#2280 := (and #1704 #1709 #1732)
+#2546 := (or #2280 #2543)
+#2549 := (and #1047 #2546)
+#1680 := (not #1679)
+#1681 := (and #1673 #1680)
+#1682 := (not #1681)
+#2269 := (or #1682 #2266)
+#2272 := (not #2269)
+#2552 := (or #2272 #2549)
+#2555 := (and #1021 #2552)
+#1647 := (not #1646)
+#1650 := (and #1647 #1649)
+#1651 := (not #1650)
+#1657 := (or #1651 #1656)
+#1658 := (not #1657)
+#2558 := (or #1658 #2555)
+#2561 := (and #1006 #2558)
+#2564 := (or #1631 #2561)
+#2567 := (and #81 #2564)
+#2570 := (or #946 #2567)
+#3209 := (iff #2570 #3208)
+#3206 := (iff #2567 #3203)
+#3198 := (and #81 #3195)
+#3204 := (iff #3198 #3203)
+#3205 := [rewrite]: #3204
+#3199 := (iff #2567 #3198)
+#3196 := (iff #2564 #3195)
+#3193 := (iff #2561 #3190)
+#3185 := (and #1006 #3182)
+#3191 := (iff #3185 #3190)
+#3192 := [rewrite]: #3191
+#3186 := (iff #2561 #3185)
+#3183 := (iff #2558 #3182)
+#3180 := (iff #2555 #3177)
+#3171 := (and #2589 #3168)
+#3178 := (iff #3171 #3177)
+#3179 := [rewrite]: #3178
+#3172 := (iff #2555 #3171)
+#3169 := (iff #2552 #3168)
+#3166 := (iff #2549 #3163)
+#3157 := (and #2634 #3154)
+#3164 := (iff #3157 #3163)
+#3165 := [rewrite]: #3164
+#3158 := (iff #2549 #3157)
+#3155 := (iff #2546 #3154)
+#3152 := (iff #2543 #3149)
+#3140 := (and #2687 #111 #1486 #2710 #2732 #2760 #3137)
+#3150 := (iff #3140 #3149)
+#3151 := [rewrite]: #3150
+#3141 := (iff #2543 #3140)
+#3138 := (iff #2540 #3137)
+#3135 := (iff #2535 #3132)
+#3125 := (and #1964 #1969 #189 #1412 #1404 #199 #2930 #2936 #3122)
+#3133 := (iff #3125 #3132)
+#3134 := [rewrite]: #3133
+#3126 := (iff #2535 #3125)
+#3123 := (iff #2529 #3122)
+#3120 := (iff #2526 #3117)
+#3112 := (and #1259 #3109)
+#3118 := (iff #3112 #3117)
+#3119 := [rewrite]: #3118
+#3113 := (iff #2526 #3112)
+#3110 := (iff #2523 #3109)
+#3107 := (iff #2520 #3104)
+#3099 := (and #606 #3096)
+#3105 := (iff #3099 #3104)
+#3106 := [rewrite]: #3105
+#3100 := (iff #2520 #3099)
+#3097 := (iff #2517 #3096)
+#3094 := (iff #2514 #3091)
+#3086 := (and #222 #3083)
+#3092 := (iff #3086 #3091)
+#3093 := [rewrite]: #3092
+#3087 := (iff #2514 #3086)
+#3084 := (iff #2511 #3083)
+#3081 := (iff #2508 #3078)
+#3073 := (and #1268 #3070)
+#3079 := (iff #3073 #3078)
+#3080 := [rewrite]: #3079
+#3074 := (iff #2508 #3073)
+#3071 := (iff #2505 #3070)
+#3068 := (iff #2502 #3065)
+#3059 := (and #2982 #3056)
+#3066 := (iff #3059 #3065)
+#3067 := [rewrite]: #3066
+#3060 := (iff #2502 #3059)
+#3057 := (iff #2499 #3056)
+#3054 := (iff #2494 #3051)
+#3045 := (and #3027 #2125 #2130 #3042)
+#3052 := (iff #3045 #3051)
+#3053 := [rewrite]: #3052
+#3046 := (iff #2494 #3045)
+#3043 := (iff #2485 #3042)
+#3040 := (iff #2482 #3031)
+#3032 := (not #3031)
+#3035 := (not #3032)
+#3038 := (iff #3035 #3031)
+#3039 := [rewrite]: #3038
+#3036 := (iff #2482 #3035)
+#3033 := (iff #2479 #3032)
+#3034 := [rewrite]: #3033
+#3037 := [monotonicity #3034]: #3036
+#3041 := [trans #3037 #3039]: #3040
+#3044 := [quant-intro #3041]: #3043
+#3028 := (iff #1301 #3027)
+#3025 := (iff #1298 #3022)
+#3008 := (or #225 #1029)
+#3019 := (or #3008 #1294)
+#3023 := (iff #3019 #3022)
+#3024 := [rewrite]: #3023
+#3020 := (iff #1298 #3019)
+#3017 := (iff #1291 #3008)
+#3009 := (not #3008)
+#3012 := (not #3009)
+#3015 := (iff #3012 #3008)
+#3016 := [rewrite]: #3015
+#3013 := (iff #1291 #3012)
+#3010 := (iff #1288 #3009)
+#3011 := [rewrite]: #3010
+#3014 := [monotonicity #3011]: #3013
+#3018 := [trans #3014 #3016]: #3017
+#3021 := [monotonicity #3018]: #3020
+#3026 := [trans #3021 #3024]: #3025
+#3029 := [quant-intro #3026]: #3028
+#3047 := [monotonicity #3029 #3044]: #3046
+#3055 := [trans #3047 #3053]: #3054
+#3006 := (iff #2111 #3005)
+#3003 := (iff #2110 #3000)
+#2986 := (or #2985 #2100)
+#2997 := (or #2986 #2109)
+#3001 := (iff #2997 #3000)
+#3002 := [rewrite]: #3001
+#2998 := (iff #2110 #2997)
+#2995 := (iff #2103 #2986)
+#2987 := (not #2986)
+#2990 := (not #2987)
+#2993 := (iff #2990 #2986)
+#2994 := [rewrite]: #2993
+#2991 := (iff #2103 #2990)
+#2988 := (iff #2102 #2987)
+#2989 := [rewrite]: #2988
+#2992 := [monotonicity #2989]: #2991
+#2996 := [trans #2992 #2994]: #2995
+#2999 := [monotonicity #2996]: #2998
+#3004 := [trans #2999 #3002]: #3003
+#3007 := [monotonicity #3004]: #3006
+#3058 := [monotonicity #3007 #3055]: #3057
+#2983 := (iff #1282 #2982)
+#2980 := (iff #1279 #2977)
+#2963 := (or #218 #2962)
+#2974 := (or #2963 #1274)
+#2978 := (iff #2974 #2977)
+#2979 := [rewrite]: #2978
+#2975 := (iff #1279 #2974)
+#2972 := (iff #609 #2963)
+#2964 := (not #2963)
+#2967 := (not #2964)
+#2970 := (iff #2967 #2963)
+#2971 := [rewrite]: #2970
+#2968 := (iff #609 #2967)
+#2965 := (iff #228 #2964)
+#2966 := [rewrite]: #2965
+#2969 := [monotonicity #2966]: #2968
+#2973 := [trans #2969 #2971]: #2972
+#2976 := [monotonicity #2973]: #2975
+#2981 := [trans #2976 #2979]: #2980
+#2984 := [quant-intro #2981]: #2983
+#3061 := [monotonicity #2984 #3058]: #3060
+#3069 := [trans #3061 #3067]: #3068
+#2960 := (iff #2079 #2959)
+#2957 := (iff #2078 #2954)
+#2940 := (or #2067 #2939)
+#2951 := (or #2940 #2077)
+#2955 := (iff #2951 #2954)
+#2956 := [rewrite]: #2955
+#2952 := (iff #2078 #2951)
+#2949 := (iff #2072 #2940)
+#2941 := (not #2940)
+#2944 := (not #2941)
+#2947 := (iff #2944 #2940)
+#2948 := [rewrite]: #2947
+#2945 := (iff #2072 #2944)
+#2942 := (iff #2071 #2941)
+#2943 := [rewrite]: #2942
+#2946 := [monotonicity #2943]: #2945
+#2950 := [trans #2946 #2948]: #2949
+#2953 := [monotonicity #2950]: #2952
+#2958 := [trans #2953 #2956]: #2957
+#2961 := [monotonicity #2958]: #2960
+#3072 := [monotonicity #2961 #3069]: #3071
+#3075 := [monotonicity #3072]: #3074
+#3082 := [trans #3075 #3080]: #3081
+#3085 := [monotonicity #3082]: #3084
+#3088 := [monotonicity #3085]: #3087
+#3095 := [trans #3088 #3093]: #3094
+#3098 := [monotonicity #3095]: #3097
+#3101 := [monotonicity #3098]: #3100
+#3108 := [trans #3101 #3106]: #3107
+#3111 := [monotonicity #3108]: #3110
+#3114 := [monotonicity #3111]: #3113
+#3121 := [trans #3114 #3119]: #3120
+#3124 := [monotonicity #3121]: #3123
+#2937 := (iff #1248 #2936)
+#2934 := (iff #1245 #2933)
+#2913 := (iff #1242 #2912)
+#2914 := [rewrite]: #2913
+#2935 := [monotonicity #2914]: #2934
+#2938 := [quant-intro #2935]: #2937
+#2931 := (iff #1391 #2930)
+#2928 := (iff #1388 #2925)
+#2922 := (or #2911 #1383)
+#2926 := (iff #2922 #2925)
+#2927 := [rewrite]: #2926
+#2923 := (iff #1388 #2922)
+#2920 := (iff #1380 #2911)
+#2915 := (not #2912)
+#2918 := (iff #2915 #2911)
+#2919 := [rewrite]: #2918
+#2916 := (iff #1380 #2915)
+#2917 := [monotonicity #2914]: #2916
+#2921 := [trans #2917 #2919]: #2920
+#2924 := [monotonicity #2921]: #2923
+#2929 := [trans #2924 #2927]: #2928
+#2932 := [quant-intro #2929]: #2931
+#3127 := [monotonicity #2932 #2938 #3124]: #3126
+#3136 := [trans #3127 #3134]: #3135
+#2909 := (iff #2463 #2906)
+#2900 := (and #2774 #144 #147 #149 #152 #2897)
+#2907 := (iff #2900 #2906)
+#2908 := [rewrite]: #2907
+#2901 := (iff #2463 #2900)
+#2898 := (iff #2460 #2897)
+#2895 := (iff #2457 #2892)
+#2886 := (and #2826 #2883)
+#2893 := (iff #2886 #2892)
+#2894 := [rewrite]: #2893
+#2887 := (iff #2457 #2886)
+#2884 := (iff #2454 #2883)
+#2881 := (iff #1941 #2878)
+#2873 := (and #2870 #1938)
+#2879 := (iff #2873 #2878)
+#2880 := [rewrite]: #2879
+#2874 := (iff #1941 #2873)
+#2871 := (iff #1185 #2870)
+#2868 := (iff #1182 #2865)
+#2851 := (or #1138 #1029)
+#2862 := (or #2851 #1179)
+#2866 := (iff #2862 #2865)
+#2867 := [rewrite]: #2866
+#2863 := (iff #1182 #2862)
+#2860 := (iff #1176 #2851)
+#2852 := (not #2851)
+#2855 := (not #2852)
+#2858 := (iff #2855 #2851)
+#2859 := [rewrite]: #2858
+#2856 := (iff #1176 #2855)
+#2853 := (iff #1173 #2852)
+#2854 := [rewrite]: #2853
+#2857 := [monotonicity #2854]: #2856
+#2861 := [trans #2857 #2859]: #2860
+#2864 := [monotonicity #2861]: #2863
+#2869 := [trans #2864 #2867]: #2868
+#2872 := [quant-intro #2869]: #2871
+#2875 := [monotonicity #2872]: #2874
+#2882 := [trans #2875 #2880]: #2881
+#2849 := (iff #2451 #2848)
+#2846 := (iff #2448 #2843)
+#2829 := (or #1909 #1916)
+#2840 := (or #2829 #2445)
+#2844 := (iff #2840 #2843)
+#2845 := [rewrite]: #2844
+#2841 := (iff #2448 #2840)
+#2838 := (iff #1919 #2829)
+#2830 := (not #2829)
+#2833 := (not #2830)
+#2836 := (iff #2833 #2829)
+#2837 := [rewrite]: #2836
+#2834 := (iff #1919 #2833)
+#2831 := (iff #1918 #2830)
+#2832 := [rewrite]: #2831
+#2835 := [monotonicity #2832]: #2834
+#2839 := [trans #2835 #2837]: #2838
+#2842 := [monotonicity #2839]: #2841
+#2847 := [trans #2842 #2845]: #2846
+#2850 := [monotonicity #2847]: #2849
+#2885 := [monotonicity #2850 #2882]: #2884
+#2827 := (iff #2438 #2826)
+#2824 := (iff #2435 #2821)
+#2802 := (or #66 #1138)
+#2818 := (or #2802 #2815)
+#2822 := (iff #2818 #2821)
+#2823 := [rewrite]: #2822
+#2819 := (iff #2435 #2818)
+#2816 := (iff #2432 #2815)
+#2817 := [rewrite]: #2816
+#2811 := (iff #1145 #2802)
+#2803 := (not #2802)
+#2806 := (not #2803)
+#2809 := (iff #2806 #2802)
+#2810 := [rewrite]: #2809
+#2807 := (iff #1145 #2806)
+#2804 := (iff #1142 #2803)
+#2805 := [rewrite]: #2804
+#2808 := [monotonicity #2805]: #2807
+#2812 := [trans #2808 #2810]: #2811
+#2820 := [monotonicity #2812 #2817]: #2819
+#2825 := [trans #2820 #2823]: #2824
+#2828 := [quant-intro #2825]: #2827
+#2888 := [monotonicity #2828 #2885]: #2887
+#2896 := [trans #2888 #2894]: #2895
+#2800 := (iff #2396 #2797)
+#2792 := (and #1842 #1847 #2789)
+#2798 := (iff #2792 #2797)
+#2799 := [rewrite]: #2798
+#2793 := (iff #2396 #2792)
+#2790 := (iff #2390 #2789)
+#2787 := (iff #2387 #2778)
+#2779 := (not #2778)
+#2782 := (not #2779)
+#2785 := (iff #2782 #2778)
+#2786 := [rewrite]: #2785
+#2783 := (iff #2387 #2782)
+#2780 := (iff #2384 #2779)
+#2781 := [rewrite]: #2780
+#2784 := [monotonicity #2781]: #2783
+#2788 := [trans #2784 #2786]: #2787
+#2791 := [quant-intro #2788]: #2790
+#2794 := [monotonicity #2791]: #2793
+#2801 := [trans #2794 #2799]: #2800
+#2899 := [monotonicity #2801 #2896]: #2898
+#2775 := (iff #1825 #2774)
+#2772 := (iff #1822 #2763)
+#2764 := (not #2763)
+#2767 := (not #2764)
+#2770 := (iff #2767 #2763)
+#2771 := [rewrite]: #2770
+#2768 := (iff #1822 #2767)
+#2765 := (iff #1203 #2764)
+#2766 := [rewrite]: #2765
+#2769 := [monotonicity #2766]: #2768
+#2773 := [trans #2769 #2771]: #2772
+#2776 := [quant-intro #2773]: #2775
+#2902 := [monotonicity #2776 #2899]: #2901
+#2910 := [trans #2902 #2908]: #2909
+#3139 := [monotonicity #2910 #3136]: #3138
+#2761 := (iff #2364 #2760)
+#2758 := (iff #2361 #2755)
+#2735 := (or #66 #1093)
+#2752 := (or #2735 #2749)
+#2756 := (iff #2752 #2755)
+#2757 := [rewrite]: #2756
+#2753 := (iff #2361 #2752)
+#2750 := (iff #2358 #2749)
+#2751 := [rewrite]: #2750
+#2744 := (iff #1100 #2735)
+#2736 := (not #2735)
+#2739 := (not #2736)
+#2742 := (iff #2739 #2735)
+#2743 := [rewrite]: #2742
+#2740 := (iff #1100 #2739)
+#2737 := (iff #1097 #2736)
+#2738 := [rewrite]: #2737
+#2741 := [monotonicity #2738]: #2740
+#2745 := [trans #2741 #2743]: #2744
+#2754 := [monotonicity #2745 #2751]: #2753
+#2759 := [trans #2754 #2757]: #2758
+#2762 := [quant-intro #2759]: #2761
+#2733 := (iff #1466 #2732)
+#2730 := (iff #1463 #2727)
+#2713 := (or #118 #1029)
+#2724 := (or #2713 #1460)
+#2728 := (iff #2724 #2727)
+#2729 := [rewrite]: #2728
+#2725 := (iff #1463 #2724)
+#2722 := (iff #1457 #2713)
+#2714 := (not #2713)
+#2717 := (not #2714)
+#2720 := (iff #2717 #2713)
+#2721 := [rewrite]: #2720
+#2718 := (iff #1457 #2717)
+#2715 := (iff #1454 #2714)
+#2716 := [rewrite]: #2715
+#2719 := [monotonicity #2716]: #2718
+#2723 := [trans #2719 #2721]: #2722
+#2726 := [monotonicity #2723]: #2725
+#2731 := [trans #2726 #2729]: #2730
+#2734 := [quant-intro #2731]: #2733
+#2711 := (iff #1477 #2710)
+#2708 := (iff #1474 #2705)
+#2691 := (or #117 #2690)
+#2702 := (or #2691 #1109)
+#2706 := (iff #2702 #2705)
+#2707 := [rewrite]: #2706
+#2703 := (iff #1474 #2702)
+#2700 := (iff #391 #2691)
+#2692 := (not #2691)
+#2695 := (not #2692)
+#2698 := (iff #2695 #2691)
+#2699 := [rewrite]: #2698
+#2696 := (iff #391 #2695)
+#2693 := (iff #121 #2692)
+#2694 := [rewrite]: #2693
+#2697 := [monotonicity #2694]: #2696
+#2701 := [trans #2697 #2699]: #2700
+#2704 := [monotonicity #2701]: #2703
+#2709 := [trans #2704 #2707]: #2708
+#2712 := [quant-intro #2709]: #2711
+#2688 := (iff #2322 #2687)
+#2685 := (iff #2319 #2682)
+#2662 := (or #66 #1055)
+#2679 := (or #2662 #2676)
+#2683 := (iff #2679 #2682)
+#2684 := [rewrite]: #2683
+#2680 := (iff #2319 #2679)
+#2677 := (iff #2316 #2676)
+#2678 := [rewrite]: #2677
+#2671 := (iff #1062 #2662)
+#2663 := (not #2662)
+#2666 := (not #2663)
+#2669 := (iff #2666 #2662)
+#2670 := [rewrite]: #2669
+#2667 := (iff #1062 #2666)
+#2664 := (iff #1059 #2663)
+#2665 := [rewrite]: #2664
+#2668 := [monotonicity #2665]: #2667
+#2672 := [trans #2668 #2670]: #2671
+#2681 := [monotonicity #2672 #2678]: #2680
+#2686 := [trans #2681 #2684]: #2685
+#2689 := [quant-intro #2686]: #2688
+#3142 := [monotonicity #2689 #2712 #2734 #2762 #3139]: #3141
+#3153 := [trans #3142 #3151]: #3152
+#2660 := (iff #2280 #2657)
+#2652 := (and #1704 #1709 #2649)
+#2658 := (iff #2652 #2657)
+#2659 := [rewrite]: #2658
+#2653 := (iff #2280 #2652)
+#2650 := (iff #1732 #2649)
+#2647 := (iff #1729 #2638)
+#2639 := (not #2638)
+#2642 := (not #2639)
+#2645 := (iff #2642 #2638)
+#2646 := [rewrite]: #2645
+#2643 := (iff #1729 #2642)
+#2640 := (iff #1720 #2639)
+#2641 := [rewrite]: #2640
+#2644 := [monotonicity #2641]: #2643
+#2648 := [trans #2644 #2646]: #2647
+#2651 := [quant-intro #2648]: #2650
+#2654 := [monotonicity #2651]: #2653
+#2661 := [trans #2654 #2659]: #2660
+#3156 := [monotonicity #2661 #3153]: #3155
+#2635 := (iff #1047 #2634)
+#2632 := (iff #1044 #2629)
+#2615 := (or #78 #1029)
+#2626 := (or #2615 #1040)
+#2630 := (iff #2626 #2629)
+#2631 := [rewrite]: #2630
+#2627 := (iff #1044 #2626)
+#2624 := (iff #1036 #2615)
+#2616 := (not #2615)
+#2619 := (not #2616)
+#2622 := (iff #2619 #2615)
+#2623 := [rewrite]: #2622
+#2620 := (iff #1036 #2619)
+#2617 := (iff #1033 #2616)
+#2618 := [rewrite]: #2617
+#2621 := [monotonicity #2618]: #2620
+#2625 := [trans #2621 #2623]: #2624
+#2628 := [monotonicity #2625]: #2627
+#2633 := [trans #2628 #2631]: #2632
+#2636 := [quant-intro #2633]: #2635
+#3159 := [monotonicity #2636 #3156]: #3158
+#3167 := [trans #3159 #3165]: #3166
+#2613 := (iff #2272 #2612)
+#2610 := (iff #2269 #2607)
+#2593 := (or #2592 #1679)
+#2604 := (or #2593 #2266)
+#2608 := (iff #2604 #2607)
+#2609 := [rewrite]: #2608
+#2605 := (iff #2269 #2604)
+#2602 := (iff #1682 #2593)
+#2594 := (not #2593)
+#2597 := (not #2594)
+#2600 := (iff #2597 #2593)
+#2601 := [rewrite]: #2600
+#2598 := (iff #1682 #2597)
+#2595 := (iff #1681 #2594)
+#2596 := [rewrite]: #2595
+#2599 := [monotonicity #2596]: #2598
+#2603 := [trans #2599 #2601]: #2602
+#2606 := [monotonicity #2603]: #2605
+#2611 := [trans #2606 #2609]: #2610
+#2614 := [monotonicity #2611]: #2613
+#3170 := [monotonicity #2614 #3167]: #3169
+#2590 := (iff #1021 #2589)
+#2587 := (iff #1018 #2584)
+#1661 := (or #77 #1760)
+#2581 := (or #1661 #1014)
+#2585 := (iff #2581 #2584)
+#2586 := [rewrite]: #2585
+#2582 := (iff #1018 #2581)
+#2579 := (iff #370 #1661)
+#1662 := (not #1661)
+#2574 := (not #1662)
+#2577 := (iff #2574 #1661)
+#2578 := [rewrite]: #2577
+#2575 := (iff #370 #2574)
+#2259 := (iff #86 #1662)
+#2573 := [rewrite]: #2259
+#2576 := [monotonicity #2573]: #2575
+#2580 := [trans #2576 #2578]: #2579
+#2583 := [monotonicity #2580]: #2582
+#2588 := [trans #2583 #2586]: #2587
+#2591 := [quant-intro #2588]: #2590
+#3173 := [monotonicity #2591 #3170]: #3172
+#3181 := [trans #3173 #3179]: #3180
+#1635 := (iff #1658 #1634)
+#1973 := (iff #1657 #2083)
+#2036 := (or #1646 #2035)
+#2114 := (or #2036 #1656)
+#2015 := (iff #2114 #2083)
+#2016 := [rewrite]: #2015
+#2115 := (iff #1657 #2114)
+#1811 := (iff #1651 #2036)
+#1893 := (not #2036)
+#1694 := (not #1893)
+#1929 := (iff #1694 #2036)
+#1930 := [rewrite]: #1929
+#2055 := (iff #1651 #1694)
+#1894 := (iff #1650 #1893)
+#1693 := [rewrite]: #1894
+#2056 := [monotonicity #1693]: #2055
+#1812 := [trans #2056 #1930]: #1811
+#2082 := [monotonicity #1812]: #2115
+#1974 := [trans #2082 #2016]: #1973
+#1759 := [monotonicity #1974]: #1635
+#3184 := [monotonicity #1759 #3181]: #3183
+#3187 := [monotonicity #3184]: #3186
+#3194 := [trans #3187 #3192]: #3193
+#3197 := [monotonicity #3194]: #3196
+#3200 := [monotonicity #3197]: #3199
+#3207 := [trans #3200 #3205]: #3206
+#3210 := [monotonicity #3207]: #3209
+#2138 := (+ #2137 #2133)
+#2139 := (= #2138 0::Int)
+#2140 := (and #2135 #218 #2139)
+#2150 := (not #2140)
+#2153 := (forall (vars (?v1 S2)) #2150)
+#2131 := (and #2125 #2130)
+#2132 := (not #2131)
+#2147 := (not #2132)
+#2157 := (and #2147 #2153)
+#2162 := (and #1301 #2157)
+#2166 := (or #2111 #2162)
+#2170 := (and #1282 #2166)
+#2174 := (or #2079 #2170)
+#2178 := (and #1268 #2174)
+#2182 := (or #2052 #2178)
+#2046 := (not #713)
+#2186 := (and #2046 #2182)
+#2190 := (or #713 #2186)
+#2194 := (and #606 #2190)
+#2198 := (or #2032 #2194)
+#2202 := (and #1259 #2198)
+#2206 := (or #2012 #2202)
+#1989 := (not #778)
+#1970 := (and #1964 #1969)
+#2210 := (and #1970 #189 #1412 #1404 #1989 #1391 #1248 #2206)
+#1922 := (+ #1906 #1921)
+#1923 := (+ #1913 #1922)
+#1924 := (>= #1923 0::Int)
+#1925 := (or #1919 #1924)
+#1926 := (not #1925)
+#1945 := (or #1926 #1941)
+#1882 := (+ #1881 #1136)
+#1888 := (+ #1887 #1882)
+#1889 := (= #1888 0::Int)
+#1883 := (>= #1882 0::Int)
+#1884 := (not #1883)
+#1890 := (and #1884 #1889)
+#1895 := (or #1145 #1890)
+#1898 := (forall (vars (?v0 S2)) #1895)
+#1949 := (and #1898 #1945)
+#1855 := (+ #1854 #1850)
+#1856 := (= #1855 0::Int)
+#1857 := (and #1852 #1856)
+#1866 := (not #1857)
+#1869 := (forall (vars (?v1 S2)) #1866)
+#1848 := (and #1842 #1847)
+#1849 := (not #1848)
+#1863 := (not #1849)
+#1873 := (and #1863 #1869)
+#1953 := (or #1873 #1949)
+#1837 := (not #522)
+#1834 := (not #531)
+#1831 := (not #540)
+#1828 := (not #549)
+#1957 := (and #1825 #1828 #1831 #1834 #1837 #1953)
+#2214 := (or #1957 #2210)
+#1798 := (+ #1797 #1091)
+#1806 := (+ #1805 #1798)
+#1807 := (= #1806 0::Int)
+#1799 := (>= #1798 0::Int)
+#1800 := (not #1799)
+#1808 := (and #1800 #1802 #1807)
+#1813 := (or #1100 #1808)
+#1816 := (forall (vars (?v0 S2)) #1813)
+#1770 := (not #878)
+#1753 := (+ #1053 #1752)
+#1754 := (+ #1744 #1753)
+#1755 := (= #1754 0::Int)
+#1745 := (+ #1744 #1053)
+#1746 := (>= #1745 0::Int)
+#1747 := (not #1746)
+#1756 := (and #1747 #1749 #1755)
+#1761 := (or #1062 #1756)
+#1764 := (forall (vars (?v0 S2)) #1761)
+#2218 := (and #1764 #1770 #1486 #1477 #1466 #1816 #2214)
+#1710 := (and #1704 #1709)
+#1711 := (not #1710)
+#1726 := (not #1711)
+#1736 := (and #1726 #1732)
+#2222 := (or #1736 #2218)
+#2226 := (and #1047 #2222)
+#1686 := (+ #1685 #1676)
+#1687 := (+ #1683 #1686)
+#1688 := (>= #1687 0::Int)
+#1689 := (or #1682 #1688)
+#1690 := (not #1689)
+#2230 := (or #1690 #2226)
+#2234 := (and #1021 #2230)
+#2238 := (or #1658 #2234)
+#2242 := (and #1006 #2238)
+#2246 := (or #1631 #2242)
+#1593 := (not #946)
+#2250 := (and #1593 #2246)
+#2254 := (or #946 #2250)
+#2571 := (iff #2254 #2570)
+#2568 := (iff #2250 #2567)
+#2565 := (iff #2246 #2564)
+#2562 := (iff #2242 #2561)
+#2559 := (iff #2238 #2558)
+#2556 := (iff #2234 #2555)
+#2553 := (iff #2230 #2552)
+#2550 := (iff #2226 #2549)
+#2547 := (iff #2222 #2546)
+#2544 := (iff #2218 #2543)
+#2541 := (iff #2214 #2540)
+#2538 := (iff #2210 #2535)
+#2532 := (and #1970 #189 #1412 #1404 #199 #1391 #1248 #2529)
+#2536 := (iff #2532 #2535)
+#2537 := [rewrite]: #2536
+#2533 := (iff #2210 #2532)
+#2530 := (iff #2206 #2529)
+#2527 := (iff #2202 #2526)
+#2524 := (iff #2198 #2523)
+#2521 := (iff #2194 #2520)
+#2518 := (iff #2190 #2517)
+#2515 := (iff #2186 #2514)
+#2512 := (iff #2182 #2511)
+#2509 := (iff #2178 #2508)
+#2506 := (iff #2174 #2505)
+#2503 := (iff #2170 #2502)
+#2500 := (iff #2166 #2499)
+#2497 := (iff #2162 #2494)
+#2488 := (and #2131 #2485)
+#2491 := (and #1301 #2488)
+#2495 := (iff #2491 #2494)
+#2496 := [rewrite]: #2495
+#2492 := (iff #2162 #2491)
+#2489 := (iff #2157 #2488)
+#2486 := (iff #2153 #2485)
+#2483 := (iff #2150 #2482)
+#2480 := (iff #2140 #2479)
+#2477 := (iff #2139 #2476)
+#2474 := (= #2138 #2473)
+#2475 := [rewrite]: #2474
+#2478 := [monotonicity #2475]: #2477
+#2481 := [monotonicity #2478]: #2480
+#2484 := [monotonicity #2481]: #2483
+#2487 := [quant-intro #2484]: #2486
+#2470 := (iff #2147 #2131)
+#2471 := [rewrite]: #2470
+#2490 := [monotonicity #2471 #2487]: #2489
+#2493 := [monotonicity #2490]: #2492
+#2498 := [trans #2493 #2496]: #2497
+#2501 := [monotonicity #2498]: #2500
+#2504 := [monotonicity #2501]: #2503
+#2507 := [monotonicity #2504]: #2506
+#2510 := [monotonicity #2507]: #2509
+#2513 := [monotonicity #2510]: #2512
+#2468 := (iff #2046 #222)
+#2469 := [rewrite]: #2468
+#2516 := [monotonicity #2469 #2513]: #2515
+#2519 := [monotonicity #2516]: #2518
+#2522 := [monotonicity #2519]: #2521
+#2525 := [monotonicity #2522]: #2524
+#2528 := [monotonicity #2525]: #2527
+#2531 := [monotonicity #2528]: #2530
+#2466 := (iff #1989 #199)
+#2467 := [rewrite]: #2466
+#2534 := [monotonicity #2467 #2531]: #2533
+#2539 := [trans #2534 #2537]: #2538
+#2464 := (iff #1957 #2463)
+#2461 := (iff #1953 #2460)
+#2458 := (iff #1949 #2457)
+#2455 := (iff #1945 #2454)
+#2452 := (iff #1926 #2451)
+#2449 := (iff #1925 #2448)
+#2446 := (iff #1924 #2445)
+#2443 := (= #1923 #2442)
+#2444 := [rewrite]: #2443
+#2447 := [monotonicity #2444]: #2446
+#2450 := [monotonicity #2447]: #2449
+#2453 := [monotonicity #2450]: #2452
+#2456 := [monotonicity #2453]: #2455
+#2439 := (iff #1898 #2438)
+#2436 := (iff #1895 #2435)
+#2433 := (iff #1890 #2432)
+#2430 := (iff #1889 #2427)
+#2417 := (+ #1881 #1887)
+#2418 := (+ #1136 #2417)
+#2421 := (= #2418 0::Int)
+#2428 := (iff #2421 #2427)
+#2429 := [rewrite]: #2428
+#2422 := (iff #1889 #2421)
+#2419 := (= #1888 #2418)
+#2420 := [rewrite]: #2419
+#2423 := [monotonicity #2420]: #2422
+#2431 := [trans #2423 #2429]: #2430
+#2415 := (iff #1884 #2414)
+#2412 := (iff #1883 #2409)
+#2401 := (+ #1136 #1881)
+#2404 := (>= #2401 0::Int)
+#2410 := (iff #2404 #2409)
+#2411 := [rewrite]: #2410
+#2405 := (iff #1883 #2404)
+#2402 := (= #1882 #2401)
+#2403 := [rewrite]: #2402
+#2406 := [monotonicity #2403]: #2405
+#2413 := [trans #2406 #2411]: #2412
+#2416 := [monotonicity #2413]: #2415
+#2434 := [monotonicity #2416 #2431]: #2433
+#2437 := [monotonicity #2434]: #2436
+#2440 := [quant-intro #2437]: #2439
+#2459 := [monotonicity #2440 #2456]: #2458
+#2399 := (iff #1873 #2396)
+#2393 := (and #1848 #2390)
+#2397 := (iff #2393 #2396)
+#2398 := [rewrite]: #2397
+#2394 := (iff #1873 #2393)
+#2391 := (iff #1869 #2390)
+#2388 := (iff #1866 #2387)
+#2385 := (iff #1857 #2384)
+#2382 := (iff #1856 #2381)
+#2379 := (= #1855 #2378)
+#2380 := [rewrite]: #2379
+#2383 := [monotonicity #2380]: #2382
+#2386 := [monotonicity #2383]: #2385
+#2389 := [monotonicity #2386]: #2388
+#2392 := [quant-intro #2389]: #2391
+#2375 := (iff #1863 #1848)
+#2376 := [rewrite]: #2375
+#2395 := [monotonicity #2376 #2392]: #2394
+#2400 := [trans #2395 #2398]: #2399
+#2462 := [monotonicity #2400 #2459]: #2461
+#2373 := (iff #1837 #152)
+#2374 := [rewrite]: #2373
+#2371 := (iff #1834 #149)
+#2372 := [rewrite]: #2371
+#2369 := (iff #1831 #147)
+#2370 := [rewrite]: #2369
+#2367 := (iff #1828 #144)
+#2368 := [rewrite]: #2367
+#2465 := [monotonicity #2368 #2370 #2372 #2374 #2462]: #2464
+#2542 := [monotonicity #2465 #2539]: #2541
+#2365 := (iff #1816 #2364)
+#2362 := (iff #1813 #2361)
+#2359 := (iff #1808 #2358)
+#2356 := (iff #1807 #2353)
+#2343 := (+ #1797 #1805)
+#2344 := (+ #1091 #2343)
+#2347 := (= #2344 0::Int)
+#2354 := (iff #2347 #2353)
+#2355 := [rewrite]: #2354
+#2348 := (iff #1807 #2347)
+#2345 := (= #1806 #2344)
+#2346 := [rewrite]: #2345
+#2349 := [monotonicity #2346]: #2348
+#2357 := [trans #2349 #2355]: #2356
+#2341 := (iff #1800 #2340)
+#2338 := (iff #1799 #2335)
+#2327 := (+ #1091 #1797)
+#2330 := (>= #2327 0::Int)
+#2336 := (iff #2330 #2335)
+#2337 := [rewrite]: #2336
+#2331 := (iff #1799 #2330)
+#2328 := (= #1798 #2327)
+#2329 := [rewrite]: #2328
+#2332 := [monotonicity #2329]: #2331
+#2339 := [trans #2332 #2337]: #2338
+#2342 := [monotonicity #2339]: #2341
+#2360 := [monotonicity #2342 #2357]: #2359
+#2363 := [monotonicity #2360]: #2362
+#2366 := [quant-intro #2363]: #2365
+#2325 := (iff #1770 #111)
+#2326 := [rewrite]: #2325
+#2323 := (iff #1764 #2322)
+#2320 := (iff #1761 #2319)
+#2317 := (iff #1756 #2316)
+#2314 := (iff #1755 #2311)
+#2301 := (+ #1744 #1752)
+#2302 := (+ #1053 #2301)
+#2305 := (= #2302 0::Int)
+#2312 := (iff #2305 #2311)
+#2313 := [rewrite]: #2312
+#2306 := (iff #1755 #2305)
+#2303 := (= #1754 #2302)
+#2304 := [rewrite]: #2303
+#2307 := [monotonicity #2304]: #2306
+#2315 := [trans #2307 #2313]: #2314
+#2299 := (iff #1747 #2298)
+#2296 := (iff #1746 #2293)
+#2285 := (+ #1053 #1744)
+#2288 := (>= #2285 0::Int)
+#2294 := (iff #2288 #2293)
+#2295 := [rewrite]: #2294
+#2289 := (iff #1746 #2288)
+#2286 := (= #1745 #2285)
+#2287 := [rewrite]: #2286
+#2290 := [monotonicity #2287]: #2289
+#2297 := [trans #2290 #2295]: #2296
+#2300 := [monotonicity #2297]: #2299
+#2318 := [monotonicity #2300 #2315]: #2317
+#2321 := [monotonicity #2318]: #2320
+#2324 := [quant-intro #2321]: #2323
+#2545 := [monotonicity #2324 #2326 #2366 #2542]: #2544
+#2283 := (iff #1736 #2280)
+#2277 := (and #1710 #1732)
+#2281 := (iff #2277 #2280)
+#2282 := [rewrite]: #2281
+#2278 := (iff #1736 #2277)
+#2275 := (iff #1726 #1710)
+#2276 := [rewrite]: #2275
+#2279 := [monotonicity #2276]: #2278
+#2284 := [trans #2279 #2282]: #2283
+#2548 := [monotonicity #2284 #2545]: #2547
+#2551 := [monotonicity #2548]: #2550
+#2273 := (iff #1690 #2272)
+#2270 := (iff #1689 #2269)
+#2267 := (iff #1688 #2266)
+#2264 := (= #1687 #2263)
+#2265 := [rewrite]: #2264
+#2268 := [monotonicity #2265]: #2267
+#2271 := [monotonicity #2268]: #2270
+#2274 := [monotonicity #2271]: #2273
+#2554 := [monotonicity #2274 #2551]: #2553
+#2557 := [monotonicity #2554]: #2556
+#2560 := [monotonicity #2557]: #2559
+#2563 := [monotonicity #2560]: #2562
+#2566 := [monotonicity #2563]: #2565
+#2260 := (iff #1593 #81)
+#2261 := [rewrite]: #2260
+#2569 := [monotonicity #2261 #2566]: #2568
+#2572 := [monotonicity #2569]: #2571
+#1568 := (not #1542)
+#2255 := (~ #1568 #2254)
+#2251 := (not #1539)
+#2252 := (~ #2251 #2250)
+#2247 := (not #1536)
+#2248 := (~ #2247 #2246)
+#2243 := (not #1533)
+#2244 := (~ #2243 #2242)
+#2239 := (not #1530)
+#2240 := (~ #2239 #2238)
+#2235 := (not #1527)
+#2236 := (~ #2235 #2234)
+#2231 := (not #1524)
+#2232 := (~ #2231 #2230)
+#2227 := (not #1521)
+#2228 := (~ #2227 #2226)
+#2223 := (not #1518)
+#2224 := (~ #2223 #2222)
+#2219 := (not #1513)
+#2220 := (~ #2219 #2218)
+#2215 := (not #1451)
+#2216 := (~ #2215 #2214)
+#2211 := (not #1446)
+#2212 := (~ #2211 #2210)
+#2207 := (not #1377)
+#2208 := (~ #2207 #2206)
+#2203 := (not #1374)
+#2204 := (~ #2203 #2202)
+#2199 := (not #1371)
+#2200 := (~ #2199 #2198)
+#2195 := (not #1368)
+#2196 := (~ #2195 #2194)
+#2191 := (not #1365)
+#2192 := (~ #2191 #2190)
+#2187 := (not #1362)
+#2188 := (~ #2187 #2186)
+#2183 := (not #1359)
+#2184 := (~ #2183 #2182)
+#2179 := (not #1356)
+#2180 := (~ #2179 #2178)
+#2175 := (not #1353)
+#2176 := (~ #2175 #2174)
+#2171 := (not #1350)
+#2172 := (~ #2171 #2170)
+#2167 := (not #1347)
+#2168 := (~ #2167 #2166)
+#2163 := (not #1344)
+#2164 := (~ #2163 #2162)
+#2144 := (not #1341)
+#2160 := (~ #2144 #2157)
+#2141 := (exists (vars (?v1 S2)) #2140)
+#2142 := (or #2132 #2141)
+#2143 := (not #2142)
+#2158 := (~ #2143 #2157)
+#2154 := (not #2141)
+#2155 := (~ #2154 #2153)
+#2151 := (~ #2150 #2150)
+#2152 := [refl]: #2151
+#2156 := [nnf-neg #2152]: #2155
+#2148 := (~ #2147 #2147)
+#2149 := [refl]: #2148
+#2159 := [nnf-neg #2149 #2156]: #2158
+#2145 := (~ #2144 #2143)
+#2146 := [sk]: #2145
+#2161 := [trans #2146 #2159]: #2160
+#2120 := (not #1304)
+#2121 := (~ #2120 #1301)
+#2118 := (~ #1301 #1301)
+#2116 := (~ #1298 #1298)
+#2117 := [refl]: #2116
+#2119 := [nnf-pos #2117]: #2118
+#2122 := [nnf-neg #2119]: #2121
+#2165 := [nnf-neg #2122 #2161]: #2164
+#2112 := (~ #1304 #2111)
+#2113 := [sk]: #2112
+#2169 := [nnf-neg #2113 #2165]: #2168
+#2088 := (not #1285)
+#2089 := (~ #2088 #1282)
+#2086 := (~ #1282 #1282)
+#2084 := (~ #1279 #1279)
+#2085 := [refl]: #2084
+#2087 := [nnf-pos #2085]: #2086
+#2090 := [nnf-neg #2087]: #2089
+#2173 := [nnf-neg #2090 #2169]: #2172
+#2080 := (~ #1285 #2079)
+#2081 := [sk]: #2080
+#2177 := [nnf-neg #2081 #2173]: #2176
+#2061 := (not #1271)
+#2062 := (~ #2061 #1268)
+#2059 := (~ #1268 #1268)
+#2057 := (~ #1265 #1265)
+#2058 := [refl]: #2057
+#2060 := [nnf-pos #2058]: #2059
+#2063 := [nnf-neg #2060]: #2062
+#2181 := [nnf-neg #2063 #2177]: #2180
+#2053 := (~ #1271 #2052)
+#2054 := [sk]: #2053
+#2185 := [nnf-neg #2054 #2181]: #2184
+#2047 := (~ #2046 #2046)
+#2048 := [refl]: #2047
+#2189 := [nnf-neg #2048 #2185]: #2188
+#2044 := (~ #713 #713)
+#2045 := [refl]: #2044
+#2193 := [nnf-neg #2045 #2189]: #2192
+#2041 := (not #736)
+#2042 := (~ #2041 #606)
+#2039 := (~ #606 #606)
+#2037 := (~ #603 #603)
+#2038 := [refl]: #2037
+#2040 := [nnf-pos #2038]: #2039
+#2043 := [nnf-neg #2040]: #2042
+#2197 := [nnf-neg #2043 #2193]: #2196
+#2033 := (~ #736 #2032)
+#2034 := [sk]: #2033
+#2201 := [nnf-neg #2034 #2197]: #2200
+#2021 := (not #1262)
+#2022 := (~ #2021 #1259)
+#2019 := (~ #1259 #1259)
+#2017 := (~ #1254 #1254)
+#2018 := [refl]: #2017
+#2020 := [nnf-pos #2018]: #2019
+#2023 := [nnf-neg #2020]: #2022
+#2205 := [nnf-neg #2023 #2201]: #2204
+#2013 := (~ #1262 #2012)
+#2014 := [sk]: #2013
+#2209 := [nnf-neg #2014 #2205]: #2208
+#2003 := (not #1251)
+#2004 := (~ #2003 #1248)
+#2001 := (~ #1248 #1248)
+#1999 := (~ #1245 #1245)
+#2000 := [refl]: #1999
+#2002 := [nnf-pos #2000]: #2001
+#2005 := [nnf-neg #2002]: #2004
+#1996 := (not #1394)
+#1997 := (~ #1996 #1391)
+#1994 := (~ #1391 #1391)
+#1992 := (~ #1388 #1388)
+#1993 := [refl]: #1992
+#1995 := [nnf-pos #1993]: #1994
+#1998 := [nnf-neg #1995]: #1997
+#1990 := (~ #1989 #1989)
+#1991 := [refl]: #1990
+#1986 := (not #1407)
+#1987 := (~ #1986 #1404)
+#1984 := (~ #1404 #1404)
+#1982 := (~ #1401 #1401)
+#1983 := [refl]: #1982
+#1985 := [nnf-pos #1983]: #1984
+#1988 := [nnf-neg #1985]: #1987
+#1980 := (~ #1412 #1412)
+#1981 := [refl]: #1980
+#1978 := (~ #189 #189)
+#1979 := [refl]: #1978
+#1975 := (not #1422)
+#1976 := (~ #1975 #1970)
+#1971 := (~ #1206 #1970)
+#1972 := [sk]: #1971
+#1977 := [nnf-neg #1972]: #1976
+#2213 := [nnf-neg #1977 #1979 #1981 #1988 #1991 #1998 #2005 #2209]: #2212
+#1958 := (not #1224)
+#1959 := (~ #1958 #1957)
+#1954 := (not #1200)
+#1955 := (~ #1954 #1953)
+#1950 := (not #1197)
+#1951 := (~ #1950 #1949)
+#1946 := (not #1194)
+#1947 := (~ #1946 #1945)
+#1942 := (not #1191)
+#1943 := (~ #1942 #1941)
+#1939 := (~ #1938 #1938)
+#1940 := [refl]: #1939
+#1935 := (not #1188)
+#1936 := (~ #1935 #1185)
+#1933 := (~ #1185 #1185)
+#1931 := (~ #1182 #1182)
+#1932 := [refl]: #1931
+#1934 := [nnf-pos #1932]: #1933
+#1937 := [nnf-neg #1934]: #1936
+#1944 := [nnf-neg #1937 #1940]: #1943
+#1927 := (~ #1188 #1926)
+#1928 := [sk]: #1927
+#1948 := [nnf-neg #1928 #1944]: #1947
+#1901 := (not #1170)
+#1902 := (~ #1901 #1898)
+#1899 := (~ #1167 #1898)
+#1896 := (~ #1164 #1895)
+#1891 := (~ #1161 #1890)
+#1892 := [sk]: #1891
+#1878 := (~ #1145 #1145)
+#1879 := [refl]: #1878
+#1897 := [monotonicity #1879 #1892]: #1896
+#1900 := [nnf-pos #1897]: #1899
+#1903 := [nnf-neg #1900]: #1902
+#1952 := [nnf-neg #1903 #1948]: #1951
+#1876 := (~ #1170 #1873)
+#1858 := (exists (vars (?v1 S2)) #1857)
+#1859 := (or #1849 #1858)
+#1860 := (not #1859)
+#1874 := (~ #1860 #1873)
+#1870 := (not #1858)
+#1871 := (~ #1870 #1869)
+#1867 := (~ #1866 #1866)
+#1868 := [refl]: #1867
+#1872 := [nnf-neg #1868]: #1871
+#1864 := (~ #1863 #1863)
+#1865 := [refl]: #1864
+#1875 := [nnf-neg #1865 #1872]: #1874
+#1861 := (~ #1170 #1860)
+#1862 := [sk]: #1861
+#1877 := [trans #1862 #1875]: #1876
+#1956 := [nnf-neg #1877 #1952]: #1955
+#1838 := (~ #1837 #1837)
+#1839 := [refl]: #1838
+#1835 := (~ #1834 #1834)
+#1836 := [refl]: #1835
+#1832 := (~ #1831 #1831)
+#1833 := [refl]: #1832
+#1829 := (~ #1828 #1828)
+#1830 := [refl]: #1829
+#1826 := (~ #1422 #1825)
+#1823 := (~ #1822 #1822)
+#1824 := [refl]: #1823
+#1827 := [nnf-neg #1824]: #1826
+#1960 := [nnf-neg #1827 #1830 #1833 #1836 #1839 #1956]: #1959
+#2217 := [nnf-neg #1960 #2213]: #2216
+#1819 := (not #1133)
+#1820 := (~ #1819 #1816)
+#1817 := (~ #1130 #1816)
+#1814 := (~ #1127 #1813)
+#1809 := (~ #1124 #1808)
+#1810 := [sk]: #1809
+#1794 := (~ #1100 #1100)
+#1795 := [refl]: #1794
+#1815 := [monotonicity #1795 #1810]: #1814
+#1818 := [nnf-pos #1815]: #1817
+#1821 := [nnf-neg #1818]: #1820
+#1791 := (not #1469)
+#1792 := (~ #1791 #1466)
+#1789 := (~ #1466 #1466)
+#1787 := (~ #1463 #1463)
+#1788 := [refl]: #1787
+#1790 := [nnf-pos #1788]: #1789
+#1793 := [nnf-neg #1790]: #1792
+#1784 := (not #1480)
+#1785 := (~ #1784 #1477)
+#1782 := (~ #1477 #1477)
+#1780 := (~ #1474 #1474)
+#1781 := [refl]: #1780
+#1783 := [nnf-pos #1781]: #1782
+#1786 := [nnf-neg #1783]: #1785
+#1777 := (not #1489)
+#1778 := (~ #1777 #1486)
+#1775 := (~ #1486 #1486)
+#1773 := (~ #1483 #1483)
+#1774 := [refl]: #1773
+#1776 := [nnf-pos #1774]: #1775
+#1779 := [nnf-neg #1776]: #1778
+#1771 := (~ #1770 #1770)
+#1772 := [refl]: #1771
+#1767 := (not #1492)
+#1768 := (~ #1767 #1764)
+#1765 := (~ #1088 #1764)
+#1762 := (~ #1085 #1761)
+#1757 := (~ #1082 #1756)
+#1758 := [sk]: #1757
+#1741 := (~ #1062 #1062)
+#1742 := [refl]: #1741
+#1763 := [monotonicity #1742 #1758]: #1762
+#1766 := [nnf-pos #1763]: #1765
+#1769 := [nnf-neg #1766]: #1768
+#2221 := [nnf-neg #1769 #1772 #1779 #1786 #1793 #1821 #2217]: #2220
+#1739 := (~ #1492 #1736)
+#1721 := (exists (vars (?v1 S2)) #1720)
+#1722 := (or #1711 #1721)
+#1723 := (not #1722)
+#1737 := (~ #1723 #1736)
+#1733 := (not #1721)
+#1734 := (~ #1733 #1732)
+#1730 := (~ #1729 #1729)
+#1731 := [refl]: #1730
+#1735 := [nnf-neg #1731]: #1734
+#1727 := (~ #1726 #1726)
+#1728 := [refl]: #1727
+#1738 := [nnf-neg #1728 #1735]: #1737
+#1724 := (~ #1492 #1723)
+#1725 := [sk]: #1724
+#1740 := [trans #1725 #1738]: #1739
+#2225 := [nnf-neg #1740 #2221]: #2224
+#1699 := (not #1050)
+#1700 := (~ #1699 #1047)
+#1697 := (~ #1047 #1047)
+#1695 := (~ #1044 #1044)
+#1696 := [refl]: #1695
+#1698 := [nnf-pos #1696]: #1697
+#1701 := [nnf-neg #1698]: #1700
+#2229 := [nnf-neg #1701 #2225]: #2228
+#1691 := (~ #1050 #1690)
+#1692 := [sk]: #1691
+#2233 := [nnf-neg #1692 #2229]: #2232
+#1667 := (not #1024)
+#1668 := (~ #1667 #1021)
+#1665 := (~ #1021 #1021)
+#1663 := (~ #1018 #1018)
+#1664 := [refl]: #1663
+#1666 := [nnf-pos #1664]: #1665
+#1669 := [nnf-neg #1666]: #1668
+#2237 := [nnf-neg #1669 #2233]: #2236
+#1659 := (~ #1024 #1658)
+#1660 := [sk]: #1659
+#2241 := [nnf-neg #1660 #2237]: #2240
+#1640 := (not #1009)
+#1641 := (~ #1640 #1006)
+#1638 := (~ #1006 #1006)
+#1636 := (~ #1005 #1005)
+#1637 := [refl]: #1636
+#1639 := [nnf-pos #1637]: #1638
+#1642 := [nnf-neg #1639]: #1641
+#2245 := [nnf-neg #1642 #2241]: #2244
+#1632 := (~ #1009 #1631)
+#1633 := [sk]: #1632
+#2249 := [nnf-neg #1633 #2245]: #2248
+#1594 := (~ #1593 #1593)
+#1627 := [refl]: #1594
+#2253 := [nnf-neg #1627 #2249]: #2252
+#1625 := (~ #946 #946)
+#1626 := [refl]: #1625
+#2256 := [nnf-neg #1626 #2253]: #2255
+#1569 := [not-or-elim #1564]: #1568
+#2257 := [mp~ #1569 #2256]: #2254
+#2258 := [mp #2257 #2572]: #2570
+#3211 := [mp #2258 #3210]: #3208
+#4126 := [mp #3211 #4125]: #4123
+#7087 := [unit-resolution #4126 #4239]: #4120
+#3450 := (or #4117 #4111)
+#3440 := [def-axiom]: #3450
+#7088 := [unit-resolution #3440 #7087]: #4111
+#3446 := (or #4114 #1631 #4108)
+#3448 := [def-axiom]: #3446
+#7089 := [unit-resolution #3448 #7088 #4166]: #4108
+#3444 := (or #4105 #4099)
+#3447 := [def-axiom]: #3444
+#7090 := [unit-resolution #3447 #7089]: #4099
+#3306 := (or #4102 #1634 #4096)
+#3464 := [def-axiom]: #3306
+#7091 := [unit-resolution #3464 #7090]: #4099
+#7092 := [unit-resolution #7091 #7086]: #4096
+#3486 := (or #4093 #4087)
+#3456 := [def-axiom]: #3486
+#7093 := [unit-resolution #3456 #7092]: #4087
+#7095 := (or #4090 #4084)
+#6151 := [hypothesis]: #1673
+#4285 := (or #7029 #2592)
+#4289 := [quant-inst #1670]: #4285
+#6152 := [unit-resolution #4289 #3755 #6151]: false
+#6170 := [lemma #6152]: #2592
+#3366 := (or #2607 #1673)
+#3363 := [def-axiom]: #3366
+#7094 := [unit-resolution #3363 #6170]: #2607
+#3483 := (or #4090 #2612 #4084)
+#3484 := [def-axiom]: #3483
+#7096 := [unit-resolution #3484 #7094]: #7095
+#7097 := [unit-resolution #7096 #7093]: #4084
+#3467 := (or #4081 #4075)
+#3474 := [def-axiom]: #3467
+#7098 := [unit-resolution #3474 #7097]: #4075
+#3504 := (or #4078 #3794 #4072)
+#3489 := [def-axiom]: #3504
+#7099 := [unit-resolution #3489 #7098 #6210]: #4072
+#3496 := (or #4069 #4063)
+#3497 := [def-axiom]: #3496
+#8263 := [unit-resolution #3497 #7099]: #4063
+#6420 := (f19 f20 ?v0!8)
+#6418 := (* -1::Int #6420)
+#6421 := (+ f14 #6418)
+#6440 := (<= #6421 0::Int)
+#6559 := (?v1!7 ?v0!8)
+#6669 := (f6 f7 #6559)
+#6677 := (f5 #6669 ?v0!8)
+#6678 := (f15 #6677)
+#6676 := (* -1::Int #6678)
+#6561 := (f19 f20 #6559)
+#6563 := (* -1::Int #6561)
+#6673 := (+ #6563 #6676)
+#6661 := (+ #6420 #6673)
+#6662 := (= #6661 0::Int)
+#6715 := (not #6662)
+#6565 := (f9 f21 #6559)
+#6571 := (= #6565 f1)
+#6660 := (not #6571)
+#6564 := (+ #6420 #6563)
+#6562 := (<= #6564 0::Int)
+#6716 := (or #6562 #6660 #6715)
+#7070 := [hypothesis]: #3906
+#3648 := (or #3903 #149)
+#3643 := [def-axiom]: #3648
+#7106 := [unit-resolution #3643 #7070]: #149
+#3491 := (or #3903 #3897)
+#3492 := [def-axiom]: #3491
+#7107 := [unit-resolution #3492 #7070]: #3897
+#3520 := (or #4069 #111)
+#3521 := [def-axiom]: #3520
+#7324 := [unit-resolution #3521 #7099]: #111
+#4279 := (or #531 #169 #878)
+#4208 := [hypothesis]: #111
+#4276 := (= #168 #110)
+#4275 := [hypothesis]: #149
+#4274 := [monotonicity #4275]: #4276
+#4277 := [trans #4274 #4208]: #169
+#4174 := [hypothesis]: #1938
+#4278 := [unit-resolution #4174 #4277]: false
+#4292 := [lemma #4278]: #4279
+#7108 := [unit-resolution #4292 #7106 #7324]: #169
+#3387 := (or #3879 #1938)
+#3388 := [def-axiom]: #3387
+#7066 := [unit-resolution #3388 #7108]: #3879
+#3644 := (or #3903 #3840)
+#3645 := [def-axiom]: #3644
+#7069 := [unit-resolution #3645 #7070]: #3840
+#7013 := (or #2843 #3845 #531)
+#6412 := (f19 f20 ?v0!11)
+#6414 := (* -1::Int #6412)
+#6787 := (+ #1920 #6414)
+#6788 := (<= #6787 0::Int)
+#6785 := (= #1920 #6412)
+#6862 := (= #6412 #1920)
+#6860 := (= f20 f25)
+#6861 := [symm #4275]: #6860
+#6895 := [monotonicity #6861]: #6862
+#6896 := [symm #6895]: #6785
+#6897 := (not #6785)
+#6898 := (or #6897 #6788)
+#6854 := [th-lemma arith triangle-eq]: #6898
+#6855 := [unit-resolution #6854 #6896]: #6788
+#6195 := (f19 f20 ?v1!10)
+#6193 := (* -1::Int #6195)
+#6285 := (+ #1906 #6193)
+#6781 := (>= #6285 0::Int)
+#6295 := (= #1906 #6195)
+#6853 := (= #6195 #1906)
+#6856 := [monotonicity #6861]: #6853
+#6857 := [symm #6856]: #6295
+#6852 := (not #6295)
+#4251 := (or #6852 #6781)
+#4280 := [th-lemma arith triangle-eq]: #4251
+#4281 := [unit-resolution #4280 #6857]: #6781
+#3675 := (not #2445)
+#4345 := [hypothesis]: #2848
+#3673 := (or #2843 #3675)
+#3676 := [def-axiom]: #3673
+#4346 := [unit-resolution #3676 #4345]: #3675
+#7082 := [hypothesis]: #3840
+#3314 := (or #2843 #1917)
+#3315 := [def-axiom]: #3314
+#4379 := [unit-resolution #3315 #4345]: #1917
+#6179 := (+ f14 #6193)
+#6184 := (<= #6179 0::Int)
+#7080 := (not #6184)
+#3672 := (or #2843 #1910)
+#3674 := [def-axiom]: #3672
+#4380 := [unit-resolution #3674 #4345]: #1910
+#7076 := (not #6781)
+#4409 := (or #7080 #1909 #7076)
+#4410 := [th-lemma arith assign-bounds -1 -1]: #4409
+#7011 := [unit-resolution #4410 #4380 #4281]: #7080
+#7075 := (not #6788)
+#7104 := (or #6184 #1916 #3845 #2445 #7076 #7075)
+#6667 := (+ #6195 #6414)
+#6670 := (+ #1913 #6667)
+#6694 := (>= #6670 0::Int)
+#7074 := (not #6694)
+#7071 := [hypothesis]: #6788
+#7072 := [hypothesis]: #6781
+#7073 := [hypothesis]: #3675
+#7077 := (or #7074 #7075 #2445 #7076)
+#7078 := [th-lemma arith assign-bounds -1 -1 1]: #7077
+#7079 := [unit-resolution #7078 #7073 #7072 #7071]: #7074
+#6164 := (f9 f21 ?v1!10)
+#4586 := (= #6164 f1)
+#7081 := [hypothesis]: #7080
+#6183 := (or #4586 #6184)
+#6211 := (or #3845 #4586 #6184)
+#6212 := (or #3845 #6183)
+#6286 := (iff #6212 #6211)
+#6287 := [rewrite]: #6286
+#6280 := [quant-inst #1904]: #6212
+#6288 := [mp #6280 #6287]: #6211
+#7083 := [unit-resolution #6288 #7082]: #6183
+#7084 := [unit-resolution #7083 #7081]: #4586
+#6476 := (not #4586)
+#7101 := (or #6476 #6694)
+#7085 := [hypothesis]: #1917
+#3488 := (or #4069 #3823)
+#3493 := [def-axiom]: #3488
+#7100 := [unit-resolution #3493 #7099]: #3823
+#6719 := (or #3828 #6476 #1916 #6694)
+#6695 := (or #6476 #1916 #6694)
+#6714 := (or #3828 #6695)
+#6721 := (iff #6714 #6719)
+#6722 := [rewrite]: #6721
+#6720 := [quant-inst #1905 #1904]: #6714
+#6723 := [mp #6720 #6722]: #6719
+#7102 := [unit-resolution #6723 #7100 #7085]: #7101
+#7103 := [unit-resolution #7102 #7084 #7079]: false
+#7105 := [lemma #7103]: #7104
+#7012 := [unit-resolution #7105 #7011 #4379 #7082 #4346 #4281 #6855]: false
+#7019 := [lemma #7012]: #7013
+#7109 := [unit-resolution #7019 #7069 #7106]: #2843
+#3660 := (or #3888 #2848 #3882)
+#3657 := [def-axiom]: #3660
+#7110 := [unit-resolution #3657 #7109 #7066]: #3888
+#3372 := (or #3891 #3885)
+#3373 := [def-axiom]: #3372
+#7111 := [unit-resolution #3373 #7110]: #3891
+#3651 := (or #3900 #3860 #3894)
+#3655 := [def-axiom]: #3651
+#7112 := [unit-resolution #3655 #7111 #7107]: #3860
+#3323 := (or #3857 #3849)
+#3664 := [def-axiom]: #3323
+#7113 := [unit-resolution #3664 #7112]: #3849
+#7512 := (or #6716 #3854 #531)
+#6821 := (f19 f25 #6559)
+#7034 := (* -1::Int #6821)
+#7035 := (+ #1843 #7034)
+#7036 := (<= #7035 0::Int)
+#7057 := (+ #6676 #7034)
+#7058 := (+ #1843 #7057)
+#7059 := (= #7058 0::Int)
+#7307 := (+ #6561 #7034)
+#7253 := (>= #7307 0::Int)
+#7306 := (= #6561 #6821)
+#7446 := (= #6821 #6561)
+#7447 := [monotonicity #4275]: #7446
+#7448 := [symm #7447]: #7306
+#7449 := (not #7306)
+#7450 := (or #7449 #7253)
+#7451 := [th-lemma arith triangle-eq]: #7450
+#7452 := [unit-resolution #7451 #7448]: #7253
+#6279 := (+ #1843 #6418)
+#6798 := (>= #6279 0::Int)
+#5095 := (= #1843 #6420)
+#7453 := (= #6420 #1843)
+#7438 := [monotonicity #6861]: #7453
+#7439 := [symm #7438]: #5095
+#7437 := (not #5095)
+#7440 := (or #7437 #6798)
+#7441 := [th-lemma arith triangle-eq]: #7440
+#7442 := [unit-resolution #7441 #7439]: #6798
+#6767 := (>= #6661 0::Int)
+#6490 := (not #6716)
+#7455 := [hypothesis]: #6490
+#6120 := (or #6716 #6662)
+#6113 := [def-axiom]: #6120
+#7456 := [unit-resolution #6113 #7455]: #6662
+#7476 := (or #6715 #6767)
+#7477 := [th-lemma arith triangle-eq]: #7476
+#7478 := [unit-resolution #7477 #7456]: #6767
+#7252 := (<= #7307 0::Int)
+#7479 := (or #7449 #7252)
+#7480 := [th-lemma arith triangle-eq]: #7479
+#7475 := [unit-resolution #7480 #7448]: #7252
+#6792 := (<= #6279 0::Int)
+#7481 := (or #7437 #6792)
+#7482 := [th-lemma arith triangle-eq]: #7481
+#7483 := [unit-resolution #7482 #7439]: #6792
+#6766 := (<= #6661 0::Int)
+#7484 := (or #6715 #6766)
+#7485 := [th-lemma arith triangle-eq]: #7484
+#7506 := [unit-resolution #7485 #7456]: #6766
+#7400 := (not #7253)
+#7405 := (not #6798)
+#7404 := (not #6767)
+#7553 := (not #7252)
+#7337 := (not #6792)
+#7552 := (not #6766)
+#7410 := (or #7059 #7552 #7337 #7553 #7404 #7405 #7400)
+#7550 := [hypothesis]: #7252
+#7330 := [hypothesis]: #6792
+#7551 := [hypothesis]: #6766
+#6858 := (not #7059)
+#7548 := [hypothesis]: #6858
+#7185 := (>= #7058 0::Int)
+#7401 := [hypothesis]: #7253
+#7402 := [hypothesis]: #6798
+#7403 := [hypothesis]: #6767
+#7406 := (or #7185 #7404 #7405 #7400)
+#7407 := [th-lemma arith assign-bounds -1 -1 -1]: #7406
+#7408 := [unit-resolution #7407 #7403 #7402 #7401]: #7185
+#7558 := (not #7185)
+#7562 := (or #7558 #7059 #7552 #7337 #7553)
+#7549 := [hypothesis]: #7185
+#7184 := (<= #7058 0::Int)
+#7554 := (or #7184 #7552 #7337 #7553)
+#7555 := [th-lemma arith assign-bounds -1 -1 -1]: #7554
+#7556 := [unit-resolution #7555 #7551 #7330 #7550]: #7184
+#7557 := (not #7184)
+#7559 := (or #7059 #7557 #7558)
+#7560 := [th-lemma arith triangle-eq]: #7559
+#7561 := [unit-resolution #7560 #7556 #7549 #7548]: false
+#7563 := [lemma #7561]: #7562
+#7409 := [unit-resolution #7563 #7408 #7548 #7551 #7330 #7550]: false
+#7445 := [lemma #7409]: #7410
+#7507 := [unit-resolution #7445 #7506 #7483 #7475 #7478 #7442 #7452]: #7059
+#4250 := (or #7036 #6858)
+#7508 := [hypothesis]: #3849
+#7148 := (or #3854 #7036 #6858)
+#6893 := (+ #1844 #6678)
+#6894 := (+ #6821 #6893)
+#6886 := (= #6894 0::Int)
+#6904 := (not #6886)
+#6822 := (+ #6821 #1844)
+#6278 := (>= #6822 0::Int)
+#6907 := (or #6278 #6904)
+#7149 := (or #3854 #6907)
+#7182 := (iff #7149 #7148)
+#7158 := (or #3854 #4250)
+#7180 := (iff #7158 #7148)
+#7181 := [rewrite]: #7180
+#7159 := (iff #7149 #7158)
+#7060 := (iff #6907 #4250)
+#6859 := (iff #6904 #6858)
+#7067 := (iff #6886 #7059)
+#7045 := (+ #6678 #6821)
+#7048 := (+ #1844 #7045)
+#7055 := (= #7048 0::Int)
+#7063 := (iff #7055 #7059)
+#7064 := [rewrite]: #7063
+#7056 := (iff #6886 #7055)
+#7049 := (= #6894 #7048)
+#7050 := [rewrite]: #7049
+#7054 := [monotonicity #7050]: #7056
+#7068 := [trans #7054 #7064]: #7067
+#4217 := [monotonicity #7068]: #6859
+#7046 := (iff #6278 #7036)
+#7021 := (+ #1844 #6821)
+#7026 := (>= #7021 0::Int)
+#7037 := (iff #7026 #7036)
+#7038 := [rewrite]: #7037
+#7033 := (iff #6278 #7026)
+#7022 := (= #6822 #7021)
+#7025 := [rewrite]: #7022
+#6959 := [monotonicity #7025]: #7033
+#7047 := [trans #6959 #7038]: #7046
+#7065 := [monotonicity #7047 #4217]: #7060
+#7179 := [monotonicity #7065]: #7159
+#7183 := [trans #7179 #7181]: #7182
+#7145 := [quant-inst #6559]: #7149
+#7178 := [mp #7145 #7183]: #7148
+#7509 := [unit-resolution #7178 #7508]: #4250
+#7510 := [unit-resolution #7509 #7507]: #7036
+#6768 := (not #6562)
+#6392 := (or #6716 #6768)
+#6778 := [def-axiom]: #6392
+#7505 := [unit-resolution #6778 #7455]: #6768
+#7511 := [th-lemma arith farkas -1 -1 -1 1 #7442 #7505 #7452 #7510]: false
+#7513 := [lemma #7511]: #7512
+#7151 := [unit-resolution #7513 #7113 #7106]: #6716
+#7153 := (or #6440 #6490)
+#3678 := (or #3857 #1842)
+#3343 := [def-axiom]: #3678
+#7152 := [unit-resolution #3343 #7112]: #1842
+#3494 := (or #4069 #3831)
+#3495 := [def-axiom]: #3494
+#7150 := [unit-resolution #3495 #7099]: #3831
+#6491 := (or #3836 #1841 #6440 #6490)
+#6489 := (or #1841 #6440 #6490)
+#6492 := (or #3836 #6489)
+#6718 := (iff #6492 #6491)
+#6381 := [rewrite]: #6718
+#6717 := [quant-inst #1840]: #6492
+#6724 := [mp #6717 #6381]: #6491
+#7154 := [unit-resolution #6724 #7150 #7152]: #7153
+#7155 := [unit-resolution #7154 #7151]: #6440
+#3338 := (or #3857 #1847)
+#3680 := [def-axiom]: #3338
+#7156 := [unit-resolution #3680 #7112]: #1847
+#7141 := [symm #7106]: #6860
+#7142 := [monotonicity #7141]: #7453
+#7140 := [symm #7142]: #5095
+#7143 := [unit-resolution #7441 #7140]: #6798
+#7144 := [th-lemma arith farkas -1 -1 1 #7143 #7156 #7155]: false
+#7216 := [lemma #7144]: #3903
+#3508 := (or #4066 #3906 #4060)
+#3510 := [def-axiom]: #3508
+#8264 := [unit-resolution #3510 #7216 #8263]: #4060
+#3548 := (or #4057 #199)
+#3553 := [def-axiom]: #3548
+#9701 := [unit-resolution #3553 #8264]: #199
+#9297 := [symm #9701]: #9268
+#16690 := [monotonicity #9297]: #16482
+#16476 := [monotonicity #16690]: #16582
+#16737 := [symm #16476]: #15124
+#15539 := [monotonicity #16737]: #15122
+#19098 := (not #5176)
+#15519 := [hypothesis]: #19098
+#5179 := (or #4630 #5176)
+#7694 := (f5 #200 ?v0!14)
+#7695 := (f15 #7694)
+#7647 := (* -1::Int #2029)
+#7713 := (+ #7647 #7695)
+#7714 := (+ #190 #7713)
+#7715 := (>= #7714 0::Int)
+#8867 := (not #7715)
+#7696 := (* -1::Int #7695)
+#7697 := (+ f14 #7696)
+#7698 := (<= #7697 0::Int)
+#7746 := (or #7698 #7715)
+#7749 := (not #7746)
+#3637 := (not #2030)
+#10219 := [hypothesis]: #2032
+#3631 := (or #2031 #3637)
+#3638 := [def-axiom]: #3631
+#10218 := [unit-resolution #3638 #10219]: #3637
+#7752 := (or #7749 #2030)
+#7911 := [hypothesis]: #7746
+#8011 := [hypothesis]: #3637
+#3534 := (or #4057 #3927)
+#3515 := [def-axiom]: #3534
+#8861 := [unit-resolution #3515 #8264]: #3927
+#7814 := (or #3932 #7749 #2030)
+#7699 := (+ #1235 #7696)
+#7700 := (+ #2029 #7699)
+#7701 := (<= #7700 0::Int)
+#7743 := (or #7698 #7701)
+#7744 := (not #7743)
+#7745 := (or #7744 #2030)
+#7816 := (or #3932 #7745)
+#7852 := (iff #7816 #7814)
+#7842 := (or #3932 #7752)
+#7846 := (iff #7842 #7814)
+#7850 := [rewrite]: #7846
+#7813 := (iff #7816 #7842)
+#7753 := (iff #7745 #7752)
+#7750 := (iff #7744 #7749)
+#7747 := (iff #7743 #7746)
+#7718 := (iff #7701 #7715)
+#7706 := (+ #2029 #7696)
+#7707 := (+ #1235 #7706)
+#7710 := (<= #7707 0::Int)
+#7716 := (iff #7710 #7715)
+#7717 := [rewrite]: #7716
+#7711 := (iff #7701 #7710)
+#7708 := (= #7700 #7707)
+#7709 := [rewrite]: #7708
+#7712 := [monotonicity #7709]: #7711
+#7719 := [trans #7712 #7717]: #7718
+#7748 := [monotonicity #7719]: #7747
+#7751 := [monotonicity #7748]: #7750
+#7754 := [monotonicity #7751]: #7753
+#7843 := [monotonicity #7754]: #7813
+#7853 := [trans #7843 #7850]: #7852
+#7817 := [quant-inst #2024]: #7816
+#7881 := [mp #7817 #7853]: #7814
+#7907 := [unit-resolution #7881 #8861 #8011 #7911]: false
+#7918 := [lemma #7907]: #7752
+#10226 := [unit-resolution #7918 #10218]: #7749
+#7767 := (or #7746 #8867)
+#7768 := [def-axiom]: #7767
+#10265 := [unit-resolution #7768 #10226]: #8867
+#7674 := (+ #190 #7647)
+#7981 := (>= #7674 0::Int)
+#7663 := (f9 f21 ?v0!14)
+#7664 := (= #7663 f1)
+#7818 := (= ?v0!14 f28)
+#7841 := (not #7818)
+#9402 := (or #7841 #2030)
+#8045 := (= #190 #2029)
+#8033 := (= #2029 #190)
+#8023 := [hypothesis]: #7818
+#9294 := [monotonicity #8023]: #8033
+#9295 := [symm #9294]: #8045
+#8124 := (= #2028 #190)
+#4167 := (f30 f28)
+#4220 := (= #4167 #190)
+#4171 := (f5 #200 f28)
+#4172 := (f15 #4171)
+#4190 := (>= #4172 0::Int)
+#4175 := (* -1::Int #4172)
+#4176 := (+ f14 #4175)
+#4177 := (<= #4176 0::Int)
+#4222 := (or #4177 #4190)
+#7990 := (= #4172 0::Int)
+#8751 := (not #7990)
+#8752 := [hypothesis]: #8751
+#10 := (f6 f7 #9)
+#12 := (f5 #10 #11)
+#3689 := (pattern #12)
+#57 := (f15 #12)
+#58 := (= #57 0::Int)
+#56 := (= #9 #11)
+#61 := (not #56)
+#325 := (or #61 #58)
+#3724 := (forall (vars (?v0 S2) (?v1 S2)) (:pat #3689) #325)
+#333 := (forall (vars (?v0 S2) (?v1 S2)) #325)
+#3727 := (iff #333 #3724)
+#3725 := (iff #325 #325)
+#3726 := [refl]: #3725
+#3728 := [quant-intro #3726]: #3727
+#1581 := (~ #333 #333)
+#1612 := (~ #325 #325)
+#1613 := [refl]: #1612
+#1582 := [nnf-pos #1613]: #1581
+#59 := (implies #56 #58)
+#60 := (forall (vars (?v0 S2) (?v1 S2)) #59)
+#334 := (iff #60 #333)
+#331 := (iff #59 #325)
+#332 := [rewrite]: #331
+#335 := [quant-intro #332]: #334
+#324 := [asserted]: #60
+#338 := [mp #324 #335]: #333
+#1583 := [mp~ #338 #1582]: #333
+#3729 := [mp #1583 #3728]: #3724
+#6738 := (not #3724)
+#8741 := (or #6738 #7990)
+#4492 := (= f28 f28)
+#7989 := (not #4492)
+#7997 := (or #7989 #7990)
+#8742 := (or #6738 #7997)
+#8744 := (iff #8742 #8741)
+#8746 := (iff #8741 #8741)
+#8747 := [rewrite]: #8746
+#8007 := (iff #7997 #7990)
+#8002 := (or false #7990)
+#8005 := (iff #8002 #7990)
+#8006 := [rewrite]: #8005
+#8003 := (iff #7997 #8002)
+#8000 := (iff #7989 false)
+#7998 := (iff #7989 #6991)
+#4495 := (iff #4492 true)
+#4496 := [rewrite]: #4495
+#7999 := [monotonicity #4496]: #7998
+#8001 := [trans #7999 #6995]: #8000
+#8004 := [monotonicity #8001]: #8003
+#8008 := [trans #8004 #8006]: #8007
+#8745 := [monotonicity #8008]: #8744
+#8748 := [trans #8745 #8747]: #8744
+#8743 := [quant-inst #186 #186]: #8742
+#8749 := [mp #8743 #8748]: #8741
+#8757 := [unit-resolution #8749 #3729 #8752]: false
+#8758 := [lemma #8757]: #7990
+#9347 := (or #8751 #4190)
+#9298 := [th-lemma arith triangle-eq]: #9347
+#8814 := [unit-resolution #9298 #8758]: #4190
+#7298 := (not #4190)
+#7299 := (or #4222 #7298)
+#7300 := [def-axiom]: #7299
+#8812 := [unit-resolution #7300 #8814]: #4222
+#4225 := (not #4222)
+#4228 := (or #4225 #4220)
+#7231 := (or #3932 #4225 #4220)
+#4178 := (+ #1235 #4175)
+#4179 := (+ #190 #4178)
+#4180 := (<= #4179 0::Int)
+#4218 := (or #4177 #4180)
+#4219 := (not #4218)
+#4221 := (or #4219 #4220)
+#7236 := (or #3932 #4221)
+#7240 := (iff #7236 #7231)
+#7237 := (or #3932 #4228)
+#7239 := (iff #7237 #7231)
+#7186 := [rewrite]: #7239
+#7235 := (iff #7236 #7237)
+#4229 := (iff #4221 #4228)
+#4226 := (iff #4219 #4225)
+#4223 := (iff #4218 #4222)
+#4193 := (iff #4180 #4190)
+#4187 := (<= #4175 0::Int)
+#4191 := (iff #4187 #4190)
+#4192 := [rewrite]: #4191
+#4188 := (iff #4180 #4187)
+#4185 := (= #4179 #4175)
+#4186 := [rewrite]: #4185
+#4189 := [monotonicity #4186]: #4188
+#4194 := [trans #4189 #4192]: #4193
+#4224 := [monotonicity #4194]: #4223
+#4227 := [monotonicity #4224]: #4226
+#4230 := [monotonicity #4227]: #4229
+#7238 := [monotonicity #4230]: #7235
+#7244 := [trans #7238 #7186]: #7240
+#7187 := [quant-inst #186]: #7236
+#7245 := [mp #7187 #7244]: #7231
+#8876 := [unit-resolution #7245 #8861]: #4228
+#9179 := [unit-resolution #8876 #8812]: #4220
+#8016 := (= #2028 #4167)
+#9174 := [monotonicity #8023]: #8016
+#9263 := [trans #9174 #9179]: #8124
+#9301 := [trans #9263 #9295]: #2030
+#9380 := [unit-resolution #8011 #9301]: false
+#9334 := [lemma #9380]: #9402
+#10264 := [unit-resolution #9334 #10218]: #7841
+#7824 := (or #7818 #7664)
+#3636 := (or #2031 #2026)
+#3632 := [def-axiom]: #3636
+#10266 := [unit-resolution #3632 #10219]: #2026
+#8848 := (or #2027 #7824)
+#7797 := (f9 #198 ?v0!14)
+#7815 := (= #7797 f1)
+#9264 := [hypothesis]: #2026
+#7840 := (= #7797 #2025)
+#7882 := [monotonicity #9297]: #7840
+#8403 := [trans #7882 #9264]: #7815
+#9164 := (not #7815)
+#7829 := (iff #7815 #7824)
+#8915 := (or #7628 #7829)
+#7819 := (if #7818 #4146 #7664)
+#7820 := (iff #7815 #7819)
+#8856 := (or #7628 #7820)
+#9309 := (iff #8856 #8915)
+#9313 := (iff #8915 #8915)
+#9314 := [rewrite]: #9313
+#7830 := (iff #7820 #7829)
+#7827 := (iff #7819 #7824)
+#7821 := (if #7818 true #7664)
+#7825 := (iff #7821 #7824)
+#7826 := [rewrite]: #7825
+#7822 := (iff #7819 #7821)
+#7823 := [monotonicity #4149]: #7822
+#7828 := [trans #7823 #7826]: #7827
+#7831 := [monotonicity #7828]: #7830
+#9244 := [monotonicity #7831]: #9309
+#8881 := [trans #9244 #9314]: #9309
+#9311 := [quant-inst #115 #186 #3 #2024]: #8856
+#8878 := [mp #9311 #8881]: #8915
+#9183 := [unit-resolution #8878 #3723]: #7829
+#8883 := (not #7829)
+#9266 := (or #8883 #9164)
+#7847 := (not #7824)
+#9239 := [hypothesis]: #7847
+#8857 := (or #8883 #9164 #7824)
+#8858 := [def-axiom]: #8857
+#9241 := [unit-resolution #8858 #9239]: #9266
+#9302 := [unit-resolution #9241 #9183]: #9164
+#8636 := [unit-resolution #9302 #8403]: false
+#8809 := [lemma #8636]: #8848
+#10473 := [unit-resolution #8809 #10266]: #7824
+#7848 := (or #7847 #7818 #7664)
+#7849 := [def-axiom]: #7848
+#10408 := [unit-resolution #7849 #10473 #10264]: #7664
+#7844 := (not #7664)
+#9165 := (or #7844 #7981)
+#8853 := [hypothesis]: #7664
+#8272 := (not #7981)
+#8886 := [hypothesis]: #8272
+#3545 := (or #4057 #189)
+#3546 := [def-axiom]: #3545
+#8131 := [unit-resolution #3546 #8264]: #189
+#3532 := (or #4069 #3815)
+#3487 := [def-axiom]: #3532
+#8132 := [unit-resolution #3487 #7099]: #3815
+#7991 := (or #3820 #188 #7844 #7981)
+#7982 := (or #188 #7844 #7981)
+#7996 := (or #3820 #7982)
+#8203 := (iff #7996 #7991)
+#8204 := [rewrite]: #8203
+#8202 := [quant-inst #2024 #186]: #7996
+#8205 := [mp #8202 #8204]: #7991
+#8936 := [unit-resolution #8205 #8132 #8131 #8886 #8853]: false
+#9170 := [lemma #8936]: #9165
+#10474 := [unit-resolution #9170 #10408]: #7981
+#10516 := (or #7715 #8272)
+#8693 := (>= #7695 0::Int)
+#7897 := (= #7695 0::Int)
+#9389 := (not #7897)
+#9660 := (not #8693)
+#9386 := [hypothesis]: #9660
+#9403 := (or #9389 #8693)
+#9379 := [th-lemma arith triangle-eq]: #9403
+#9404 := [unit-resolution #9379 #9386]: #9389
+#7892 := (= f28 ?v0!14)
+#7893 := (<= #7695 0::Int)
+#9385 := (or #8693 #7893)
+#9405 := [th-lemma arith farkas 1 1]: #9385
+#9406 := [unit-resolution #9405 #9386]: #7893
+#7894 := (not #7893)
+#7895 := (or #7892 #7894)
+#344 := (<= #57 0::Int)
+#345 := (not #344)
+#348 := (or #56 #345)
+#3730 := (forall (vars (?v0 S2) (?v1 S2)) (:pat #3689) #348)
+#351 := (forall (vars (?v0 S2) (?v1 S2)) #348)
+#3733 := (iff #351 #3730)
+#3731 := (iff #348 #348)
+#3732 := [refl]: #3731
+#3734 := [quant-intro #3732]: #3733
+#1585 := (~ #351 #351)
+#1584 := (~ #348 #348)
+#1614 := [refl]: #1584
+#1586 := [nnf-pos #1614]: #1585
+#62 := (< 0::Int #57)
+#63 := (implies #61 #62)
+#64 := (forall (vars (?v0 S2) (?v1 S2)) #63)
+#354 := (iff #64 #351)
+#337 := (or #56 #62)
+#341 := (forall (vars (?v0 S2) (?v1 S2)) #337)
+#352 := (iff #341 #351)
+#349 := (iff #337 #348)
+#346 := (iff #62 #345)
+#347 := [rewrite]: #346
+#350 := [monotonicity #347]: #349
+#353 := [quant-intro #350]: #352
+#342 := (iff #64 #341)
+#339 := (iff #63 #337)
+#340 := [rewrite]: #339
+#343 := [quant-intro #340]: #342
+#355 := [trans #343 #353]: #354
+#336 := [asserted]: #64
+#356 := [mp #336 #355]: #351
+#1615 := [mp~ #356 #1586]: #351
+#3735 := [mp #1615 #3734]: #3730
+#6342 := (not #3730)
+#8933 := (or #6342 #7892 #7894)
+#8931 := (or #6342 #7895)
+#8940 := (iff #8931 #8933)
+#8941 := [rewrite]: #8940
+#8926 := [quant-inst #186 #2024]: #8931
+#8934 := [mp #8926 #8941]: #8933
+#9408 := [unit-resolution #8934 #3735]: #7895
+#9410 := [unit-resolution #9408 #9406]: #7892
+#7896 := (not #7892)
+#7904 := (or #7896 #7897)
+#8960 := (or #6738 #7896 #7897)
+#8945 := (or #6738 #7904)
+#8665 := (iff #8945 #8960)
+#8668 := [rewrite]: #8665
+#8954 := [quant-inst #186 #2024]: #8945
+#8958 := [mp #8954 #8668]: #8960
+#9411 := [unit-resolution #8958 #3729]: #7904
+#9400 := [unit-resolution #9411 #9410 #9404]: false
+#9401 := [lemma #9400]: #8693
+#9661 := (or #9660 #7715 #8272)
+#8269 := [hypothesis]: #7981
+#9623 := [hypothesis]: #8867
+#9624 := [hypothesis]: #8693
+#9659 := [th-lemma arith farkas 1 -1 1 #9624 #9623 #8269]: false
+#9662 := [lemma #9659]: #9661
+#10513 := [unit-resolution #9662 #9401]: #10516
+#10254 := [unit-resolution #10513 #10474 #10265]: false
+#10267 := [lemma #10254]: #2031
+#3539 := (or #4057 #4051)
+#3540 := [def-axiom]: #3539
+#9888 := [unit-resolution #3540 #8264]: #4051
+#3533 := (or #4057 #3919)
+#3479 := [def-axiom]: #3533
+#8832 := [unit-resolution #3479 #8264]: #3919
+#4211 := (or #2011 #3932 #3924)
+#5422 := [hypothesis]: #3919
+#5349 := [hypothesis]: #3927
+#5149 := [hypothesis]: #2012
+#4951 := (<= #2010 0::Int)
+#4210 := (or #4951 #2011)
+#4205 := [th-lemma arith farkas 1 1]: #4210
+#4212 := [unit-resolution #4205 #5149]: #4951
+#5428 := (not #4951)
+#5460 := (or #5428 #3924 #3932 #2011)
+#4742 := (f5 #200 ?v0!13)
+#4743 := (f15 #4742)
+#4824 := (+ #2009 #4743)
+#4825 := (+ #190 #4824)
+#4953 := (>= #4825 0::Int)
+#4826 := (= #4825 0::Int)
+#4764 := (* -1::Int #4743)
+#4765 := (+ f14 #4764)
+#4766 := (<= #4765 0::Int)
+#4886 := (not #4766)
+#4685 := (* -1::Int #2007)
+#4797 := (+ #4685 #4743)
+#4798 := (+ #190 #4797)
+#4799 := (>= #4798 0::Int)
+#4959 := (or #4766 #4799)
+#4964 := (not #4959)
+#4960 := (= #2008 #2007)
+#5304 := (not #4960)
+#4946 := (= #2007 #2008)
+#5150 := (not #4946)
+#5348 := (iff #5150 #5304)
+#5128 := (iff #4946 #4960)
+#5347 := [commutativity]: #5128
+#5343 := [monotonicity #5347]: #5348
+#5151 := (or #5150 #2011)
+#5345 := [th-lemma arith triangle-eq]: #5151
+#5346 := [unit-resolution #5345 #5149]: #5150
+#5127 := [mp #5346 #5343]: #5304
+#4968 := (or #4964 #4960)
+#4973 := (or #3932 #4964 #4960)
+#4767 := (+ #1235 #4764)
+#4762 := (+ #2007 #4767)
+#4763 := (<= #4762 0::Int)
+#4954 := (or #4766 #4763)
+#4955 := (not #4954)
+#4961 := (or #4955 #4960)
+#4978 := (or #3932 #4961)
+#4878 := (iff #4978 #4973)
+#4880 := (or #3932 #4968)
+#4882 := (iff #4880 #4973)
+#4883 := [rewrite]: #4882
+#4881 := (iff #4978 #4880)
+#4971 := (iff #4961 #4968)
+#4969 := (iff #4955 #4964)
+#4962 := (iff #4954 #4959)
+#4822 := (iff #4763 #4799)
+#4772 := (+ #2007 #4764)
+#4793 := (+ #1235 #4772)
+#4796 := (<= #4793 0::Int)
+#4800 := (iff #4796 #4799)
+#4801 := [rewrite]: #4800
+#4791 := (iff #4763 #4796)
+#4794 := (= #4762 #4793)
+#4795 := [rewrite]: #4794
+#4792 := [monotonicity #4795]: #4791
+#4823 := [trans #4792 #4801]: #4822
+#4963 := [monotonicity #4823]: #4962
+#4970 := [monotonicity #4963]: #4969
+#4972 := [monotonicity #4970]: #4971
+#4879 := [monotonicity #4972]: #4881
+#4884 := [trans #4879 #4883]: #4878
+#4979 := [quant-inst #2006]: #4978
+#4885 := [mp #4979 #4884]: #4973
+#5350 := [unit-resolution #4885 #5349]: #4968
+#5351 := [unit-resolution #5350 #5127]: #4964
+#4980 := (or #4959 #4886)
+#4943 := [def-axiom]: #4980
+#5420 := [unit-resolution #4943 #5351]: #4886
+#4941 := (not #4799)
+#4942 := (or #4959 #4941)
+#4944 := [def-axiom]: #4942
+#5421 := [unit-resolution #4944 #5351]: #4941
+#4829 := (or #4766 #4799 #4826)
+#4852 := (or #3924 #4766 #4799 #4826)
+#4768 := (+ #4743 #2009)
+#4769 := (+ #190 #4768)
+#4770 := (= #4769 0::Int)
+#4771 := (or #4766 #4763 #4770)
+#4853 := (or #3924 #4771)
+#4858 := (iff #4853 #4852)
+#4849 := (or #3924 #4829)
+#4856 := (iff #4849 #4852)
+#4857 := [rewrite]: #4856
+#4850 := (iff #4853 #4849)
+#4830 := (iff #4771 #4829)
+#4827 := (iff #4770 #4826)
+#4820 := (= #4769 #4825)
+#4821 := [rewrite]: #4820
+#4828 := [monotonicity #4821]: #4827
+#4851 := [monotonicity #4823 #4828]: #4830
+#4855 := [monotonicity #4851]: #4850
+#4859 := [trans #4855 #4857]: #4858
+#4854 := [quant-inst #2006]: #4853
+#4887 := [mp #4854 #4859]: #4852
+#5423 := [unit-resolution #4887 #5422]: #4829
+#5418 := [unit-resolution #5423 #5421 #5420]: #4826
+#5424 := (not #4826)
+#5395 := (or #5424 #4953)
+#5419 := [th-lemma arith triangle-eq]: #5395
+#5425 := [unit-resolution #5419 #5418]: #4953
+#5426 := [hypothesis]: #4951
+#5427 := [th-lemma arith farkas 1 -1 1 #5426 #5421 #5425]: false
+#5480 := [lemma #5427]: #5460
+#4213 := [unit-resolution #5480 #4212 #5149 #5349 #5422]: false
+#4215 := [lemma #4213]: #4211
+#9889 := [unit-resolution #4215 #8861 #8832]: #2011
+#3538 := (or #4054 #2012 #4048)
+#3431 := [def-axiom]: #3538
+#9893 := [unit-resolution #3431 #9889 #9888]: #4048
+#3559 := (or #4045 #4039)
+#3560 := [def-axiom]: #3559
+#18769 := [unit-resolution #3560 #9893]: #4039
+#3558 := (or #4042 #2032 #4036)
+#3554 := [def-axiom]: #3558
+#18770 := [unit-resolution #3554 #18769]: #4039
+#18771 := [unit-resolution #18770 #10267]: #4036
+#3586 := (or #4033 #3944)
+#3564 := [def-axiom]: #3586
+#18772 := [unit-resolution #3564 #18771]: #3944
+#11863 := (or #3949 #4630 #5176)
+#11888 := (or #3949 #5179)
+#11865 := (iff #11888 #11863)
+#11884 := [rewrite]: #11865
+#11905 := [quant-inst #2123]: #11888
+#11867 := [mp #11905 #11884]: #11863
+#10037 := [unit-resolution #11867 #18772]: #5179
+#15919 := [unit-resolution #10037 #15519]: #4630
+#15588 := [mp #15919 #15539]: #15284
+#15473 := (not #14478)
+#15461 := (or #15473 #14450 #15350)
+#15360 := [def-axiom]: #15461
+#15572 := [unit-resolution #15360 #15588 #16371]: #15350
+#15307 := (or #14460 #15274)
+#15417 := [def-axiom]: #15307
+#15639 := [unit-resolution #15417 #15572]: #15274
+#15258 := [mp #15639 #14829]: #19613
+#5210 := (f5 #200 ?v0!20)
+#5211 := (f15 #5210)
+#19610 := (<= #5211 0::Int)
+#19614 := (= #5211 0::Int)
+#5267 := (+ #2127 #5211)
+#5268 := (+ #190 #5267)
+#14690 := (<= #5268 0::Int)
+#5271 := (= #5268 0::Int)
+#5228 := (+ #5194 #5211)
+#5229 := (+ #190 #5228)
+#5230 := (>= #5229 0::Int)
+#5212 := (* -1::Int #5211)
+#5213 := (+ f14 #5212)
+#5214 := (<= #5213 0::Int)
+#5235 := (or #5214 #5230)
+#5238 := (not #5235)
+#5241 := (or #5238 #5176)
+#11930 := (or #3932 #5238 #5176)
+#5215 := (+ #1235 #5212)
+#5216 := (+ #5169 #5215)
+#5217 := (<= #5216 0::Int)
+#5218 := (or #5214 #5217)
+#5219 := (not #5218)
+#5220 := (or #5219 #5176)
+#11948 := (or #3932 #5220)
+#11916 := (iff #11948 #11930)
+#11956 := (or #3932 #5241)
+#11947 := (iff #11956 #11930)
+#11957 := [rewrite]: #11947
+#11952 := (iff #11948 #11956)
+#5242 := (iff #5220 #5241)
+#5239 := (iff #5219 #5238)
+#5236 := (iff #5218 #5235)
+#5233 := (iff #5217 #5230)
+#5221 := (+ #5169 #5212)
+#5222 := (+ #1235 #5221)
+#5225 := (<= #5222 0::Int)
+#5231 := (iff #5225 #5230)
+#5232 := [rewrite]: #5231
+#5226 := (iff #5217 #5225)
+#5223 := (= #5216 #5222)
+#5224 := [rewrite]: #5223
+#5227 := [monotonicity #5224]: #5226
+#5234 := [trans #5227 #5232]: #5233
+#5237 := [monotonicity #5234]: #5236
+#5240 := [monotonicity #5237]: #5239
+#5243 := [monotonicity #5240]: #5242
+#11958 := [monotonicity #5243]: #11952
+#11917 := [trans #11958 #11957]: #11916
+#11951 := [quant-inst #2123]: #11948
+#11918 := [mp #11951 #11917]: #11930
+#14638 := [unit-resolution #11918 #8861]: #5241
+#15554 := [unit-resolution #14638 #15519]: #5238
+#19224 := (or #5235 #5271)
+#19170 := (not #5271)
+#19168 := [hypothesis]: #19170
+#11915 := (not #5214)
+#19162 := [hypothesis]: #5238
+#11961 := (or #5235 #11915)
+#11877 := [def-axiom]: #11961
+#19173 := [unit-resolution #11877 #19162]: #11915
+#11896 := (not #5230)
+#11943 := (or #5235 #11896)
+#11880 := [def-axiom]: #11943
+#19184 := [unit-resolution #11880 #19162]: #11896
+#5274 := (or #5214 #5230 #5271)
+#11968 := (or #3924 #5214 #5230 #5271)
+#5263 := (+ #5211 #2127)
+#5264 := (+ #190 #5263)
+#5265 := (= #5264 0::Int)
+#5266 := (or #5214 #5217 #5265)
+#11881 := (or #3924 #5266)
+#11987 := (iff #11881 #11968)
+#11986 := (or #3924 #5274)
+#11985 := (iff #11986 #11968)
+#11984 := [rewrite]: #11985
+#11929 := (iff #11881 #11986)
+#5275 := (iff #5266 #5274)
+#5272 := (iff #5265 #5271)
+#5269 := (= #5264 #5268)
+#5270 := [rewrite]: #5269
+#5273 := [monotonicity #5270]: #5272
+#5276 := [monotonicity #5234 #5273]: #5275
+#11965 := [monotonicity #5276]: #11929
+#11971 := [trans #11965 #11984]: #11987
+#11962 := [quant-inst #2123]: #11881
+#11989 := [mp #11962 #11971]: #11968
+#19203 := [unit-resolution #11989 #8832]: #5274
+#19204 := [unit-resolution #19203 #19184 #19173 #19168]: false
+#19220 := [lemma #19204]: #19224
+#15563 := [unit-resolution #19220 #15554]: #5271
+#14689 := (or #19170 #14690)
+#14714 := [th-lemma arith triangle-eq]: #14689
+#15571 := [unit-resolution #14714 #15563]: #14690
+#14679 := (>= #5268 0::Int)
+#14688 := (or #19170 #14679)
+#12799 := [th-lemma arith triangle-eq]: #14688
+#15545 := [unit-resolution #12799 #15563]: #14679
+#4168 := (* -1::Int #4167)
+#4169 := (+ #190 #4168)
+#7297 := (<= #4169 0::Int)
+#7302 := (= #190 #4167)
+#18257 := (iff #4220 #7302)
+#18255 := (iff #7302 #4220)
+#18256 := [commutativity]: #18255
+#18258 := [symm #18256]: #18257
+#18259 := [mp #9179 #18258]: #7302
+#18260 := (not #7302)
+#18261 := (or #18260 #7297)
+#18262 := [th-lemma arith triangle-eq]: #18261
+#18263 := [unit-resolution #18262 #18259]: #7297
+#4170 := (>= #4169 0::Int)
+#3555 := (or #4045 #3935)
+#3556 := [def-axiom]: #3555
+#9894 := [unit-resolution #3556 #9893]: #3935
+#7218 := (or #3940 #4170)
+#7219 := [quant-inst #186]: #7218
+#10752 := [unit-resolution #7219 #9894]: #4170
+#5157 := (+ #2126 #4168)
+#5318 := (<= #5157 0::Int)
+#5330 := (+ #4168 #5212)
+#5331 := (+ #2126 #5330)
+#5332 := (= #5331 0::Int)
+#14652 := (>= #5331 0::Int)
+#14681 := (not #14690)
+#15108 := (or #14681 #14652)
+#10393 := (not #4170)
+#15550 := (or #14681 #10393 #14652)
+#15472 := [th-lemma arith assign-bounds -1 1]: #15550
+#15621 := [unit-resolution #15472 #10752]: #15108
+#15637 := [unit-resolution #15621 #15571]: #14652
+#14720 := (<= #5331 0::Int)
+#12661 := (not #7297)
+#14678 := (not #14679)
+#15123 := (or #14720 #14678 #12661)
+#15620 := [th-lemma arith assign-bounds 1 -1]: #15123
+#13505 := [unit-resolution #15620 #15545 #18263]: #14720
+#19121 := (not #14652)
+#12624 := (not #14720)
+#15596 := (or #5332 #12624 #19121)
+#15676 := [th-lemma arith triangle-eq]: #15596
+#15984 := [unit-resolution #15676 #13505 #15637]: #5332
+#5337 := (not #5332)
+#15716 := (or #5318 #5337)
+#4518 := (f9 f29 f28)
+#4519 := (= #4518 f1)
+#4144 := (f9 #198 f28)
+#4145 := (= #4144 f1)
+#31 := (:var 0 S1)
+#28 := (:var 2 S7)
+#29 := (f12 f13 #28)
+#30 := (f11 #29 #9)
+#32 := (f10 #30 #31)
+#3710 := (pattern #32)
+#35 := (= #31 f1)
+#33 := (f9 #32 #9)
+#34 := (= #33 f1)
+#36 := (iff #34 #35)
+#3711 := (forall (vars (?v0 S7) (?v1 S2) (?v2 S1)) (:pat #3710) #36)
+#37 := (forall (vars (?v0 S7) (?v1 S2) (?v2 S1)) #36)
+#3714 := (iff #37 #3711)
+#3712 := (iff #36 #36)
+#3713 := [refl]: #3712
+#3715 := [quant-intro #3713]: #3714
+#1577 := (~ #37 #37)
+#1606 := (~ #36 #36)
+#1607 := [refl]: #1606
+#1578 := [nnf-pos #1607]: #1577
+#321 := [asserted]: #37
+#1608 := [mp~ #321 #1578]: #37
+#3716 := [mp #1608 #3715]: #3711
+#6379 := (not #3711)
+#6555 := (or #6379 #4145)
+#4147 := (iff #4145 #4146)
+#6775 := (or #6379 #4147)
+#7023 := (iff #6775 #6555)
+#7024 := (iff #6555 #6555)
+#3437 := [rewrite]: #7024
+#4155 := (iff #4147 #4145)
+#4150 := (iff #4145 true)
+#4153 := (iff #4150 #4145)
+#4154 := [rewrite]: #4153
+#4151 := (iff #4147 #4150)
+#4152 := [monotonicity #4149]: #4151
+#4156 := [trans #4152 #4154]: #4155
+#7020 := [monotonicity #4156]: #7023
+#3310 := [trans #7020 #3437]: #7023
+#7010 := [quant-inst #115 #186 #3]: #6775
+#6101 := [mp #7010 #3310]: #6555
+#12643 := [unit-resolution #6101 #3716]: #4145
+#12662 := (= #4518 #4144)
+#12669 := [monotonicity #9701]: #12662
+#12665 := [trans #12669 #12643]: #4519
+#4447 := (= #221 #110)
+#8042 := (= #190 #110)
+#7927 := (= #110 #190)
+#4428 := (+ #110 #1235)
+#4429 := (>= #4428 0::Int)
+#4424 := (f9 f21 f16)
+#4425 := (= #4424 f1)
+#7254 := (not #4425)
+#4381 := (= ?v0!12 f16)
+#4382 := (?v1!7 ?v0!12)
+#4390 := (f6 f7 #4382)
+#4391 := (f5 #4390 ?v0!12)
+#4392 := (f15 #4391)
+#4393 := (* -1::Int #4392)
+#4383 := (f19 f20 #4382)
+#4384 := (* -1::Int #4383)
+#4394 := (+ #4384 #4393)
+#4395 := (+ #1965 #4394)
+#4396 := (= #4395 0::Int)
+#4397 := (not #4396)
+#4387 := (f9 f21 #4382)
+#4388 := (= #4387 f1)
+#4389 := (not #4388)
+#4385 := (+ #1965 #4384)
+#4386 := (<= #4385 0::Int)
+#4398 := (or #4386 #4389 #4397)
+#4252 := (= f28 f16)
+#4431 := (f5 #200 f16)
+#4432 := (f15 #4431)
+#4439 := (* -1::Int #4432)
+#4442 := (+ #1235 #4439)
+#4443 := (+ #110 #4442)
+#4444 := (<= #4443 0::Int)
+#7610 := (not #4444)
+#4440 := (+ f14 #4439)
+#4441 := (<= #4440 0::Int)
+#4445 := (or #4441 #4444)
+#4446 := (not #4445)
+#9205 := (not #4447)
+#5924 := (iff #713 #9205)
+#5922 := (iff #222 #4447)
+#5921 := (iff #4447 #222)
+#5919 := [monotonicity #7324]: #5921
+#5923 := [symm #5919]: #5922
+#6064 := [monotonicity #5923]: #5924
+#5920 := [hypothesis]: #713
+#6065 := [mp #5920 #6064]: #9205
+#4448 := (or #4446 #4447)
+#7601 := (or #3932 #4446 #4447)
+#7602 := (or #3932 #4448)
+#7604 := (iff #7602 #7601)
+#7605 := [rewrite]: #7604
+#7603 := [quant-inst #65]: #7602
+#7606 := [mp #7603 #7605]: #7601
+#6066 := [unit-resolution #7606 #8861]: #4448
+#9900 := [unit-resolution #6066 #6065]: #4446
+#7611 := (or #4445 #7610)
+#7612 := [def-axiom]: #7611
+#9807 := [unit-resolution #7612 #9900]: #7610
+#8599 := (or #4444 #4252)
+#8514 := (<= #4432 0::Int)
+#8515 := (not #8514)
+#8195 := (not #4252)
+#8549 := [hypothesis]: #8195
+#8513 := (or #6342 #4252 #8515)
+#8516 := (or #4252 #8515)
+#8519 := (or #6342 #8516)
+#8521 := (iff #8519 #8513)
+#8522 := [rewrite]: #8521
+#8520 := [quant-inst #186 #65]: #8519
+#8523 := [mp #8520 #8522]: #8513
+#8552 := [unit-resolution #8523 #3735 #8549]: #8515
+#4483 := (<= #110 0::Int)
+#7325 := (or #878 #4483)
+#7328 := [th-lemma arith triangle-eq]: #7325
+#7329 := [unit-resolution #7328 #7324]: #4483
+#4253 := (?v1!7 f28)
+#4254 := (f19 f20 #4253)
+#4255 := (* -1::Int #4254)
+#4256 := (+ #190 #4255)
+#7963 := (>= #4256 0::Int)
+#4257 := (<= #4256 0::Int)
+#7326 := (not #4257)
+#4261 := (f6 f7 #4253)
+#4262 := (f5 #4261 f28)
+#4263 := (f15 #4262)
+#4264 := (* -1::Int #4263)
+#4265 := (+ #4255 #4264)
+#4266 := (+ #190 #4265)
+#4267 := (= #4266 0::Int)
+#4268 := (not #4267)
+#4258 := (f9 f21 #4253)
+#4259 := (= #4258 f1)
+#4260 := (not #4259)
+#4269 := (or #4257 #4260 #4268)
+#4270 := (not #4269)
+#8572 := (or #4252 #4270)
+#3547 := (or #4057 #1412)
+#3550 := [def-axiom]: #3547
+#8267 := [unit-resolution #3550 #8264]: #1412
+#7305 := (or #3836 #4252 #1411 #4270)
+#4271 := (or #4252 #1411 #4270)
+#7316 := (or #3836 #4271)
+#7315 := (iff #7316 #7305)
+#7317 := [rewrite]: #7315
+#7314 := [quant-inst #186]: #7316
+#7318 := [mp #7314 #7317]: #7305
+#8573 := [unit-resolution #7318 #7150 #8267]: #8572
+#8574 := [unit-resolution #8573 #8549]: #4270
+#7322 := (or #4269 #7326)
+#7327 := [def-axiom]: #7322
+#8575 := [unit-resolution #7327 #8574]: #7326
+#8570 := (or #7963 #4257)
+#8576 := [th-lemma arith farkas 1 1]: #8570
+#8577 := [unit-resolution #8576 #8575]: #7963
+#8444 := (>= #4254 0::Int)
+#3531 := (or #4069 #3806)
+#3511 := [def-axiom]: #3531
+#8408 := [unit-resolution #3511 #7099]: #3806
+#7889 := (or #3811 #8444)
+#7890 := [quant-inst #4253]: #7889
+#8578 := [unit-resolution #7890 #8408]: #8444
+#8579 := [hypothesis]: #7610
+#8580 := [th-lemma arith farkas 1 1 1 1 1 #8579 #8578 #8577 #7329 #8552]: false
+#8600 := [lemma #8580]: #8599
+#9903 := [unit-resolution #8600 #9807]: #4252
+#11140 := (or #4398 #8195)
+#4399 := (not #4398)
+#11053 := [hypothesis]: #4399
+#7585 := (or #4398 #4388)
+#7586 := [def-axiom]: #7585
+#11054 := [unit-resolution #7586 #11053]: #4388
+#11117 := (= #187 #4387)
+#8916 := (= f28 #4382)
+#10478 := (= f16 #4382)
+#8133 := (= #4382 f16)
+#8136 := (?v1!7 #4382)
+#8144 := (f6 f7 #8136)
+#8145 := (f5 #8144 #4382)
+#8146 := (f15 #8145)
+#8147 := (* -1::Int #8146)
+#8137 := (f19 f20 #8136)
+#8138 := (* -1::Int #8137)
+#8148 := (+ #8138 #8147)
+#8149 := (+ #4383 #8148)
+#8150 := (= #8149 0::Int)
+#8151 := (not #8150)
+#8141 := (f9 f21 #8136)
+#8142 := (= #8141 f1)
+#8143 := (not #8142)
+#8139 := (+ #4383 #8138)
+#8140 := (<= #8139 0::Int)
+#8152 := (or #8140 #8143 #8151)
+#9559 := (+ #110 #4384)
+#9591 := (>= #9559 0::Int)
+#8009 := [hypothesis]: #4252
+#8040 := [monotonicity #8009]: #8042
+#8044 := [symm #8040]: #7927
+#7986 := (not #7927)
+#7987 := (or #7986 #4429)
+#7985 := [th-lemma arith triangle-eq]: #7987
+#8266 := [unit-resolution #7985 #8044]: #4429
+#8029 := (+ #190 #4384)
+#8030 := (>= #8029 0::Int)
+#9838 := (or #3820 #188 #4389 #8030)
+#8037 := (or #188 #4389 #8030)
+#9839 := (or #3820 #8037)
+#9884 := (iff #9839 #9838)
+#9885 := [rewrite]: #9884
+#9597 := [quant-inst #4382 #186]: #9839
+#9897 := [mp #9597 #9885]: #9838
+#11049 := [unit-resolution #9897 #8132 #8131 #11054]: #8030
+#8411 := (not #4429)
+#11055 := (not #8030)
+#11056 := (or #9591 #11055 #8411)
+#11057 := [th-lemma arith assign-bounds -1 -1]: #11056
+#11058 := [unit-resolution #11057 #11049 #8266]: #9591
+#10720 := (not #9591)
+#11080 := (or #10720 #8140)
+#9223 := (>= #8137 0::Int)
+#10115 := (not #9223)
+#10116 := [hypothesis]: #10115
+#10093 := (or #3811 #9223)
+#10094 := [quant-inst #8136]: #10093
+#10134 := [unit-resolution #10094 #8408 #10116]: false
+#10135 := [lemma #10134]: #9223
+#7335 := (not #4483)
+#11059 := (or #10720 #7335 #10115 #8140)
+#11079 := [th-lemma arith assign-bounds -1 1 1]: #11059
+#11081 := [unit-resolution #11079 #10135 #7329]: #11080
+#11082 := [unit-resolution #11081 #11058]: #8140
+#10256 := (not #8140)
+#10257 := (or #8152 #10256)
+#10263 := [def-axiom]: #10257
+#11083 := [unit-resolution #10263 #11082]: #8152
+#8153 := (not #8152)
+#11107 := (or #8133 #8153)
+#8134 := (+ f14 #4384)
+#8135 := (<= #8134 0::Int)
+#11087 := (not #8135)
+#8025 := (>= #4385 0::Int)
+#7582 := (not #4386)
+#7583 := (or #4398 #7582)
+#7584 := [def-axiom]: #7583
+#11078 := [unit-resolution #7584 #11053]: #7582
+#11084 := (or #8025 #4386)
+#11085 := [th-lemma arith farkas 1 1]: #11084
+#11086 := [unit-resolution #11085 #11078]: #8025
+#11088 := (not #8025)
+#11110 := (or #11087 #11088)
+#3544 := (or #4057 #1969)
+#3549 := [def-axiom]: #3544
+#9691 := [unit-resolution #3549 #8264]: #1969
+#11108 := (or #11087 #1968 #11088)
+#11109 := [th-lemma arith assign-bounds 1 1]: #11108
+#11111 := [unit-resolution #11109 #9691]: #11110
+#11112 := [unit-resolution #11111 #11086]: #11087
+#10255 := (or #3836 #8133 #8135 #8153)
+#8154 := (or #8133 #8135 #8153)
+#10258 := (or #3836 #8154)
+#10260 := (iff #10258 #10255)
+#10261 := [rewrite]: #10260
+#10259 := [quant-inst #4382]: #10258
+#10262 := [mp #10259 #10261]: #10255
+#11113 := [unit-resolution #10262 #7150 #11112]: #11107
+#11114 := [unit-resolution #11113 #11083]: #8133
+#11115 := [symm #11114]: #10478
+#11116 := [trans #8009 #11115]: #8916
+#11137 := [monotonicity #11116]: #11117
+#11138 := [trans #11137 #11054]: #188
+#11139 := [unit-resolution #8131 #11138]: false
+#11141 := [lemma #11139]: #11140
+#9883 := [unit-resolution #11141 #9903]: #4398
+#8305 := (or #4381 #4399)
+#7574 := (or #3836 #4381 #1968 #4399)
+#4400 := (or #4381 #1968 #4399)
+#7575 := (or #3836 #4400)
+#7577 := (iff #7575 #7574)
+#7578 := [rewrite]: #7577
+#7576 := [quant-inst #1961]: #7575
+#7579 := [mp #7576 #7578]: #7574
+#8780 := [unit-resolution #7579 #7150 #9691]: #8305
+#18251 := [unit-resolution #8780 #9883]: #4381
+#9007 := (not #4381)
+#9005 := (or #9007 #7254)
+#7350 := [hypothesis]: #4425
+#8961 := (= #1962 #4424)
+#8959 := [hypothesis]: #4381
+#8962 := [monotonicity #8959]: #8961
+#9006 := [trans #8962 #7350]: #1963
+#3542 := (or #4057 #1964)
+#3543 := [def-axiom]: #3542
+#8944 := [unit-resolution #3543 #8264]: #1964
+#8888 := [unit-resolution #8944 #9006]: false
+#9008 := [lemma #8888]: #9005
+#9652 := [unit-resolution #9008 #18251]: #7254
+#4430 := (or #4425 #4429)
+#3551 := (or #4057 #3909)
+#3552 := [def-axiom]: #3551
+#8603 := [unit-resolution #3552 #8264]: #3909
+#7595 := (or #3914 #4425 #4429)
+#7596 := (or #3914 #4430)
+#7598 := (iff #7596 #7595)
+#7599 := [rewrite]: #7598
+#7597 := [quant-inst #65]: #7596
+#7600 := [mp #7597 #7599]: #7595
+#9803 := [unit-resolution #7600 #8603]: #4430
+#9901 := [unit-resolution #9803 #9652]: #4429
+#9733 := (or #7927 #8411)
+#8590 := (<= #4428 0::Int)
+#4272 := (>= #190 0::Int)
+#7146 := (or #3811 #4272)
+#7514 := [quant-inst #186]: #7146
+#7845 := [unit-resolution #7514 #8408]: #4272
+#8544 := (not #8590)
+#7810 := [hypothesis]: #8544
+#7888 := [th-lemma arith farkas 1 -1 1 #7810 #7329 #7845]: false
+#7759 := [lemma #7888]: #8590
+#9935 := (or #7927 #8544 #8411)
+#9936 := [th-lemma arith triangle-eq]: #9935
+#9606 := [unit-resolution #9936 #7759]: #9733
+#9645 := [unit-resolution #9606 #9901]: #7927
+#9892 := [symm #9645]: #8042
+#9902 := (= #221 #190)
+#8871 := (= #221 #4167)
+#8010 := (= f16 f28)
+#9669 := [symm #9903]: #8010
+#9549 := [monotonicity #9669]: #8871
+#9612 := [trans #9549 #9179]: #9902
+#9668 := [trans #9612 #9892]: #4447
+#9570 := [trans #9668 #7324]: #222
+#9595 := [unit-resolution #5920 #9570]: false
+#9632 := [lemma #9595]: #222
+#3565 := (or #4033 #4027)
+#3567 := [def-axiom]: #3565
+#14028 := [unit-resolution #3567 #18771]: #4027
+#3585 := (or #4030 #713 #4024)
+#3575 := [def-axiom]: #3585
+#14128 := [unit-resolution #3575 #14028]: #4027
+#14105 := [unit-resolution #14128 #9632]: #4024
+#3577 := (or #4021 #4015)
+#3578 := [def-axiom]: #3577
+#14080 := [unit-resolution #3578 #14105]: #4015
+#14243 := (or #4018 #4012)
+#8617 := (or #8544 #2051)
+#6626 := (f5 #200 ?v0!15)
+#6627 := (f15 #6626)
+#9520 := (= #6627 0::Int)
+#8602 := (not #9520)
+#9750 := (>= #6627 0::Int)
+#8541 := (not #9750)
+#8393 := [hypothesis]: #8590
+#6707 := [hypothesis]: #2052
+#3322 := (>= #110 0::Int)
+#6072 := (or #3811 #3322)
+#6127 := [quant-inst #65]: #6072
+#8487 := [unit-resolution #6127 #8408]: #3322
+#4577 := (* -1::Int #2050)
+#6683 := (+ #4577 #6627)
+#6684 := (+ #190 #6683)
+#9543 := (<= #6684 0::Int)
+#6687 := (= #6684 0::Int)
+#6628 := (* -1::Int #6627)
+#6629 := (+ f14 #6628)
+#6630 := (<= #6629 0::Int)
+#9343 := (not #6630)
+#6585 := (f19 f20 ?v0!15)
+#6610 := (* -1::Int #6585)
+#6644 := (+ #6610 #6627)
+#6645 := (+ #190 #6644)
+#6646 := (>= #6645 0::Int)
+#6651 := (or #6630 #6646)
+#6654 := (not #6651)
+#6586 := (= #2050 #6585)
+#9841 := (not #6586)
+#6611 := (+ #2050 #6610)
+#9310 := (>= #6611 0::Int)
+#9745 := (not #9310)
+#9746 := (or #9745 #2051)
+#9740 := [hypothesis]: #9310
+#9233 := (>= #6585 0::Int)
+#9593 := (or #3811 #9233)
+#9579 := [quant-inst #2049]: #9593
+#9741 := [unit-resolution #9579 #8408]: #9233
+#9744 := [th-lemma arith farkas -1 1 1 #6707 #9741 #9740]: false
+#9747 := [lemma #9744]: #9746
+#9587 := [unit-resolution #9747 #6707]: #9745
+#9589 := (or #9841 #9310)
+#9611 := [th-lemma arith triangle-eq]: #9589
+#8492 := [unit-resolution #9611 #9587]: #9841
+#9172 := (or #3932 #6654 #6586)
+#6631 := (+ #1235 #6628)
+#6632 := (+ #6585 #6631)
+#6633 := (<= #6632 0::Int)
+#6634 := (or #6630 #6633)
+#6635 := (not #6634)
+#6636 := (or #6635 #6586)
+#9240 := (or #3932 #6636)
+#9341 := (iff #9240 #9172)
+#6657 := (or #6654 #6586)
+#8908 := (or #3932 #6657)
+#9265 := (iff #8908 #9172)
+#9339 := [rewrite]: #9265
+#8877 := (iff #9240 #8908)
+#6658 := (iff #6636 #6657)
+#6655 := (iff #6635 #6654)
+#6652 := (iff #6634 #6651)
+#6649 := (iff #6633 #6646)
+#6637 := (+ #6585 #6628)
+#6638 := (+ #1235 #6637)
+#6641 := (<= #6638 0::Int)
+#6647 := (iff #6641 #6646)
+#6648 := [rewrite]: #6647
+#6642 := (iff #6633 #6641)
+#6639 := (= #6632 #6638)
+#6640 := [rewrite]: #6639
+#6643 := [monotonicity #6640]: #6642
+#6650 := [trans #6643 #6648]: #6649
+#6653 := [monotonicity #6650]: #6652
+#6656 := [monotonicity #6653]: #6655
+#6659 := [monotonicity #6656]: #6658
+#8889 := [monotonicity #6659]: #8877
+#9337 := [trans #8889 #9339]: #9341
+#8828 := [quant-inst #2049]: #9240
+#9342 := [mp #8828 #9337]: #9172
+#8493 := [unit-resolution #9342 #8861 #8492]: #6654
+#9344 := (or #6651 #9343)
+#9299 := [def-axiom]: #9344
+#8494 := [unit-resolution #9299 #8493]: #9343
+#9330 := (not #6646)
+#8854 := (or #6651 #9330)
+#8855 := [def-axiom]: #8854
+#8307 := [unit-resolution #8855 #8493]: #9330
+#6690 := (or #6630 #6646 #6687)
+#9368 := (or #3924 #6630 #6646 #6687)
+#6679 := (+ #6627 #4577)
+#6680 := (+ #190 #6679)
+#6681 := (= #6680 0::Int)
+#6682 := (or #6630 #6633 #6681)
+#8704 := (or #3924 #6682)
+#9540 := (iff #8704 #9368)
+#9390 := (or #3924 #6690)
+#9425 := (iff #9390 #9368)
+#9539 := [rewrite]: #9425
+#9413 := (iff #8704 #9390)
+#6691 := (iff #6682 #6690)
+#6688 := (iff #6681 #6687)
+#6685 := (= #6680 #6684)
+#6686 := [rewrite]: #6685
+#6689 := [monotonicity #6686]: #6688
+#6692 := [monotonicity #6650 #6689]: #6691
+#9423 := [monotonicity #6692]: #9413
+#9536 := [trans #9423 #9539]: #9540
+#9412 := [quant-inst #2049]: #8704
+#9541 := [mp #9412 #9536]: #9368
+#9868 := [unit-resolution #9541 #8832]: #6690
+#8535 := [unit-resolution #9868 #8307 #8494]: #6687
+#9870 := (not #6687)
+#9871 := (or #9870 #9543)
+#9872 := [th-lemma arith triangle-eq]: #9871
+#8540 := [unit-resolution #9872 #8535]: #9543
+#8548 := (not #3322)
+#9874 := (not #9543)
+#8539 := (or #8541 #2051 #9874 #8544 #8548)
+#8530 := [th-lemma arith assign-bounds -1 -1 -1 1]: #8539
+#8601 := [unit-resolution #8530 #8540 #8487 #6707 #8393]: #8541
+#8604 := (or #8602 #9750)
+#8571 := [th-lemma arith triangle-eq]: #8604
+#8605 := [unit-resolution #8571 #8601]: #8602
+#9672 := (= f28 ?v0!15)
+#9673 := (<= #6627 0::Int)
+#8598 := (or #9673 #2051 #9874 #8544 #8548)
+#8613 := [th-lemma arith assign-bounds 1 1 1 1]: #8598
+#8611 := [unit-resolution #8613 #8540 #8487 #6707 #8393]: #9673
+#9674 := (not #9673)
+#9609 := (or #6342 #9672 #9674)
+#9675 := (or #9672 #9674)
+#9610 := (or #6342 #9675)
+#9622 := (iff #9610 #9609)
+#9618 := [rewrite]: #9622
+#9604 := [quant-inst #186 #2049]: #9610
+#9619 := [mp #9604 #9618]: #9609
+#8607 := [unit-resolution #9619 #3735 #8611]: #9672
+#9676 := (not #9672)
+#8402 := (or #6738 #9676 #9520)
+#9726 := (or #9676 #9520)
+#8488 := (or #6738 #9726)
+#8484 := (iff #8488 #8402)
+#8485 := [rewrite]: #8484
+#8489 := [quant-inst #186 #2049]: #8488
+#8486 := [mp #8489 #8485]: #8402
+#8619 := [unit-resolution #8486 #3729 #8607 #8605]: false
+#8618 := [lemma #8619]: #8617
+#14242 := [unit-resolution #8618 #7759]: #2051
+#3593 := (or #4018 #2052 #4012)
+#3573 := [def-axiom]: #3593
+#14109 := [unit-resolution #3573 #14242]: #14243
+#14124 := [unit-resolution #14109 #14080]: #4012
+#3596 := (or #4009 #4003)
+#3601 := [def-axiom]: #3596
+#14006 := [unit-resolution #3601 #14124]: #4003
+#14245 := (or #4006 #4000)
+#6060 := [hypothesis]: #2959
+#3627 := (not #2077)
+#3630 := (or #2954 #3627)
+#3514 := [def-axiom]: #3630
+#6061 := [unit-resolution #3514 #6060]: #3627
+#10344 := (or #2954 #2077)
+#5944 := (f19 f20 ?v1!16)
+#5961 := (* -1::Int #5944)
+#5013 := (+ #190 #5961)
+#5014 := (<= #5013 0::Int)
+#5817 := (f9 f21 ?v1!16)
+#5818 := (= #5817 f1)
+#10018 := (not #5818)
+#5816 := (= ?v1!16 f28)
+#5824 := (or #5816 #5818)
+#10022 := (not #5824)
+#5814 := (f9 #198 ?v1!16)
+#5815 := (= #5814 f1)
+#5829 := (iff #5815 #5824)
+#9980 := (or #7628 #5829)
+#5819 := (if #5816 #4146 #5818)
+#5820 := (iff #5815 #5819)
+#9981 := (or #7628 #5820)
+#10002 := (iff #9981 #9980)
+#10009 := (iff #9980 #9980)
+#10010 := [rewrite]: #10009
+#5830 := (iff #5820 #5829)
+#5827 := (iff #5819 #5824)
+#5821 := (if #5816 true #5818)
+#5825 := (iff #5821 #5824)
+#5826 := [rewrite]: #5825
+#5822 := (iff #5819 #5821)
+#5823 := [monotonicity #4149]: #5822
+#5828 := [trans #5823 #5826]: #5827
+#5831 := [monotonicity #5828]: #5830
+#10008 := [monotonicity #5831]: #10002
+#10011 := [trans #10008 #10010]: #10002
+#10007 := [quant-inst #115 #186 #3 #2064]: #9981
+#10012 := [mp #10007 #10011]: #9980
+#10496 := [unit-resolution #10012 #3723]: #5829
+#10031 := (not #5815)
+#10269 := (iff #2068 #10031)
+#10268 := (iff #2067 #5815)
+#10548 := (iff #5815 #2067)
+#10495 := (= #5814 #2066)
+#10497 := [monotonicity #9297]: #10495
+#10549 := [monotonicity #10497]: #10548
+#10526 := [symm #10549]: #10268
+#10270 := [monotonicity #10526]: #10269
+#3625 := (or #2954 #2068)
+#3628 := [def-axiom]: #3625
+#6063 := [unit-resolution #3628 #6060]: #2068
+#10271 := [mp #6063 #10270]: #10031
+#10024 := (not #5829)
+#10025 := (or #10024 #5815 #10022)
+#10026 := [def-axiom]: #10025
+#10272 := [unit-resolution #10026 #10271 #10496]: #10022
+#10019 := (or #5824 #10018)
+#10020 := [def-axiom]: #10019
+#10273 := [unit-resolution #10020 #10272]: #10018
+#5037 := (or #5818 #5014)
+#10145 := (or #3914 #5818 #5014)
+#4981 := (+ #5944 #1235)
+#4982 := (>= #4981 0::Int)
+#5007 := (or #5818 #4982)
+#10146 := (or #3914 #5007)
+#10172 := (iff #10146 #10145)
+#10167 := (or #3914 #5037)
+#10170 := (iff #10167 #10145)
+#10171 := [rewrite]: #10170
+#10168 := (iff #10146 #10167)
+#5038 := (iff #5007 #5037)
+#5035 := (iff #4982 #5014)
+#5008 := (+ #1235 #5944)
+#5011 := (>= #5008 0::Int)
+#5015 := (iff #5011 #5014)
+#5016 := [rewrite]: #5015
+#5006 := (iff #4982 #5011)
+#5009 := (= #4981 #5008)
+#5010 := [rewrite]: #5009
+#5012 := [monotonicity #5010]: #5006
+#5036 := [trans #5012 #5016]: #5035
+#5039 := [monotonicity #5036]: #5038
+#10169 := [monotonicity #5039]: #10168
+#10173 := [trans #10169 #10171]: #10172
+#10166 := [quant-inst #2064]: #10146
+#10174 := [mp #10166 #10173]: #10145
+#10537 := [unit-resolution #10174 #8603]: #5037
+#10535 := [unit-resolution #10537 #10273]: #5014
+#5741 := (f19 f20 ?v0!17)
+#5634 := (* -1::Int #5741)
+#5694 := (+ #2074 #5634)
+#5699 := (<= #5694 0::Int)
+#10122 := (or #3940 #5699)
+#5671 := (+ #5741 #2075)
+#5684 := (>= #5671 0::Int)
+#10124 := (or #3940 #5684)
+#10127 := (iff #10124 #10122)
+#10130 := (iff #10122 #10122)
+#10131 := [rewrite]: #10130
+#5701 := (iff #5684 #5699)
+#5685 := (+ #2075 #5741)
+#5689 := (>= #5685 0::Int)
+#5700 := (iff #5689 #5699)
+#5698 := [rewrite]: #5700
+#5692 := (iff #5684 #5689)
+#5690 := (= #5671 #5685)
+#5691 := [rewrite]: #5690
+#5693 := [monotonicity #5691]: #5692
+#5702 := [trans #5693 #5698]: #5701
+#10128 := [monotonicity #5702]: #10127
+#10132 := [trans #10128 #10131]: #10127
+#10126 := [quant-inst #2065]: #10124
+#10133 := [mp #10126 #10132]: #10122
+#10770 := [unit-resolution #10133 #9894]: #5699
+#10716 := [hypothesis]: #3627
+#5629 := (+ #190 #5634)
+#10727 := (>= #5629 0::Int)
+#5771 := (f9 f21 ?v0!17)
+#5772 := (= #5771 f1)
+#5770 := (= ?v0!17 f28)
+#5778 := (or #5770 #5772)
+#5760 := (f9 #198 ?v0!17)
+#5761 := (= #5760 f1)
+#5783 := (iff #5761 #5778)
+#10038 := (or #7628 #5783)
+#5773 := (if #5770 #4146 #5772)
+#5774 := (iff #5761 #5773)
+#10036 := (or #7628 #5774)
+#10046 := (iff #10036 #10038)
+#10028 := (iff #10038 #10038)
+#10048 := [rewrite]: #10028
+#5784 := (iff #5774 #5783)
+#5781 := (iff #5773 #5778)
+#5775 := (if #5770 true #5772)
+#5779 := (iff #5775 #5778)
+#5780 := [rewrite]: #5779
+#5776 := (iff #5773 #5775)
+#5777 := [monotonicity #4149]: #5776
+#5782 := [trans #5777 #5780]: #5781
+#5785 := [monotonicity #5782]: #5784
+#10047 := [monotonicity #5785]: #10046
+#10049 := [trans #10047 #10048]: #10046
+#10045 := [quant-inst #115 #186 #3 #2065]: #10036
+#10050 := [mp #10045 #10049]: #10038
+#10530 := [unit-resolution #10050 #3723]: #5783
+#3626 := (or #2954 #2070)
+#3629 := [def-axiom]: #3626
+#6062 := [unit-resolution #3629 #6060]: #2070
+#10538 := (= #5760 #2069)
+#10540 := [monotonicity #9297]: #10538
+#10544 := [trans #10540 #6062]: #5761
+#10058 := (not #5761)
+#10073 := (not #5783)
+#10059 := (or #10073 #10058 #5778)
+#10060 := [def-axiom]: #10059
+#10536 := [unit-resolution #10060 #10544 #10530]: #5778
+#10051 := (not #5770)
+#10683 := (= #2074 #4167)
+#10731 := (not #10683)
+#6074 := (+ #2074 #4168)
+#10684 := (<= #6074 0::Int)
+#10570 := (not #10684)
+#5977 := (f5 #200 ?v1!16)
+#5978 := (f15 #5977)
+#10503 := (<= #5978 0::Int)
+#10504 := (not #10503)
+#10502 := (= f28 ?v1!16)
+#10419 := (not #10502)
+#10016 := (not #5816)
+#10017 := (or #5824 #10016)
+#10015 := [def-axiom]: #10017
+#10374 := [unit-resolution #10015 #10272]: #10016
+#10357 := (or #10419 #5816)
+#10399 := [hypothesis]: #10502
+#10410 := [symm #10399]: #5816
+#10367 := [hypothesis]: #10016
+#10417 := [unit-resolution #10367 #10410]: false
+#10416 := [lemma #10417]: #10357
+#10418 := [unit-resolution #10416 #10374]: #10419
+#10343 := (or #10504 #10502)
+#10409 := [hypothesis]: #10419
+#10318 := [hypothesis]: #10503
+#10322 := (or #6342 #10502 #10504)
+#10505 := (or #10502 #10504)
+#10338 := (or #6342 #10505)
+#10339 := (iff #10338 #10322)
+#10340 := [rewrite]: #10339
+#10295 := [quant-inst #186 #2064]: #10338
+#10341 := [mp #10295 #10340]: #10322
+#10422 := [unit-resolution #10341 #3735 #10318 #10409]: false
+#10472 := [lemma #10422]: #10343
+#10405 := [unit-resolution #10472 #10418]: #10504
+#10571 := (not #5014)
+#10467 := (or #10570 #10503 #2077 #10571)
+#4581 := (* -1::Int #2073)
+#6036 := (+ #4581 #5978)
+#6037 := (+ #190 #6036)
+#6040 := (= #6037 0::Int)
+#10441 := (not #6040)
+#10120 := (<= #6037 0::Int)
+#10719 := (not #10120)
+#10714 := [hypothesis]: #10504
+#10567 := [hypothesis]: #10684
+#10406 := (or #10503 #2077 #10719 #10570 #10393)
+#10337 := [th-lemma arith assign-bounds 1 1 1 1]: #10406
+#10342 := [unit-resolution #10337 #10567 #10714 #10752 #10716]: #10719
+#10420 := (or #10441 #10120)
+#10421 := [th-lemma arith triangle-eq]: #10420
+#10378 := [unit-resolution #10421 #10342]: #10441
+#5996 := (+ #5961 #5978)
+#5997 := (+ #190 #5996)
+#5998 := (>= #5997 0::Int)
+#5979 := (* -1::Int #5978)
+#5980 := (+ f14 #5979)
+#5981 := (<= #5980 0::Int)
+#6003 := (or #5981 #5998)
+#6006 := (not #6003)
+#5987 := (= #2073 #5944)
+#10379 := (not #5987)
+#5962 := (+ #2073 #5961)
+#10107 := (>= #5962 0::Int)
+#10569 := (not #10107)
+#10566 := [hypothesis]: #5014
+#10572 := (or #10569 #2077 #10570 #10571)
+#10562 := [hypothesis]: #10107
+#10568 := [th-lemma arith farkas -1 1 1 -1 1 #10562 #10716 #10567 #10752 #10566]: false
+#10576 := [lemma #10568]: #10572
+#10358 := [unit-resolution #10576 #10567 #10716 #10566]: #10569
+#10380 := (or #10379 #10107)
+#10381 := [th-lemma arith triangle-eq]: #10380
+#10359 := [unit-resolution #10381 #10358]: #10379
+#6009 := (or #6006 #5987)
+#10077 := (or #3932 #6006 #5987)
+#5982 := (+ #1235 #5979)
+#5983 := (+ #5944 #5982)
+#5984 := (<= #5983 0::Int)
+#5985 := (or #5981 #5984)
+#5986 := (not #5985)
+#5988 := (or #5986 #5987)
+#10078 := (or #3932 #5988)
+#10096 := (iff #10078 #10077)
+#10079 := (or #3932 #6009)
+#10075 := (iff #10079 #10077)
+#10095 := [rewrite]: #10075
+#10080 := (iff #10078 #10079)
+#6010 := (iff #5988 #6009)
+#6007 := (iff #5986 #6006)
+#6004 := (iff #5985 #6003)
+#6001 := (iff #5984 #5998)
+#5989 := (+ #5944 #5979)
+#5990 := (+ #1235 #5989)
+#5993 := (<= #5990 0::Int)
+#5999 := (iff #5993 #5998)
+#6000 := [rewrite]: #5999
+#5994 := (iff #5984 #5993)
+#5991 := (= #5983 #5990)
+#5992 := [rewrite]: #5991
+#5995 := [monotonicity #5992]: #5994
+#6002 := [trans #5995 #6000]: #6001
+#6005 := [monotonicity #6002]: #6004
+#6008 := [monotonicity #6005]: #6007
+#6011 := [monotonicity #6008]: #6010
+#10081 := [monotonicity #6011]: #10080
+#10097 := [trans #10081 #10095]: #10096
+#10076 := [quant-inst #2064]: #10078
+#10098 := [mp #10076 #10097]: #10077
+#10377 := [unit-resolution #10098 #8861]: #6009
+#10383 := [unit-resolution #10377 #10359]: #6006
+#10564 := (or #6003 #6040)
+#10442 := [hypothesis]: #10441
+#10074 := (not #5981)
+#10558 := [hypothesis]: #6006
+#10099 := (or #6003 #10074)
+#10100 := [def-axiom]: #10099
+#10559 := [unit-resolution #10100 #10558]: #10074
+#10101 := (not #5998)
+#10102 := (or #6003 #10101)
+#10103 := [def-axiom]: #10102
+#10560 := [unit-resolution #10103 #10558]: #10101
+#6043 := (or #5981 #5998 #6040)
+#10108 := (or #3924 #5981 #5998 #6040)
+#6032 := (+ #5978 #4581)
+#6033 := (+ #190 #6032)
+#6034 := (= #6033 0::Int)
+#6035 := (or #5981 #5984 #6034)
+#10109 := (or #3924 #6035)
+#10118 := (iff #10109 #10108)
+#10110 := (or #3924 #6043)
+#10113 := (iff #10110 #10108)
+#10114 := [rewrite]: #10113
+#10111 := (iff #10109 #10110)
+#6044 := (iff #6035 #6043)
+#6041 := (iff #6034 #6040)
+#6038 := (= #6033 #6037)
+#6039 := [rewrite]: #6038
+#6042 := [monotonicity #6039]: #6041
+#6045 := [monotonicity #6002 #6042]: #6044
+#10112 := [monotonicity #6045]: #10111
+#10119 := [trans #10112 #10114]: #10118
+#10104 := [quant-inst #2064]: #10109
+#10117 := [mp #10104 #10119]: #10108
+#10561 := [unit-resolution #10117 #8832]: #6043
+#10563 := [unit-resolution #10561 #10560 #10559 #10442]: false
+#10565 := [lemma #10563]: #10564
+#10382 := [unit-resolution #10565 #10383 #10378]: false
+#10471 := [lemma #10382]: #10467
+#10423 := [unit-resolution #10471 #10405 #10716 #10535]: #10570
+#10738 := (or #10731 #10684)
+#10737 := [th-lemma arith triangle-eq]: #10738
+#10531 := [unit-resolution #10737 #10423]: #10731
+#10368 := (or #10051 #10683)
+#10401 := [hypothesis]: #5770
+#10400 := [monotonicity #10401]: #10683
+#10370 := [hypothesis]: #10731
+#10402 := [unit-resolution #10370 #10400]: false
+#10369 := [lemma #10402]: #10368
+#10426 := [unit-resolution #10369 #10531]: #10051
+#10070 := (not #5778)
+#10071 := (or #10070 #5770 #5772)
+#10072 := [def-axiom]: #10071
+#10427 := [unit-resolution #10072 #10426 #10536]: #5772
+#10067 := (not #5772)
+#10803 := (or #10067 #10727)
+#10782 := (not #10727)
+#10800 := [hypothesis]: #10782
+#10801 := [hypothesis]: #5772
+#10785 := (or #3820 #188 #10067 #10727)
+#10728 := (or #188 #10067 #10727)
+#10786 := (or #3820 #10728)
+#10788 := (iff #10786 #10785)
+#10789 := [rewrite]: #10788
+#10787 := [quant-inst #2065 #186]: #10786
+#10790 := [mp #10787 #10789]: #10785
+#10802 := [unit-resolution #10790 #8132 #8131 #10801 #10800]: false
+#10804 := [lemma #10802]: #10803
+#10428 := [unit-resolution #10804 #10427]: #10727
+#10429 := (not #5699)
+#10276 := (or #10782 #2077 #10719 #10503 #10429)
+#10278 := [th-lemma arith assign-bounds 1 1 1 1]: #10276
+#10279 := [unit-resolution #10278 #10428 #10770 #10405 #10716]: #10719
+#10280 := [unit-resolution #10421 #10279]: #10441
+#10275 := [unit-resolution #10565 #10280]: #6003
+#10281 := [unit-resolution #10377 #10275]: #5987
+#10277 := [unit-resolution #10381 #10281]: #10107
+#10282 := [th-lemma arith farkas -1 -1 1 1 1 #10277 #10428 #10716 #10770 #10535]: false
+#10345 := [lemma #10282]: #10344
+#11340 := [unit-resolution #10345 #6061 #6060]: false
+#11363 := [lemma #11340]: #2954
+#3597 := (or #4006 #2959 #4000)
+#3598 := [def-axiom]: #3597
+#14089 := [unit-resolution #3598 #11363]: #14245
+#14298 := [unit-resolution #14089 #14006]: #4000
+#3606 := (or #3997 #3991)
+#3607 := [def-axiom]: #3606
+#17543 := [unit-resolution #3607 #14298]: #3991
+#16954 := [hypothesis]: #3005
+#3619 := (or #3000 #2101)
+#3622 := [def-axiom]: #3619
+#16970 := [unit-resolution #3622 #16954]: #2101
+#6572 := (f5 #200 ?v0!19)
+#6570 := (f15 #6572)
+#6573 := (* -1::Int #6570)
+#16948 := (+ #2097 #6573)
+#16952 := (>= #16948 0::Int)
+#16938 := (= #2097 #6570)
+#17002 := (= #2096 #6572)
+#16996 := (= #2095 #200)
+#5496 := (= ?v1!18 f28)
+#5497 := (f9 f21 ?v1!18)
+#5498 := (= #5497 f1)
+#6712 := (not #5498)
+#6463 := (f19 f20 ?v0!19)
+#6534 := (* -1::Int #6463)
+#5451 := (f19 f20 ?v1!18)
+#6728 := (+ #5451 #6534)
+#6729 := (+ #2097 #6728)
+#6730 := (>= #6729 0::Int)
+#16957 := (not #6730)
+#3507 := (not #2109)
+#3522 := (or #3000 #3507)
+#3524 := [def-axiom]: #3522
+#16955 := [unit-resolution #3524 #16954]: #3507
+#5548 := (* -1::Int #5451)
+#5549 := (+ #2104 #5548)
+#12637 := (>= #5549 0::Int)
+#5469 := (= #2104 #5451)
+#3523 := (or #3000 #2094)
+#3618 := [def-axiom]: #3523
+#16956 := [unit-resolution #3618 #16954]: #2094
+#16334 := (or #3949 #2985 #5469)
+#5472 := (or #2985 #5469)
+#16340 := (or #3949 #5472)
+#16341 := (iff #16340 #16334)
+#16068 := [rewrite]: #16341
+#16335 := [quant-inst #2091]: #16340
+#16150 := [mp #16335 #16068]: #16334
+#16959 := [unit-resolution #16150 #18772 #16956]: #5469
+#16963 := (not #5469)
+#16964 := (or #16963 #12637)
+#16966 := [th-lemma arith triangle-eq]: #16964
+#16967 := [unit-resolution #16966 #16959]: #12637
+#16961 := (not #12637)
+#12548 := (or #16957 #16961 #2109)
+#6535 := (+ #2105 #6534)
+#6536 := (<= #6535 0::Int)
+#18328 := (not #6536)
+#18338 := [hypothesis]: #18328
+#17088 := (or #3940 #6536)
+#6478 := (+ #6463 #2106)
+#6488 := (>= #6478 0::Int)
+#17089 := (or #3940 #6488)
+#18276 := (iff #17089 #17088)
+#18278 := (iff #17088 #17088)
+#18289 := [rewrite]: #18278
+#6557 := (iff #6488 #6536)
+#6529 := (+ #2106 #6463)
+#6532 := (>= #6529 0::Int)
+#6537 := (iff #6532 #6536)
+#6556 := [rewrite]: #6537
+#6527 := (iff #6488 #6532)
+#6530 := (= #6478 #6529)
+#6531 := [rewrite]: #6530
+#6533 := [monotonicity #6531]: #6527
+#6558 := [trans #6533 #6556]: #6557
+#18277 := [monotonicity #6558]: #18276
+#18290 := [trans #18277 #18289]: #18276
+#18275 := [quant-inst #2092]: #17089
+#18291 := [mp #18275 #18290]: #17088
+#18339 := [unit-resolution #18291 #9894 #18338]: false
+#18340 := [lemma #18339]: #6536
+#16962 := (or #16957 #16961 #18328 #2109)
+#16968 := [th-lemma arith assign-bounds 1 -1 -1]: #16962
+#12549 := [unit-resolution #16968 #18340]: #12548
+#12627 := [unit-resolution #12549 #16967 #16955]: #16957
+#16932 := (or #6712 #6730)
+#16777 := (or #3828 #6712 #2100 #6730)
+#6731 := (or #6712 #2100 #6730)
+#16782 := (or #3828 #6731)
+#16783 := (iff #16782 #16777)
+#16775 := [rewrite]: #16783
+#16780 := [quant-inst #2092 #2091]: #16782
+#16784 := [mp #16780 #16775]: #16777
+#16975 := [unit-resolution #16784 #7100 #16970]: #16932
+#12604 := [unit-resolution #16975 #12627]: #6712
+#5504 := (or #5496 #5498)
+#5486 := (f9 #198 ?v1!18)
+#5487 := (= #5486 f1)
+#5509 := (iff #5487 #5504)
+#16182 := (or #7628 #5509)
+#5499 := (if #5496 #4146 #5498)
+#5500 := (iff #5487 #5499)
+#16247 := (or #7628 #5500)
+#16275 := (iff #16247 #16182)
+#16296 := (iff #16182 #16182)
+#15136 := [rewrite]: #16296
+#5510 := (iff #5500 #5509)
+#5507 := (iff #5499 #5504)
+#5501 := (if #5496 true #5498)
+#5505 := (iff #5501 #5504)
+#5506 := [rewrite]: #5505
+#5502 := (iff #5499 #5501)
+#5503 := [monotonicity #4149]: #5502
+#5508 := [trans #5503 #5506]: #5507
+#5511 := [monotonicity #5508]: #5510
+#15135 := [monotonicity #5511]: #16275
+#15137 := [trans #15135 #15136]: #16275
+#16092 := [quant-inst #115 #186 #3 #2091]: #16247
+#16292 := [mp #16092 #15137]: #16182
+#16991 := [unit-resolution #16292 #3723]: #5509
+#16993 := (= #5486 #2093)
+#16994 := [monotonicity #9297]: #16993
+#16974 := [trans #16994 #16956]: #5487
+#16337 := (not #5487)
+#15281 := (not #5509)
+#16421 := (or #15281 #16337 #5504)
+#16336 := [def-axiom]: #16421
+#16992 := [unit-resolution #16336 #16974 #16991]: #5504
+#16342 := (not #5504)
+#16264 := (or #16342 #5496 #5498)
+#15134 := [def-axiom]: #16264
+#12567 := [unit-resolution #15134 #16992 #12604]: #5496
+#12652 := [monotonicity #12567]: #16996
+#12545 := [monotonicity #12652]: #17002
+#12613 := [monotonicity #12545]: #16938
+#16983 := (not #16938)
+#16973 := (or #16983 #16952)
+#16987 := [th-lemma arith triangle-eq]: #16973
+#12610 := [unit-resolution #16987 #12613]: #16952
+#6574 := (+ f14 #6573)
+#6575 := (<= #6574 0::Int)
+#6600 := (+ #6534 #6570)
+#6598 := (+ #190 #6600)
+#6601 := (>= #6598 0::Int)
+#12029 := (not #6601)
+#6452 := (+ #2104 #4168)
+#6453 := (>= #6452 0::Int)
+#5573 := (+ #190 #5548)
+#5574 := (<= #5573 0::Int)
+#16786 := (or #3914 #5498 #5574)
+#5564 := (+ #5451 #1235)
+#5565 := (>= #5564 0::Int)
+#5566 := (or #5498 #5565)
+#16796 := (or #3914 #5566)
+#16822 := (iff #16796 #16786)
+#5579 := (or #5498 #5574)
+#16815 := (or #3914 #5579)
+#16827 := (iff #16815 #16786)
+#16821 := [rewrite]: #16827
+#16825 := (iff #16796 #16815)
+#5580 := (iff #5566 #5579)
+#5577 := (iff #5565 #5574)
+#5567 := (+ #1235 #5451)
+#5570 := (>= #5567 0::Int)
+#5575 := (iff #5570 #5574)
+#5576 := [rewrite]: #5575
+#5571 := (iff #5565 #5570)
+#5568 := (= #5564 #5567)
+#5569 := [rewrite]: #5568
+#5572 := [monotonicity #5569]: #5571
+#5578 := [trans #5572 #5576]: #5577
+#5581 := [monotonicity #5578]: #5580
+#16826 := [monotonicity #5581]: #16825
+#16820 := [trans #16826 #16821]: #16822
+#16817 := [quant-inst #2091]: #16796
+#16823 := [mp #16817 #16820]: #16786
+#12620 := [unit-resolution #16823 #8603 #12604]: #5574
+#16986 := (not #5574)
+#12638 := (or #6453 #16986 #16961)
+#16989 := (or #6453 #16986 #16961 #10393)
+#16990 := [th-lemma arith assign-bounds 1 -1 -1]: #16989
+#12605 := [unit-resolution #16990 #10752]: #12638
+#12619 := [unit-resolution #12605 #12620 #16967]: #6453
+#12634 := (not #16952)
+#12636 := (not #6453)
+#12622 := (or #12029 #18328 #2109 #12636 #12661 #12634)
+#12656 := [th-lemma arith assign-bounds -1 -1 1 -1 1]: #12622
+#12623 := [unit-resolution #12656 #12619 #18340 #18263 #16955 #12610]: #12029
+#6617 := (or #6575 #6601)
+#6699 := (+ #2106 #6570)
+#6700 := (+ #190 #6699)
+#6703 := (= #6700 0::Int)
+#11657 := (not #6703)
+#12480 := (>= #6700 0::Int)
+#17015 := (not #12480)
+#17017 := (or #3000 #17015)
+#16969 := [unit-resolution #16968 #16967 #18340 #16955]: #16957
+#16976 := [unit-resolution #16975 #16969]: #6712
+#16995 := [unit-resolution #15134 #16992 #16976]: #5496
+#16997 := [monotonicity #16995]: #16996
+#16981 := [monotonicity #16997]: #17002
+#16982 := [monotonicity #16981]: #16938
+#16980 := [unit-resolution #16987 #16982]: #16952
+#16988 := [unit-resolution #16823 #8603 #16976]: #5574
+#17010 := [unit-resolution #16990 #16967 #10752 #16988]: #6453
+#17013 := [hypothesis]: #12480
+#17014 := [th-lemma arith farkas 1 -1 1 -1 1 #17013 #16955 #17010 #18263 #16980]: false
+#17018 := [lemma #17014]: #17017
+#12641 := [unit-resolution #17018 #16954]: #17015
+#11714 := (or #11657 #12480)
+#12635 := [th-lemma arith triangle-eq]: #11714
+#12017 := [unit-resolution #12635 #12641]: #11657
+#12092 := (or #3924 #6575 #6601 #6703)
+#6696 := (+ #6570 #2106)
+#6697 := (+ #190 #6696)
+#6698 := (= #6697 0::Int)
+#6580 := (+ #1235 #6573)
+#6581 := (+ #6463 #6580)
+#6579 := (<= #6581 0::Int)
+#6693 := (or #6575 #6579 #6698)
+#12098 := (or #3924 #6693)
+#12143 := (iff #12098 #12092)
+#5162 := (or #6575 #6601 #6703)
+#12058 := (or #3924 #5162)
+#12141 := (iff #12058 #12092)
+#12147 := [rewrite]: #12141
+#12060 := (iff #12098 #12058)
+#5352 := (iff #6693 #5162)
+#5122 := (iff #6698 #6703)
+#6701 := (= #6697 #6700)
+#6702 := [rewrite]: #6701
+#5123 := [monotonicity #6702]: #5122
+#6618 := (iff #6579 #6601)
+#6588 := (+ #6463 #6573)
+#6589 := (+ #1235 #6588)
+#6591 := (<= #6589 0::Int)
+#6602 := (iff #6591 #6601)
+#6603 := [rewrite]: #6602
+#6592 := (iff #6579 #6591)
+#6587 := (= #6581 #6589)
+#6590 := [rewrite]: #6587
+#6599 := [monotonicity #6590]: #6592
+#6619 := [trans #6599 #6603]: #6618
+#5353 := [monotonicity #6619 #5123]: #5352
+#12148 := [monotonicity #5353]: #12060
+#12142 := [trans #12148 #12147]: #12143
+#12130 := [quant-inst #2092]: #12098
+#12144 := [mp #12130 #12142]: #12092
+#12644 := [unit-resolution #12144 #8832 #12017]: #6617
+#11750 := [unit-resolution #12644 #12623]: #6575
+#12020 := [th-lemma arith farkas -1 1 1 #11750 #12610 #16970]: false
+#12047 := [lemma #12020]: #3000
+#3615 := (or #3994 #3005 #3988)
+#3616 := [def-axiom]: #3615
+#17550 := [unit-resolution #3616 #12047 #17543]: #3988
+#3620 := (or #3985 #3977)
+#3624 := [def-axiom]: #3620
+#14729 := [unit-resolution #3624 #17550]: #3977
+#4520 := (not #4519)
+#9879 := (or #3982 #5318 #4520 #5337)
+#5308 := (+ #4167 #5267)
+#5309 := (= #5308 0::Int)
+#5310 := (not #5309)
+#5298 := (+ #4167 #2127)
+#5299 := (>= #5298 0::Int)
+#5311 := (or #5299 #4520 #5310)
+#9846 := (or #3982 #5311)
+#3396 := (iff #9846 #9879)
+#5340 := (or #5318 #4520 #5337)
+#9859 := (or #3982 #5340)
+#8772 := (iff #9859 #9879)
+#8417 := [rewrite]: #8772
+#7921 := (iff #9846 #9859)
+#5341 := (iff #5311 #5340)
+#5338 := (iff #5310 #5337)
+#5335 := (iff #5309 #5332)
+#5323 := (+ #4167 #5211)
+#5324 := (+ #2127 #5323)
+#5327 := (= #5324 0::Int)
+#5333 := (iff #5327 #5332)
+#5334 := [rewrite]: #5333
+#5328 := (iff #5309 #5327)
+#5325 := (= #5308 #5324)
+#5326 := [rewrite]: #5325
+#5329 := [monotonicity #5326]: #5328
+#5336 := [trans #5329 #5334]: #5335
+#5339 := [monotonicity #5336]: #5338
+#5321 := (iff #5299 #5318)
+#5312 := (+ #2127 #4167)
+#5315 := (>= #5312 0::Int)
+#5319 := (iff #5315 #5318)
+#5320 := [rewrite]: #5319
+#5316 := (iff #5299 #5315)
+#5313 := (= #5298 #5312)
+#5314 := [rewrite]: #5313
+#5317 := [monotonicity #5314]: #5316
+#5322 := [trans #5317 #5320]: #5321
+#5342 := [monotonicity #5322 #5339]: #5341
+#7809 := [monotonicity #5342]: #7921
+#9848 := [trans #7809 #8417]: #3396
+#9851 := [quant-inst #186]: #9846
+#9882 := [mp #9851 #9848]: #9879
+#14657 := [unit-resolution #9882 #14729 #12665]: #15716
+#13744 := [unit-resolution #14657 #15984]: #5318
+#5158 := (>= #5157 0::Int)
+#15624 := (or #4623 #5158)
+#3612 := (or #3997 #3961)
+#3617 := [def-axiom]: #3612
+#14246 := [unit-resolution #3617 #14298]: #3961
+#11761 := (or #3966 #4623 #4520 #5158)
+#5159 := (or #4623 #4520 #5158)
+#10607 := (or #3966 #5159)
+#11378 := (iff #10607 #11761)
+#11392 := [rewrite]: #11378
+#10887 := [quant-inst #186 #2123]: #10607
+#11762 := [mp #10887 #11392]: #11761
+#15544 := [unit-resolution #11762 #14246 #12665]: #15624
+#13741 := [unit-resolution #15544 #15919]: #5158
+#13599 := [th-lemma arith eq-propagate 1 1 -1 -1 1 1 #13741 #13744 #10752 #18263 #15545 #15571]: #19614
+#13366 := (not #19614)
+#15910 := (or #13366 #19610)
+#15852 := [th-lemma arith triangle-eq]: #15910
+#12478 := [unit-resolution #15852 #13599]: #19610
+#19611 := (not #19610)
+#19612 := (or #19609 #19611)
+#12798 := (or #6342 #19609 #19611)
+#14603 := (or #6342 #19612)
+#14654 := (iff #14603 #12798)
+#14656 := [rewrite]: #14654
+#14666 := [quant-inst #186 #2123]: #14603
+#14616 := [mp #14666 #14656]: #12798
+#13305 := [unit-resolution #14616 #3735]: #19612
+#16503 := [unit-resolution #13305 #12478]: #19609
+#15534 := [unit-resolution #16503 #15258]: false
+#16669 := [lemma #15534]: #5176
+#19122 := (or #19098 #17503)
+#17519 := (not #17503)
+#19187 := [hypothesis]: #17519
+#19192 := [hypothesis]: #5176
+#19097 := [th-lemma arith triangle-eq]: #19122
+#19172 := [unit-resolution #19097 #19192 #19187]: false
+#19096 := [lemma #19172]: #19122
+#24633 := [unit-resolution #19096 #16669]: #17503
+#17541 := (or #17547 #17519)
+#3530 := (or #3985 #2130)
+#3623 := [def-axiom]: #3530
+#17522 := [unit-resolution #3623 #17550]: #2130
+#17561 := [hypothesis]: #17503
+#17545 := [hypothesis]: #5430
+#17548 := [th-lemma arith farkas -1 1 1 #17545 #17561 #17522]: false
+#17521 := [lemma #17548]: #17541
+#24634 := [unit-resolution #17521 #24633]: #17547
+#21003 := (or #5430 #5448)
+#3528 := (or #3985 #2125)
+#3529 := [def-axiom]: #3528
+#21002 := [unit-resolution #3529 #17550]: #2125
+#11823 := (or #3836 #2124 #5430 #5448)
+#5449 := (or #2124 #5430 #5448)
+#11543 := (or #3836 #5449)
+#11822 := (iff #11543 #11823)
+#11793 := [rewrite]: #11822
+#11827 := [quant-inst #2123]: #11543
+#11829 := [mp #11827 #11793]: #11823
+#21004 := [unit-resolution #11829 #7150 #21002]: #21003
+#24635 := [unit-resolution #21004 #24634]: #5448
+#10444 := (or #5447 #5445)
+#11271 := [def-axiom]: #10444
+#20098 := [unit-resolution #11271 #24635]: #5445
+#20099 := (or #5446 #11830)
+#20097 := [th-lemma arith triangle-eq]: #20099
+#20088 := [unit-resolution #20097 #20098]: #11830
+#18274 := (+ #5432 #17799)
+#18321 := (>= #18274 0::Int)
+#15420 := (or #3940 #18321)
+#15369 := [quant-inst #5431]: #15420
+#19872 := [unit-resolution #15369 #9894]: #18321
+#20100 := (not #18321)
+#19873 := (not #11830)
+#20101 := (or #15418 #19873 #17519 #20100)
+#20102 := [th-lemma arith assign-bounds -1 -1 -1]: #20101
+#20082 := [unit-resolution #20102 #24633 #19872 #20088]: #15418
+#18473 := (<= #18529 0::Int)
+#18450 := (+ f14 #5442)
+#18526 := (<= #18450 0::Int)
+#19955 := (not #18526)
+#18451 := (>= #5432 0::Int)
+#15358 := (or #3811 #18451)
+#15465 := [quant-inst #5431]: #15358
+#21016 := [unit-resolution #15465 #8408]: #18451
+#21023 := (not #18451)
+#19947 := (or #19955 #2129 #17519 #21023 #19873)
+#19948 := [th-lemma arith assign-bounds -1 -1 -1 -1]: #19947
+#19989 := [unit-resolution #19948 #21016 #24633 #20088 #17522]: #19955
+#18253 := (f9 f29 #5431)
+#17861 := (= #18253 f1)
+#24578 := (f9 #198 #5431)
+#24579 := (= #24578 f1)
+#24590 := (= #5431 f28)
+#24596 := (or #24590 #5437)
+#24601 := (iff #24579 #24596)
+#24604 := (or #7628 #24601)
+#24591 := (if #24590 #4146 #5437)
+#24592 := (iff #24579 #24591)
+#24605 := (or #7628 #24592)
+#24607 := (iff #24605 #24604)
+#24609 := (iff #24604 #24604)
+#24610 := [rewrite]: #24609
+#24602 := (iff #24592 #24601)
+#24599 := (iff #24591 #24596)
+#24593 := (if #24590 true #5437)
+#24597 := (iff #24593 #24596)
+#24598 := [rewrite]: #24597
+#24594 := (iff #24591 #24593)
+#24595 := [monotonicity #4149]: #24594
+#24600 := [trans #24595 #24598]: #24599
+#24603 := [monotonicity #24600]: #24602
+#24608 := [monotonicity #24603]: #24607
+#24611 := [trans #24608 #24610]: #24607
+#24606 := [quant-inst #115 #186 #3 #5431]: #24605
+#24612 := [mp #24606 #24611]: #24604
+#24632 := [unit-resolution #24612 #3723]: #24601
+#24621 := (not #24601)
+#24638 := (or #24621 #24579)
+#12006 := (or #5447 #5437)
+#12000 := [def-axiom]: #12006
+#24636 := [unit-resolution #12000 #24635]: #5437
+#24616 := (or #24596 #5438)
+#24617 := [def-axiom]: #24616
+#24637 := [unit-resolution #24617 #24636]: #24596
+#24618 := (not #24596)
+#24622 := (or #24621 #24579 #24618)
+#24623 := [def-axiom]: #24622
+#24639 := [unit-resolution #24623 #24637]: #24638
+#24640 := [unit-resolution #24639 #24632]: #24579
+#24641 := (= #18253 #24578)
+#24642 := [monotonicity #9701]: #24641
+#24643 := [trans #24642 #24640]: #17861
+#17862 := (not #17861)
+#24631 := [hypothesis]: #17862
+#24644 := [unit-resolution #24631 #24643]: false
+#24645 := [lemma #24644]: #17861
+#12615 := (or #17862 #18526 #18473)
+#3526 := (or #3985 #3969)
+#3527 := [def-axiom]: #3526
+#22006 := [unit-resolution #3527 #17550]: #3969
+#16736 := (or #3974 #17862 #18526 #18473)
+#18522 := (+ #18254 #2127)
+#18507 := (+ #5441 #18522)
+#18515 := (>= #18507 0::Int)
+#18516 := (or #17862 #18526 #18515)
+#16729 := (or #3974 #18516)
+#15322 := (iff #16729 #16736)
+#16781 := (or #3974 #12615)
+#16766 := (iff #16781 #16736)
+#16787 := [rewrite]: #16766
+#16771 := (iff #16729 #16781)
+#13633 := (iff #18516 #12615)
+#14647 := (iff #18515 #18473)
+#18513 := (+ #5441 #18254)
+#18523 := (+ #2127 #18513)
+#18470 := (>= #18523 0::Int)
+#18472 := (iff #18470 #18473)
+#18530 := [rewrite]: #18472
+#18514 := (iff #18515 #18470)
+#18525 := (= #18507 #18523)
+#18520 := [rewrite]: #18525
+#18527 := [monotonicity #18520]: #18514
+#12281 := [trans #18527 #18530]: #14647
+#14300 := [monotonicity #12281]: #13633
+#15106 := [monotonicity #14300]: #16771
+#16805 := [trans #15106 #16787]: #15322
+#16779 := [quant-inst #2123 #5431]: #16729
+#15127 := [mp #16779 #16805]: #16736
+#20083 := [unit-resolution #15127 #22006]: #12615
+#19946 := [unit-resolution #20083 #24645 #19989]: #18473
+#18690 := (= #18529 0::Int)
+#18704 := (not #18690)
+#18651 := (+ #2126 #17799)
+#18655 := (<= #18651 0::Int)
+#18872 := (not #18655)
+#11828 := (not #5435)
+#11949 := (or #5447 #11828)
+#11796 := [def-axiom]: #11949
+#19950 := [unit-resolution #11796 #24635]: #11828
+#18977 := (or #18872 #5435 #17519 #20100)
+#18978 := [th-lemma arith assign-bounds -1 -1 -1]: #18977
+#18953 := [unit-resolution #18978 #24633 #19872 #19950]: #18872
+#18707 := (or #18655 #17862 #18704)
+#16901 := (or #3982 #18655 #17862 #18704)
+#18650 := (+ #2127 #5441)
+#18592 := (+ #18254 #18650)
+#18638 := (= #18592 0::Int)
+#18644 := (not #18638)
+#18662 := (>= #18522 0::Int)
+#18615 := (or #18662 #17862 #18644)
+#16863 := (or #3982 #18615)
+#12904 := (iff #16863 #16901)
+#16890 := (or #3982 #18707)
+#16885 := (iff #16890 #16901)
+#15516 := [rewrite]: #16885
+#15161 := (iff #16863 #16890)
+#18708 := (iff #18615 #18707)
+#18692 := (iff #18644 #18704)
+#18691 := (iff #18638 #18690)
+#18658 := (= #18523 0::Int)
+#18654 := (iff #18658 #18690)
+#18693 := [rewrite]: #18654
+#18673 := (iff #18638 #18658)
+#18649 := (= #18592 #18523)
+#18653 := [rewrite]: #18649
+#18674 := [monotonicity #18653]: #18673
+#18694 := [trans #18674 #18693]: #18691
+#18689 := [monotonicity #18694]: #18692
+#18659 := (iff #18662 #18655)
+#18581 := (+ #2127 #18254)
+#18641 := (>= #18581 0::Int)
+#18666 := (iff #18641 #18655)
+#18665 := [rewrite]: #18666
+#18646 := (iff #18662 #18641)
+#18645 := (= #18522 #18581)
+#18596 := [rewrite]: #18645
+#18647 := [monotonicity #18596]: #18646
+#18664 := [trans #18647 #18665]: #18659
+#18709 := [monotonicity #18664 #18689]: #18708
+#15308 := [monotonicity #18709]: #15161
+#16914 := [trans #15308 #15516]: #12904
+#16811 := [quant-inst #5431]: #16863
+#15485 := [mp #16811 #16914]: #16901
+#19077 := [unit-resolution #15485 #14729]: #18707
+#19169 := [unit-resolution #19077 #24645 #18953]: #18704
+#19090 := (not #15418)
+#19243 := (not #18473)
+#19996 := (or #18690 #19243 #19090)
+#19953 := [th-lemma arith triangle-eq]: #19996
+[unit-resolution #19953 #19169 #19946 #20082]: false
+unsat
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/SMT_Examples/Boogie_Max.b2i	Tue Jul 23 18:36:23 2013 +0200
@@ -0,0 +1,748 @@
+vc max 1
+    implies
+    label pos 10 7
+    true
+    implies
+    <
+    int-num 0
+    var length
+      int
+    implies
+    true
+    implies
+    =
+    var max@0
+      int
+    select 2
+    var array
+      array 2
+        int
+        int
+    int-num 0
+    implies
+    and 4
+    <=
+    int-num 0
+    int-num 0
+    <=
+    int-num 0
+    int-num 0
+    <=
+    int-num 1
+    int-num 1
+    <=
+    int-num 1
+    int-num 1
+    and 2
+    label neg 14 5
+    forall 1 0 3
+      var i
+        int
+      attribute qid 1
+        string-attr BoogieMa.14:23
+      attribute uniqueId 1
+        string-attr 2
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 2
+    <=
+    int-num 0
+    var i
+      int
+    <
+    var i
+      int
+    int-num 1
+    <=
+    select 2
+    var array
+      array 2
+        int
+        int
+    var i
+      int
+    var max@0
+      int
+    implies
+    forall 1 0 3
+      var i
+        int
+      attribute qid 1
+        string-attr BoogieMa.14:23
+      attribute uniqueId 1
+        string-attr 2
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 2
+    <=
+    int-num 0
+    var i
+      int
+    <
+    var i
+      int
+    int-num 1
+    <=
+    select 2
+    var array
+      array 2
+        int
+        int
+    var i
+      int
+    var max@0
+      int
+    and 2
+    label neg 15 5
+    =
+    select 2
+    var array
+      array 2
+        int
+        int
+    int-num 0
+    var max@0
+      int
+    implies
+    =
+    select 2
+    var array
+      array 2
+        int
+        int
+    int-num 0
+    var max@0
+      int
+    implies
+    label pos 13 3
+    true
+    implies
+    and 2
+    <=
+    int-num 0
+    var k@0
+      int
+    <=
+    int-num 1
+    var p@0
+      int
+    implies
+    forall 1 0 3
+      var i
+        int
+      attribute qid 1
+        string-attr BoogieMa.14:23
+      attribute uniqueId 1
+        string-attr 2
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 2
+    <=
+    int-num 0
+    var i
+      int
+    <
+    var i
+      int
+    var p@0
+      int
+    <=
+    select 2
+    var array
+      array 2
+        int
+        int
+    var i
+      int
+    var max@1
+      int
+    implies
+    =
+    select 2
+    var array
+      array 2
+        int
+        int
+    var k@0
+      int
+    var max@1
+      int
+    implies
+    and 2
+    <=
+    int-num 0
+    var k@0
+      int
+    <=
+    int-num 1
+    var p@0
+      int
+    and 2
+    implies
+    label pos 13 3
+    true
+    implies
+    and 2
+    <=
+    int-num 0
+    var k@0
+      int
+    <=
+    int-num 1
+    var p@0
+      int
+    implies
+    >=
+    var p@0
+      int
+    var length
+      int
+    implies
+    and 2
+    <=
+    int-num 0
+    var k@0
+      int
+    <=
+    int-num 1
+    var p@0
+      int
+    implies
+    label pos 0 0
+    true
+    implies
+    =
+    var k@2
+      int
+    var k@0
+      int
+    implies
+    =
+    var max@4
+      int
+    var max@1
+      int
+    implies
+    =
+    var p@2
+      int
+    var p@0
+      int
+    implies
+    label pos 0 0
+    true
+    and 2
+    label neg 5 3
+    exists 1 0 3
+      var i
+        int
+      attribute qid 1
+        string-attr BoogieMa.5:19
+      attribute uniqueId 1
+        string-attr 1
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 2
+    <=
+    int-num 0
+    var i
+      int
+    <
+    var i
+      int
+    var length
+      int
+    =
+    select 2
+    var array
+      array 2
+        int
+        int
+    var i
+      int
+    var max@4
+      int
+    implies
+    exists 1 0 3
+      var i
+        int
+      attribute qid 1
+        string-attr BoogieMa.5:19
+      attribute uniqueId 1
+        string-attr 1
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 2
+    <=
+    int-num 0
+    var i
+      int
+    <
+    var i
+      int
+    var length
+      int
+    =
+    select 2
+    var array
+      array 2
+        int
+        int
+    var i
+      int
+    var max@4
+      int
+    and 2
+    label neg 4 3
+    forall 1 0 3
+      var i
+        int
+      attribute qid 1
+        string-attr BoogieMa.4:19
+      attribute uniqueId 1
+        string-attr 0
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 2
+    <=
+    int-num 0
+    var i
+      int
+    <
+    var i
+      int
+    var length
+      int
+    <=
+    select 2
+    var array
+      array 2
+        int
+        int
+    var i
+      int
+    var max@4
+      int
+    implies
+    forall 1 0 3
+      var i
+        int
+      attribute qid 1
+        string-attr BoogieMa.4:19
+      attribute uniqueId 1
+        string-attr 0
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 2
+    <=
+    int-num 0
+    var i
+      int
+    <
+    var i
+      int
+    var length
+      int
+    <=
+    select 2
+    var array
+      array 2
+        int
+        int
+    var i
+      int
+    var max@4
+      int
+    true
+    implies
+    label pos 17 5
+    true
+    implies
+    and 2
+    <=
+    int-num 0
+    var k@0
+      int
+    <=
+    int-num 1
+    var p@0
+      int
+    implies
+    <
+    var p@0
+      int
+    var length
+      int
+    implies
+    and 2
+    <=
+    int-num 0
+    var k@0
+      int
+    <=
+    int-num 1
+    var p@0
+      int
+    and 2
+    implies
+    label pos 17 31
+    true
+    implies
+    and 2
+    <=
+    int-num 0
+    var k@0
+      int
+    <=
+    int-num 1
+    var p@0
+      int
+    implies
+    >
+    select 2
+    var array
+      array 2
+        int
+        int
+    var p@0
+      int
+    var max@1
+      int
+    implies
+    =
+    var max@2
+      int
+    select 2
+    var array
+      array 2
+        int
+        int
+    var p@0
+      int
+    implies
+    and 2
+    <=
+    int-num 1
+    var p@0
+      int
+    <=
+    int-num 1
+    var p@0
+      int
+    implies
+    label pos 0 0
+    true
+    implies
+    =
+    var k@1
+      int
+    var p@0
+      int
+    implies
+    =
+    var max@3
+      int
+    var max@2
+      int
+    implies
+    label pos 18 7
+    true
+    implies
+    and 2
+    <=
+    int-num 0
+    var k@1
+      int
+    <=
+    int-num 1
+    var p@0
+      int
+    implies
+    =
+    var p@1
+      int
+    +
+    var p@0
+      int
+    int-num 1
+    implies
+    and 2
+    <=
+    int-num 0
+    var k@1
+      int
+    <=
+    int-num 2
+    var p@1
+      int
+    implies
+    label pos 0 0
+    true
+    and 2
+    label neg 14 5
+    forall 1 0 3
+      var i
+        int
+      attribute qid 1
+        string-attr BoogieMa.14:23
+      attribute uniqueId 1
+        string-attr 2
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 2
+    <=
+    int-num 0
+    var i
+      int
+    <
+    var i
+      int
+    var p@1
+      int
+    <=
+    select 2
+    var array
+      array 2
+        int
+        int
+    var i
+      int
+    var max@3
+      int
+    implies
+    forall 1 0 3
+      var i
+        int
+      attribute qid 1
+        string-attr BoogieMa.14:23
+      attribute uniqueId 1
+        string-attr 2
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 2
+    <=
+    int-num 0
+    var i
+      int
+    <
+    var i
+      int
+    var p@1
+      int
+    <=
+    select 2
+    var array
+      array 2
+        int
+        int
+    var i
+      int
+    var max@3
+      int
+    and 2
+    label neg 15 5
+    =
+    select 2
+    var array
+      array 2
+        int
+        int
+    var k@1
+      int
+    var max@3
+      int
+    implies
+    =
+    select 2
+    var array
+      array 2
+        int
+        int
+    var k@1
+      int
+    var max@3
+      int
+    implies
+    false
+    true
+    implies
+    label pos 17 5
+    true
+    implies
+    and 2
+    <=
+    int-num 0
+    var k@0
+      int
+    <=
+    int-num 1
+    var p@0
+      int
+    implies
+    <=
+    select 2
+    var array
+      array 2
+        int
+        int
+    var p@0
+      int
+    var max@1
+      int
+    implies
+    and 2
+    <=
+    int-num 0
+    var k@0
+      int
+    <=
+    int-num 1
+    var p@0
+      int
+    implies
+    label pos 0 0
+    true
+    implies
+    =
+    var k@1
+      int
+    var k@0
+      int
+    implies
+    =
+    var max@3
+      int
+    var max@1
+      int
+    implies
+    label pos 18 7
+    true
+    implies
+    and 2
+    <=
+    int-num 0
+    var k@1
+      int
+    <=
+    int-num 1
+    var p@0
+      int
+    implies
+    =
+    var p@1
+      int
+    +
+    var p@0
+      int
+    int-num 1
+    implies
+    and 2
+    <=
+    int-num 0
+    var k@1
+      int
+    <=
+    int-num 2
+    var p@1
+      int
+    implies
+    label pos 0 0
+    true
+    and 2
+    label neg 14 5
+    forall 1 0 3
+      var i
+        int
+      attribute qid 1
+        string-attr BoogieMa.14:23
+      attribute uniqueId 1
+        string-attr 2
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 2
+    <=
+    int-num 0
+    var i
+      int
+    <
+    var i
+      int
+    var p@1
+      int
+    <=
+    select 2
+    var array
+      array 2
+        int
+        int
+    var i
+      int
+    var max@3
+      int
+    implies
+    forall 1 0 3
+      var i
+        int
+      attribute qid 1
+        string-attr BoogieMa.14:23
+      attribute uniqueId 1
+        string-attr 2
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 2
+    <=
+    int-num 0
+    var i
+      int
+    <
+    var i
+      int
+    var p@1
+      int
+    <=
+    select 2
+    var array
+      array 2
+        int
+        int
+    var i
+      int
+    var max@3
+      int
+    and 2
+    label neg 15 5
+    =
+    select 2
+    var array
+      array 2
+        int
+        int
+    var k@1
+      int
+    var max@3
+      int
+    implies
+    =
+    select 2
+    var array
+      array 2
+        int
+        int
+    var k@1
+      int
+    var max@3
+      int
+    implies
+    false
+    true
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/SMT_Examples/Boogie_Max.certs	Tue Jul 23 18:36:23 2013 +0200
@@ -0,0 +1,2151 @@
+2fe4278cefdf713c58d9bb38d221cd0a858eee31 2150 0
+#2 := false
+#7 := 0::Int
+decl ?v0!3 :: Int
+#1165 := ?v0!3
+#688 := -1::Int
+#1310 := (* -1::Int ?v0!3)
+decl f8 :: Int
+#31 := f8
+#2386 := (+ f8 #1310)
+#2387 := (<= #2386 0::Int)
+#2492 := (not #2387)
+#2419 := (>= #2386 0::Int)
+decl f16 :: Int
+#82 := f16
+#797 := (* -1::Int f16)
+#829 := (+ f8 #797)
+#1830 := (>= #829 -1::Int)
+#828 := (= #829 -1::Int)
+decl f15 :: Int
+#78 := f15
+decl f5 :: (-> S2 Int Int)
+decl f14 :: Int
+#76 := f14
+decl f6 :: S2
+#11 := f6
+#93 := (f5 f6 f14)
+#94 := (= #93 f15)
+#20 := (:var 0 Int)
+#24 := (f5 f6 #20)
+#2148 := (pattern #24)
+#808 := (* -1::Int f15)
+#809 := (+ #24 #808)
+#810 := (<= #809 0::Int)
+#798 := (+ #20 #797)
+#796 := (>= #798 0::Int)
+#677 := (>= #20 0::Int)
+#1413 := (not #677)
+#1587 := (or #1413 #796 #810)
+#2182 := (forall (vars (?v0 Int)) (:pat #2148) #1587)
+#2187 := (not #2182)
+#2190 := (or #2187 #94)
+#2193 := (not #2190)
+#1172 := (f5 f6 ?v0!3)
+#1332 := (* -1::Int #1172)
+#1333 := (+ f15 #1332)
+#1334 := (>= #1333 0::Int)
+#1311 := (+ f16 #1310)
+#1312 := (<= #1311 0::Int)
+#1166 := (>= ?v0!3 0::Int)
+#1550 := (not #1166)
+#1565 := (or #1550 #1312 #1334)
+#1570 := (not #1565)
+#2196 := (or #1570 #2193)
+#2199 := (not #2196)
+#85 := 2::Int
+#788 := (>= f16 2::Int)
+#1612 := (not #788)
+#785 := (>= f14 0::Int)
+#1611 := (not #785)
+#832 := (not #828)
+#15 := 1::Int
+#707 := (>= f8 1::Int)
+#841 := (not #707)
+decl f9 :: Int
+#36 := f9
+#115 := (= f15 f9)
+#453 := (not #115)
+decl f7 :: Int
+#29 := f7
+#114 := (= f14 f7)
+#462 := (not #114)
+#71 := (f5 f6 f8)
+#846 := (* -1::Int #71)
+#847 := (+ f9 #846)
+#845 := (>= #847 0::Int)
+#844 := (not #845)
+#704 := (>= f7 0::Int)
+#1542 := (not #704)
+#2208 := (or #1542 #844 #462 #453 #841 #832 #1611 #1612 #2199)
+#2211 := (not #2208)
+decl f13 :: Int
+#73 := f13
+#79 := (= f15 f13)
+#386 := (not #79)
+#77 := (= f14 f8)
+#395 := (not #77)
+#74 := (= f13 #71)
+#420 := (not #74)
+#2202 := (or #1542 #845 #420 #395 #386 #841 #832 #1611 #1612 #2199)
+#2205 := (not #2202)
+#2214 := (or #2205 #2211)
+#2217 := (not #2214)
+#753 := (* -1::Int f8)
+decl f3 :: Int
+#8 := f3
+#754 := (+ f3 #753)
+#755 := (<= #754 0::Int)
+#2220 := (or #1542 #841 #755 #2217)
+#2223 := (not #2220)
+decl ?v0!2 :: Int
+#1110 := ?v0!2
+#1118 := (f5 f6 ?v0!2)
+#1263 := (* -1::Int #1118)
+decl f11 :: Int
+#45 := f11
+#1264 := (+ f11 #1263)
+#1265 := (>= #1264 0::Int)
+#1112 := (* -1::Int ?v0!2)
+#1113 := (+ f3 #1112)
+#1114 := (<= #1113 0::Int)
+#1111 := (>= ?v0!2 0::Int)
+#1503 := (not #1111)
+decl ?v0!1 :: Int
+#1092 := ?v0!1
+#1100 := (f5 f6 ?v0!1)
+#1101 := (= #1100 f11)
+#1094 := (* -1::Int ?v0!1)
+#1095 := (+ f3 #1094)
+#1096 := (<= #1095 0::Int)
+#1093 := (>= ?v0!1 0::Int)
+#1483 := (not #1093)
+#1498 := (or #1483 #1096 #1101)
+#1529 := (not #1498)
+#1530 := (or #1529 #1503 #1114 #1265)
+#1531 := (not #1530)
+#51 := (= #24 f11)
+#715 := (* -1::Int #20)
+#716 := (+ f3 #715)
+#717 := (<= #716 0::Int)
+#1472 := (or #1413 #717 #51)
+#1477 := (not #1472)
+#2165 := (forall (vars (?v0 Int)) (:pat #2148) #1477)
+#2170 := (or #2165 #1531)
+#2173 := (not #2170)
+decl f12 :: Int
+#47 := f12
+#48 := (= f12 f8)
+#235 := (not #48)
+#46 := (= f11 f9)
+#244 := (not #46)
+decl f10 :: Int
+#43 := f10
+#44 := (= f10 f7)
+#253 := (not #44)
+#758 := (not #755)
+#2176 := (or #1542 #841 #758 #253 #244 #235 #2173)
+#2179 := (not #2176)
+#2226 := (or #2179 #2223)
+#2229 := (not #2226)
+#40 := (f5 f6 f7)
+#41 := (= #40 f9)
+#556 := (not #41)
+#953 := (* -1::Int f9)
+#954 := (+ #24 #953)
+#955 := (<= #954 0::Int)
+#943 := (+ #20 #753)
+#942 := (>= #943 0::Int)
+#1450 := (or #1413 #942 #955)
+#2157 := (forall (vars (?v0 Int)) (:pat #2148) #1450)
+#2162 := (not #2157)
+decl f4 :: Int
+#10 := f4
+#12 := (f5 f6 0::Int)
+#28 := (= #12 f4)
+#589 := (not #28)
+#2232 := (or #589 #1542 #841 #2162 #556 #2229)
+#2235 := (not #2232)
+#2238 := (or #589 #2235)
+#2241 := (not #2238)
+#691 := (* -1::Int #24)
+#692 := (+ f4 #691)
+#690 := (>= #692 0::Int)
+#680 := (>= #20 1::Int)
+#1428 := (or #1413 #680 #690)
+#2149 := (forall (vars (?v0 Int)) (:pat #2148) #1428)
+#2154 := (not #2149)
+#2244 := (or #2154 #2241)
+#2247 := (not #2244)
+decl ?v0!0 :: Int
+#1040 := ?v0!0
+#1034 := (f5 f6 ?v0!0)
+#1035 := (* -1::Int #1034)
+#1032 := (+ f4 #1035)
+#1033 := (>= #1032 0::Int)
+#1042 := (>= ?v0!0 1::Int)
+#1041 := (>= ?v0!0 0::Int)
+#1179 := (not #1041)
+#1405 := (or #1179 #1042 #1033)
+#1999 := (= f4 #1034)
+#1925 := (= #12 #1034)
+#1965 := (= #1034 #12)
+#1975 := (= ?v0!0 0::Int)
+#1043 := (not #1042)
+#1410 := (not #1405)
+#1998 := [hypothesis]: #1410
+#1735 := (or #1405 #1043)
+#1820 := [def-axiom]: #1735
+#2000 := [unit-resolution #1820 #1998]: #1043
+#1734 := (or #1405 #1041)
+#1819 := [def-axiom]: #1734
+#1968 := [unit-resolution #1819 #1998]: #1041
+#1934 := [th-lemma arith eq-propagate 0 0 #1968 #2000]: #1975
+#1967 := [monotonicity #1934]: #1965
+#1926 := [symm #1967]: #1925
+#13 := (= f4 #12)
+#799 := (not #796)
+#802 := (and #677 #799)
+#805 := (not #802)
+#813 := (or #805 #810)
+#816 := (forall (vars (?v0 Int)) #813)
+#819 := (not #816)
+#822 := (or #819 #94)
+#825 := (and #816 #822)
+#790 := (and #785 #788)
+#793 := (not #790)
+#835 := (and #785 #707)
+#838 := (not #835)
+#709 := (and #704 #707)
+#712 := (not #709)
+#908 := (or #712 #844 #462 #453 #838 #832 #793 #825)
+#884 := (or #712 #845 #420 #841 #395 #386 #838 #832 #793 #825)
+#913 := (and #884 #908)
+#934 := (or #712 #755 #913)
+#736 := (* -1::Int f11)
+#737 := (+ #24 #736)
+#738 := (<= #737 0::Int)
+#718 := (not #717)
+#721 := (and #677 #718)
+#724 := (not #721)
+#741 := (or #724 #738)
+#744 := (forall (vars (?v0 Int)) #741)
+#727 := (or #724 #51)
+#730 := (exists (vars (?v0 Int)) #727)
+#733 := (not #730)
+#747 := (or #733 #744)
+#750 := (and #730 #747)
+#779 := (or #712 #758 #253 #244 #235 #750)
+#939 := (and #779 #934)
+#944 := (not #942)
+#947 := (and #677 #944)
+#950 := (not #947)
+#958 := (or #950 #955)
+#961 := (forall (vars (?v0 Int)) #958)
+#964 := (not #961)
+#982 := (or #589 #712 #964 #556 #939)
+#987 := (and #28 #982)
+#678 := (not #680)
+#682 := (and #677 #678)
+#685 := (not #682)
+#694 := (or #685 #690)
+#697 := (forall (vars (?v0 Int)) #694)
+#700 := (not #697)
+#990 := (or #700 #987)
+#993 := (and #697 #990)
+#622 := (not #13)
+#996 := (<= f3 0::Int)
+#1016 := (or #996 #622 #993)
+#1021 := (not #1016)
+#1 := true
+#95 := (implies false true)
+#96 := (implies #94 #95)
+#97 := (and #94 #96)
+#90 := (<= #24 f15)
+#88 := (< #20 f16)
+#21 := (<= 0::Int #20)
+#89 := (and #21 #88)
+#91 := (implies #89 #90)
+#92 := (forall (vars (?v0 Int)) #91)
+#98 := (implies #92 #97)
+#99 := (and #92 #98)
+#100 := (implies true #99)
+#86 := (<= 2::Int f16)
+#80 := (<= 0::Int f14)
+#87 := (and #80 #86)
+#101 := (implies #87 #100)
+#83 := (+ f8 1::Int)
+#84 := (= f16 #83)
+#102 := (implies #84 #101)
+#32 := (<= 1::Int f8)
+#81 := (and #80 #32)
+#103 := (implies #81 #102)
+#104 := (implies true #103)
+#116 := (implies #115 #104)
+#117 := (implies #114 #116)
+#118 := (implies true #117)
+#30 := (<= 0::Int f7)
+#33 := (and #30 #32)
+#119 := (implies #33 #118)
+#113 := (<= #71 f9)
+#120 := (implies #113 #119)
+#121 := (implies #33 #120)
+#122 := (implies true #121)
+#105 := (implies #79 #104)
+#106 := (implies #77 #105)
+#107 := (implies true #106)
+#75 := (and #32 #32)
+#108 := (implies #75 #107)
+#109 := (implies #74 #108)
+#72 := (< f9 #71)
+#110 := (implies #72 #109)
+#111 := (implies #33 #110)
+#112 := (implies true #111)
+#123 := (and #112 #122)
+#124 := (implies #33 #123)
+#70 := (< f8 f3)
+#125 := (implies #70 #124)
+#126 := (implies #33 #125)
+#127 := (implies true #126)
+#54 := (<= #24 f11)
+#49 := (< #20 f3)
+#50 := (and #21 #49)
+#55 := (implies #50 #54)
+#56 := (forall (vars (?v0 Int)) #55)
+#57 := (implies #56 true)
+#58 := (and #56 #57)
+#52 := (implies #50 #51)
+#53 := (exists (vars (?v0 Int)) #52)
+#59 := (implies #53 #58)
+#60 := (and #53 #59)
+#61 := (implies true #60)
+#62 := (implies #48 #61)
+#63 := (implies #46 #62)
+#64 := (implies #44 #63)
+#65 := (implies true #64)
+#66 := (implies #33 #65)
+#42 := (<= f3 f8)
+#67 := (implies #42 #66)
+#68 := (implies #33 #67)
+#69 := (implies true #68)
+#128 := (and #69 #127)
+#129 := (implies #33 #128)
+#130 := (implies #41 #129)
+#37 := (<= #24 f9)
+#34 := (< #20 f8)
+#35 := (and #21 #34)
+#38 := (implies #35 #37)
+#39 := (forall (vars (?v0 Int)) #38)
+#131 := (implies #39 #130)
+#132 := (implies #33 #131)
+#133 := (implies true #132)
+#134 := (implies #28 #133)
+#135 := (and #28 #134)
+#25 := (<= #24 f4)
+#22 := (< #20 1::Int)
+#23 := (and #21 #22)
+#26 := (implies #23 #25)
+#27 := (forall (vars (?v0 Int)) #26)
+#136 := (implies #27 #135)
+#137 := (and #27 #136)
+#16 := (<= 1::Int 1::Int)
+#17 := (and #16 #16)
+#14 := (<= 0::Int 0::Int)
+#18 := (and #14 #17)
+#19 := (and #14 #18)
+#138 := (implies #19 #137)
+#139 := (implies #13 #138)
+#140 := (implies true #139)
+#9 := (< 0::Int f3)
+#141 := (implies #9 #140)
+#142 := (implies true #141)
+#143 := (not #142)
+#1024 := (iff #143 #1021)
+#307 := (not #89)
+#308 := (or #307 #90)
+#311 := (forall (vars (?v0 Int)) #308)
+#333 := (not #311)
+#334 := (or #333 #94)
+#339 := (and #311 #334)
+#352 := (not #87)
+#353 := (or #352 #339)
+#301 := (+ 1::Int f8)
+#304 := (= f16 #301)
+#361 := (not #304)
+#362 := (or #361 #353)
+#370 := (not #81)
+#371 := (or #370 #362)
+#454 := (or #453 #371)
+#463 := (or #462 #454)
+#269 := (not #33)
+#478 := (or #269 #463)
+#486 := (not #113)
+#487 := (or #486 #478)
+#495 := (or #269 #487)
+#387 := (or #386 #371)
+#396 := (or #395 #387)
+#411 := (not #32)
+#412 := (or #411 #396)
+#421 := (or #420 #412)
+#429 := (not #72)
+#430 := (or #429 #421)
+#438 := (or #269 #430)
+#507 := (and #438 #495)
+#513 := (or #269 #507)
+#521 := (not #70)
+#522 := (or #521 #513)
+#530 := (or #269 #522)
+#186 := (not #50)
+#193 := (or #186 #54)
+#196 := (forall (vars (?v0 Int)) #193)
+#187 := (or #186 #51)
+#190 := (exists (vars (?v0 Int)) #187)
+#216 := (not #190)
+#217 := (or #216 #196)
+#222 := (and #190 #217)
+#236 := (or #235 #222)
+#245 := (or #244 #236)
+#254 := (or #253 #245)
+#270 := (or #269 #254)
+#278 := (not #42)
+#279 := (or #278 #270)
+#287 := (or #269 #279)
+#542 := (and #287 #530)
+#548 := (or #269 #542)
+#557 := (or #556 #548)
+#179 := (not #35)
+#180 := (or #179 #37)
+#183 := (forall (vars (?v0 Int)) #180)
+#565 := (not #183)
+#566 := (or #565 #557)
+#574 := (or #269 #566)
+#590 := (or #589 #574)
+#595 := (and #28 #590)
+#172 := (not #23)
+#173 := (or #172 #25)
+#176 := (forall (vars (?v0 Int)) #173)
+#601 := (not #176)
+#602 := (or #601 #595)
+#607 := (and #176 #602)
+#166 := (and #14 #16)
+#169 := (and #14 #166)
+#613 := (not #169)
+#614 := (or #613 #607)
+#623 := (or #622 #614)
+#638 := (not #9)
+#639 := (or #638 #623)
+#651 := (not #639)
+#1022 := (iff #651 #1021)
+#1019 := (iff #639 #1016)
+#1007 := (or false #993)
+#1010 := (or #622 #1007)
+#1013 := (or #996 #1010)
+#1017 := (iff #1013 #1016)
+#1018 := [rewrite]: #1017
+#1014 := (iff #639 #1013)
+#1011 := (iff #623 #1010)
+#1008 := (iff #614 #1007)
+#994 := (iff #607 #993)
+#991 := (iff #602 #990)
+#988 := (iff #595 #987)
+#985 := (iff #590 #982)
+#967 := (or #712 #939)
+#970 := (or #556 #967)
+#973 := (or #964 #970)
+#976 := (or #712 #973)
+#979 := (or #589 #976)
+#983 := (iff #979 #982)
+#984 := [rewrite]: #983
+#980 := (iff #590 #979)
+#977 := (iff #574 #976)
+#974 := (iff #566 #973)
+#971 := (iff #557 #970)
+#968 := (iff #548 #967)
+#940 := (iff #542 #939)
+#937 := (iff #530 #934)
+#925 := (or #712 #913)
+#928 := (or #755 #925)
+#931 := (or #712 #928)
+#935 := (iff #931 #934)
+#936 := [rewrite]: #935
+#932 := (iff #530 #931)
+#929 := (iff #522 #928)
+#926 := (iff #513 #925)
+#914 := (iff #507 #913)
+#911 := (iff #495 #908)
+#857 := (or #793 #825)
+#860 := (or #832 #857)
+#863 := (or #838 #860)
+#893 := (or #453 #863)
+#896 := (or #462 #893)
+#899 := (or #712 #896)
+#902 := (or #844 #899)
+#905 := (or #712 #902)
+#909 := (iff #905 #908)
+#910 := [rewrite]: #909
+#906 := (iff #495 #905)
+#903 := (iff #487 #902)
+#900 := (iff #478 #899)
+#897 := (iff #463 #896)
+#894 := (iff #454 #893)
+#864 := (iff #371 #863)
+#861 := (iff #362 #860)
+#858 := (iff #353 #857)
+#826 := (iff #339 #825)
+#823 := (iff #334 #822)
+#820 := (iff #333 #819)
+#817 := (iff #311 #816)
+#814 := (iff #308 #813)
+#811 := (iff #90 #810)
+#812 := [rewrite]: #811
+#806 := (iff #307 #805)
+#803 := (iff #89 #802)
+#800 := (iff #88 #799)
+#801 := [rewrite]: #800
+#675 := (iff #21 #677)
+#676 := [rewrite]: #675
+#804 := [monotonicity #676 #801]: #803
+#807 := [monotonicity #804]: #806
+#815 := [monotonicity #807 #812]: #814
+#818 := [quant-intro #815]: #817
+#821 := [monotonicity #818]: #820
+#824 := [monotonicity #821]: #823
+#827 := [monotonicity #818 #824]: #826
+#794 := (iff #352 #793)
+#791 := (iff #87 #790)
+#787 := (iff #86 #788)
+#789 := [rewrite]: #787
+#784 := (iff #80 #785)
+#786 := [rewrite]: #784
+#792 := [monotonicity #786 #789]: #791
+#795 := [monotonicity #792]: #794
+#859 := [monotonicity #795 #827]: #858
+#833 := (iff #361 #832)
+#830 := (iff #304 #828)
+#831 := [rewrite]: #830
+#834 := [monotonicity #831]: #833
+#862 := [monotonicity #834 #859]: #861
+#839 := (iff #370 #838)
+#836 := (iff #81 #835)
+#706 := (iff #32 #707)
+#708 := [rewrite]: #706
+#837 := [monotonicity #786 #708]: #836
+#840 := [monotonicity #837]: #839
+#865 := [monotonicity #840 #862]: #864
+#895 := [monotonicity #865]: #894
+#898 := [monotonicity #895]: #897
+#713 := (iff #269 #712)
+#710 := (iff #33 #709)
+#703 := (iff #30 #704)
+#705 := [rewrite]: #703
+#711 := [monotonicity #705 #708]: #710
+#714 := [monotonicity #711]: #713
+#901 := [monotonicity #714 #898]: #900
+#891 := (iff #486 #844)
+#889 := (iff #113 #845)
+#890 := [rewrite]: #889
+#892 := [monotonicity #890]: #891
+#904 := [monotonicity #892 #901]: #903
+#907 := [monotonicity #714 #904]: #906
+#912 := [trans #907 #910]: #911
+#887 := (iff #438 #884)
+#866 := (or #386 #863)
+#869 := (or #395 #866)
+#872 := (or #841 #869)
+#875 := (or #420 #872)
+#878 := (or #845 #875)
+#881 := (or #712 #878)
+#885 := (iff #881 #884)
+#886 := [rewrite]: #885
+#882 := (iff #438 #881)
+#879 := (iff #430 #878)
+#876 := (iff #421 #875)
+#873 := (iff #412 #872)
+#870 := (iff #396 #869)
+#867 := (iff #387 #866)
+#868 := [monotonicity #865]: #867
+#871 := [monotonicity #868]: #870
+#842 := (iff #411 #841)
+#843 := [monotonicity #708]: #842
+#874 := [monotonicity #843 #871]: #873
+#877 := [monotonicity #874]: #876
+#855 := (iff #429 #845)
+#850 := (not #844)
+#853 := (iff #850 #845)
+#854 := [rewrite]: #853
+#851 := (iff #429 #850)
+#848 := (iff #72 #844)
+#849 := [rewrite]: #848
+#852 := [monotonicity #849]: #851
+#856 := [trans #852 #854]: #855
+#880 := [monotonicity #856 #877]: #879
+#883 := [monotonicity #714 #880]: #882
+#888 := [trans #883 #886]: #887
+#915 := [monotonicity #888 #912]: #914
+#927 := [monotonicity #714 #915]: #926
+#923 := (iff #521 #755)
+#918 := (not #758)
+#921 := (iff #918 #755)
+#922 := [rewrite]: #921
+#919 := (iff #521 #918)
+#916 := (iff #70 #758)
+#917 := [rewrite]: #916
+#920 := [monotonicity #917]: #919
+#924 := [trans #920 #922]: #923
+#930 := [monotonicity #924 #927]: #929
+#933 := [monotonicity #714 #930]: #932
+#938 := [trans #933 #936]: #937
+#782 := (iff #287 #779)
+#761 := (or #235 #750)
+#764 := (or #244 #761)
+#767 := (or #253 #764)
+#770 := (or #712 #767)
+#773 := (or #758 #770)
+#776 := (or #712 #773)
+#780 := (iff #776 #779)
+#781 := [rewrite]: #780
+#777 := (iff #287 #776)
+#774 := (iff #279 #773)
+#771 := (iff #270 #770)
+#768 := (iff #254 #767)
+#765 := (iff #245 #764)
+#762 := (iff #236 #761)
+#751 := (iff #222 #750)
+#748 := (iff #217 #747)
+#745 := (iff #196 #744)
+#742 := (iff #193 #741)
+#739 := (iff #54 #738)
+#740 := [rewrite]: #739
+#725 := (iff #186 #724)
+#722 := (iff #50 #721)
+#719 := (iff #49 #718)
+#720 := [rewrite]: #719
+#723 := [monotonicity #676 #720]: #722
+#726 := [monotonicity #723]: #725
+#743 := [monotonicity #726 #740]: #742
+#746 := [quant-intro #743]: #745
+#734 := (iff #216 #733)
+#731 := (iff #190 #730)
+#728 := (iff #187 #727)
+#729 := [monotonicity #726]: #728
+#732 := [quant-intro #729]: #731
+#735 := [monotonicity #732]: #734
+#749 := [monotonicity #735 #746]: #748
+#752 := [monotonicity #732 #749]: #751
+#763 := [monotonicity #752]: #762
+#766 := [monotonicity #763]: #765
+#769 := [monotonicity #766]: #768
+#772 := [monotonicity #714 #769]: #771
+#759 := (iff #278 #758)
+#756 := (iff #42 #755)
+#757 := [rewrite]: #756
+#760 := [monotonicity #757]: #759
+#775 := [monotonicity #760 #772]: #774
+#778 := [monotonicity #714 #775]: #777
+#783 := [trans #778 #781]: #782
+#941 := [monotonicity #783 #938]: #940
+#969 := [monotonicity #714 #941]: #968
+#972 := [monotonicity #969]: #971
+#965 := (iff #565 #964)
+#962 := (iff #183 #961)
+#959 := (iff #180 #958)
+#956 := (iff #37 #955)
+#957 := [rewrite]: #956
+#951 := (iff #179 #950)
+#948 := (iff #35 #947)
+#945 := (iff #34 #944)
+#946 := [rewrite]: #945
+#949 := [monotonicity #676 #946]: #948
+#952 := [monotonicity #949]: #951
+#960 := [monotonicity #952 #957]: #959
+#963 := [quant-intro #960]: #962
+#966 := [monotonicity #963]: #965
+#975 := [monotonicity #966 #972]: #974
+#978 := [monotonicity #714 #975]: #977
+#981 := [monotonicity #978]: #980
+#986 := [trans #981 #984]: #985
+#989 := [monotonicity #986]: #988
+#701 := (iff #601 #700)
+#698 := (iff #176 #697)
+#695 := (iff #173 #694)
+#689 := (iff #25 #690)
+#693 := [rewrite]: #689
+#686 := (iff #172 #685)
+#683 := (iff #23 #682)
+#679 := (iff #22 #678)
+#681 := [rewrite]: #679
+#684 := [monotonicity #676 #681]: #683
+#687 := [monotonicity #684]: #686
+#696 := [monotonicity #687 #693]: #695
+#699 := [quant-intro #696]: #698
+#702 := [monotonicity #699]: #701
+#992 := [monotonicity #702 #989]: #991
+#995 := [monotonicity #699 #992]: #994
+#673 := (iff #613 false)
+#668 := (not true)
+#671 := (iff #668 false)
+#672 := [rewrite]: #671
+#669 := (iff #613 #668)
+#666 := (iff #169 true)
+#658 := (and true true)
+#661 := (and true #658)
+#664 := (iff #661 true)
+#665 := [rewrite]: #664
+#662 := (iff #169 #661)
+#659 := (iff #166 #658)
+#656 := (iff #16 true)
+#657 := [rewrite]: #656
+#654 := (iff #14 true)
+#655 := [rewrite]: #654
+#660 := [monotonicity #655 #657]: #659
+#663 := [monotonicity #655 #660]: #662
+#667 := [trans #663 #665]: #666
+#670 := [monotonicity #667]: #669
+#674 := [trans #670 #672]: #673
+#1009 := [monotonicity #674 #995]: #1008
+#1012 := [monotonicity #1009]: #1011
+#1005 := (iff #638 #996)
+#997 := (not #996)
+#1000 := (not #997)
+#1003 := (iff #1000 #996)
+#1004 := [rewrite]: #1003
+#1001 := (iff #638 #1000)
+#998 := (iff #9 #997)
+#999 := [rewrite]: #998
+#1002 := [monotonicity #999]: #1001
+#1006 := [trans #1002 #1004]: #1005
+#1015 := [monotonicity #1006 #1012]: #1014
+#1020 := [trans #1015 #1018]: #1019
+#1023 := [monotonicity #1020]: #1022
+#652 := (iff #143 #651)
+#649 := (iff #142 #639)
+#644 := (implies true #639)
+#647 := (iff #644 #639)
+#648 := [rewrite]: #647
+#645 := (iff #142 #644)
+#642 := (iff #141 #639)
+#635 := (implies #9 #623)
+#640 := (iff #635 #639)
+#641 := [rewrite]: #640
+#636 := (iff #141 #635)
+#633 := (iff #140 #623)
+#628 := (implies true #623)
+#631 := (iff #628 #623)
+#632 := [rewrite]: #631
+#629 := (iff #140 #628)
+#626 := (iff #139 #623)
+#619 := (implies #13 #614)
+#624 := (iff #619 #623)
+#625 := [rewrite]: #624
+#620 := (iff #139 #619)
+#617 := (iff #138 #614)
+#610 := (implies #169 #607)
+#615 := (iff #610 #614)
+#616 := [rewrite]: #615
+#611 := (iff #138 #610)
+#608 := (iff #137 #607)
+#605 := (iff #136 #602)
+#598 := (implies #176 #595)
+#603 := (iff #598 #602)
+#604 := [rewrite]: #603
+#599 := (iff #136 #598)
+#596 := (iff #135 #595)
+#593 := (iff #134 #590)
+#586 := (implies #28 #574)
+#591 := (iff #586 #590)
+#592 := [rewrite]: #591
+#587 := (iff #134 #586)
+#584 := (iff #133 #574)
+#579 := (implies true #574)
+#582 := (iff #579 #574)
+#583 := [rewrite]: #582
+#580 := (iff #133 #579)
+#577 := (iff #132 #574)
+#571 := (implies #33 #566)
+#575 := (iff #571 #574)
+#576 := [rewrite]: #575
+#572 := (iff #132 #571)
+#569 := (iff #131 #566)
+#562 := (implies #183 #557)
+#567 := (iff #562 #566)
+#568 := [rewrite]: #567
+#563 := (iff #131 #562)
+#560 := (iff #130 #557)
+#553 := (implies #41 #548)
+#558 := (iff #553 #557)
+#559 := [rewrite]: #558
+#554 := (iff #130 #553)
+#551 := (iff #129 #548)
+#545 := (implies #33 #542)
+#549 := (iff #545 #548)
+#550 := [rewrite]: #549
+#546 := (iff #129 #545)
+#543 := (iff #128 #542)
+#540 := (iff #127 #530)
+#535 := (implies true #530)
+#538 := (iff #535 #530)
+#539 := [rewrite]: #538
+#536 := (iff #127 #535)
+#533 := (iff #126 #530)
+#527 := (implies #33 #522)
+#531 := (iff #527 #530)
+#532 := [rewrite]: #531
+#528 := (iff #126 #527)
+#525 := (iff #125 #522)
+#518 := (implies #70 #513)
+#523 := (iff #518 #522)
+#524 := [rewrite]: #523
+#519 := (iff #125 #518)
+#516 := (iff #124 #513)
+#510 := (implies #33 #507)
+#514 := (iff #510 #513)
+#515 := [rewrite]: #514
+#511 := (iff #124 #510)
+#508 := (iff #123 #507)
+#505 := (iff #122 #495)
+#500 := (implies true #495)
+#503 := (iff #500 #495)
+#504 := [rewrite]: #503
+#501 := (iff #122 #500)
+#498 := (iff #121 #495)
+#492 := (implies #33 #487)
+#496 := (iff #492 #495)
+#497 := [rewrite]: #496
+#493 := (iff #121 #492)
+#490 := (iff #120 #487)
+#483 := (implies #113 #478)
+#488 := (iff #483 #487)
+#489 := [rewrite]: #488
+#484 := (iff #120 #483)
+#481 := (iff #119 #478)
+#475 := (implies #33 #463)
+#479 := (iff #475 #478)
+#480 := [rewrite]: #479
+#476 := (iff #119 #475)
+#473 := (iff #118 #463)
+#468 := (implies true #463)
+#471 := (iff #468 #463)
+#472 := [rewrite]: #471
+#469 := (iff #118 #468)
+#466 := (iff #117 #463)
+#459 := (implies #114 #454)
+#464 := (iff #459 #463)
+#465 := [rewrite]: #464
+#460 := (iff #117 #459)
+#457 := (iff #116 #454)
+#450 := (implies #115 #371)
+#455 := (iff #450 #454)
+#456 := [rewrite]: #455
+#451 := (iff #116 #450)
+#381 := (iff #104 #371)
+#376 := (implies true #371)
+#379 := (iff #376 #371)
+#380 := [rewrite]: #379
+#377 := (iff #104 #376)
+#374 := (iff #103 #371)
+#367 := (implies #81 #362)
+#372 := (iff #367 #371)
+#373 := [rewrite]: #372
+#368 := (iff #103 #367)
+#365 := (iff #102 #362)
+#358 := (implies #304 #353)
+#363 := (iff #358 #362)
+#364 := [rewrite]: #363
+#359 := (iff #102 #358)
+#356 := (iff #101 #353)
+#349 := (implies #87 #339)
+#354 := (iff #349 #353)
+#355 := [rewrite]: #354
+#350 := (iff #101 #349)
+#347 := (iff #100 #339)
+#342 := (implies true #339)
+#345 := (iff #342 #339)
+#346 := [rewrite]: #345
+#343 := (iff #100 #342)
+#340 := (iff #99 #339)
+#337 := (iff #98 #334)
+#330 := (implies #311 #94)
+#335 := (iff #330 #334)
+#336 := [rewrite]: #335
+#331 := (iff #98 #330)
+#328 := (iff #97 #94)
+#323 := (and #94 true)
+#326 := (iff #323 #94)
+#327 := [rewrite]: #326
+#324 := (iff #97 #323)
+#321 := (iff #96 true)
+#316 := (implies #94 true)
+#319 := (iff #316 true)
+#320 := [rewrite]: #319
+#317 := (iff #96 #316)
+#314 := (iff #95 true)
+#315 := [rewrite]: #314
+#318 := [monotonicity #315]: #317
+#322 := [trans #318 #320]: #321
+#325 := [monotonicity #322]: #324
+#329 := [trans #325 #327]: #328
+#312 := (iff #92 #311)
+#309 := (iff #91 #308)
+#310 := [rewrite]: #309
+#313 := [quant-intro #310]: #312
+#332 := [monotonicity #313 #329]: #331
+#338 := [trans #332 #336]: #337
+#341 := [monotonicity #313 #338]: #340
+#344 := [monotonicity #341]: #343
+#348 := [trans #344 #346]: #347
+#351 := [monotonicity #348]: #350
+#357 := [trans #351 #355]: #356
+#305 := (iff #84 #304)
+#302 := (= #83 #301)
+#303 := [rewrite]: #302
+#306 := [monotonicity #303]: #305
+#360 := [monotonicity #306 #357]: #359
+#366 := [trans #360 #364]: #365
+#369 := [monotonicity #366]: #368
+#375 := [trans #369 #373]: #374
+#378 := [monotonicity #375]: #377
+#382 := [trans #378 #380]: #381
+#452 := [monotonicity #382]: #451
+#458 := [trans #452 #456]: #457
+#461 := [monotonicity #458]: #460
+#467 := [trans #461 #465]: #466
+#470 := [monotonicity #467]: #469
+#474 := [trans #470 #472]: #473
+#477 := [monotonicity #474]: #476
+#482 := [trans #477 #480]: #481
+#485 := [monotonicity #482]: #484
+#491 := [trans #485 #489]: #490
+#494 := [monotonicity #491]: #493
+#499 := [trans #494 #497]: #498
+#502 := [monotonicity #499]: #501
+#506 := [trans #502 #504]: #505
+#448 := (iff #112 #438)
+#443 := (implies true #438)
+#446 := (iff #443 #438)
+#447 := [rewrite]: #446
+#444 := (iff #112 #443)
+#441 := (iff #111 #438)
+#435 := (implies #33 #430)
+#439 := (iff #435 #438)
+#440 := [rewrite]: #439
+#436 := (iff #111 #435)
+#433 := (iff #110 #430)
+#426 := (implies #72 #421)
+#431 := (iff #426 #430)
+#432 := [rewrite]: #431
+#427 := (iff #110 #426)
+#424 := (iff #109 #421)
+#417 := (implies #74 #412)
+#422 := (iff #417 #421)
+#423 := [rewrite]: #422
+#418 := (iff #109 #417)
+#415 := (iff #108 #412)
+#408 := (implies #32 #396)
+#413 := (iff #408 #412)
+#414 := [rewrite]: #413
+#409 := (iff #108 #408)
+#406 := (iff #107 #396)
+#401 := (implies true #396)
+#404 := (iff #401 #396)
+#405 := [rewrite]: #404
+#402 := (iff #107 #401)
+#399 := (iff #106 #396)
+#392 := (implies #77 #387)
+#397 := (iff #392 #396)
+#398 := [rewrite]: #397
+#393 := (iff #106 #392)
+#390 := (iff #105 #387)
+#383 := (implies #79 #371)
+#388 := (iff #383 #387)
+#389 := [rewrite]: #388
+#384 := (iff #105 #383)
+#385 := [monotonicity #382]: #384
+#391 := [trans #385 #389]: #390
+#394 := [monotonicity #391]: #393
+#400 := [trans #394 #398]: #399
+#403 := [monotonicity #400]: #402
+#407 := [trans #403 #405]: #406
+#299 := (iff #75 #32)
+#300 := [rewrite]: #299
+#410 := [monotonicity #300 #407]: #409
+#416 := [trans #410 #414]: #415
+#419 := [monotonicity #416]: #418
+#425 := [trans #419 #423]: #424
+#428 := [monotonicity #425]: #427
+#434 := [trans #428 #432]: #433
+#437 := [monotonicity #434]: #436
+#442 := [trans #437 #440]: #441
+#445 := [monotonicity #442]: #444
+#449 := [trans #445 #447]: #448
+#509 := [monotonicity #449 #506]: #508
+#512 := [monotonicity #509]: #511
+#517 := [trans #512 #515]: #516
+#520 := [monotonicity #517]: #519
+#526 := [trans #520 #524]: #525
+#529 := [monotonicity #526]: #528
+#534 := [trans #529 #532]: #533
+#537 := [monotonicity #534]: #536
+#541 := [trans #537 #539]: #540
+#297 := (iff #69 #287)
+#292 := (implies true #287)
+#295 := (iff #292 #287)
+#296 := [rewrite]: #295
+#293 := (iff #69 #292)
+#290 := (iff #68 #287)
+#284 := (implies #33 #279)
+#288 := (iff #284 #287)
+#289 := [rewrite]: #288
+#285 := (iff #68 #284)
+#282 := (iff #67 #279)
+#275 := (implies #42 #270)
+#280 := (iff #275 #279)
+#281 := [rewrite]: #280
+#276 := (iff #67 #275)
+#273 := (iff #66 #270)
+#266 := (implies #33 #254)
+#271 := (iff #266 #270)
+#272 := [rewrite]: #271
+#267 := (iff #66 #266)
+#264 := (iff #65 #254)
+#259 := (implies true #254)
+#262 := (iff #259 #254)
+#263 := [rewrite]: #262
+#260 := (iff #65 #259)
+#257 := (iff #64 #254)
+#250 := (implies #44 #245)
+#255 := (iff #250 #254)
+#256 := [rewrite]: #255
+#251 := (iff #64 #250)
+#248 := (iff #63 #245)
+#241 := (implies #46 #236)
+#246 := (iff #241 #245)
+#247 := [rewrite]: #246
+#242 := (iff #63 #241)
+#239 := (iff #62 #236)
+#232 := (implies #48 #222)
+#237 := (iff #232 #236)
+#238 := [rewrite]: #237
+#233 := (iff #62 #232)
+#230 := (iff #61 #222)
+#225 := (implies true #222)
+#228 := (iff #225 #222)
+#229 := [rewrite]: #228
+#226 := (iff #61 #225)
+#223 := (iff #60 #222)
+#220 := (iff #59 #217)
+#213 := (implies #190 #196)
+#218 := (iff #213 #217)
+#219 := [rewrite]: #218
+#214 := (iff #59 #213)
+#211 := (iff #58 #196)
+#206 := (and #196 true)
+#209 := (iff #206 #196)
+#210 := [rewrite]: #209
+#207 := (iff #58 #206)
+#204 := (iff #57 true)
+#199 := (implies #196 true)
+#202 := (iff #199 true)
+#203 := [rewrite]: #202
+#200 := (iff #57 #199)
+#197 := (iff #56 #196)
+#194 := (iff #55 #193)
+#195 := [rewrite]: #194
+#198 := [quant-intro #195]: #197
+#201 := [monotonicity #198]: #200
+#205 := [trans #201 #203]: #204
+#208 := [monotonicity #198 #205]: #207
+#212 := [trans #208 #210]: #211
+#191 := (iff #53 #190)
+#188 := (iff #52 #187)
+#189 := [rewrite]: #188
+#192 := [quant-intro #189]: #191
+#215 := [monotonicity #192 #212]: #214
+#221 := [trans #215 #219]: #220
+#224 := [monotonicity #192 #221]: #223
+#227 := [monotonicity #224]: #226
+#231 := [trans #227 #229]: #230
+#234 := [monotonicity #231]: #233
+#240 := [trans #234 #238]: #239
+#243 := [monotonicity #240]: #242
+#249 := [trans #243 #247]: #248
+#252 := [monotonicity #249]: #251
+#258 := [trans #252 #256]: #257
+#261 := [monotonicity #258]: #260
+#265 := [trans #261 #263]: #264
+#268 := [monotonicity #265]: #267
+#274 := [trans #268 #272]: #273
+#277 := [monotonicity #274]: #276
+#283 := [trans #277 #281]: #282
+#286 := [monotonicity #283]: #285
+#291 := [trans #286 #289]: #290
+#294 := [monotonicity #291]: #293
+#298 := [trans #294 #296]: #297
+#544 := [monotonicity #298 #541]: #543
+#547 := [monotonicity #544]: #546
+#552 := [trans #547 #550]: #551
+#555 := [monotonicity #552]: #554
+#561 := [trans #555 #559]: #560
+#184 := (iff #39 #183)
+#181 := (iff #38 #180)
+#182 := [rewrite]: #181
+#185 := [quant-intro #182]: #184
+#564 := [monotonicity #185 #561]: #563
+#570 := [trans #564 #568]: #569
+#573 := [monotonicity #570]: #572
+#578 := [trans #573 #576]: #577
+#581 := [monotonicity #578]: #580
+#585 := [trans #581 #583]: #584
+#588 := [monotonicity #585]: #587
+#594 := [trans #588 #592]: #593
+#597 := [monotonicity #594]: #596
+#177 := (iff #27 #176)
+#174 := (iff #26 #173)
+#175 := [rewrite]: #174
+#178 := [quant-intro #175]: #177
+#600 := [monotonicity #178 #597]: #599
+#606 := [trans #600 #604]: #605
+#609 := [monotonicity #178 #606]: #608
+#170 := (iff #19 #169)
+#167 := (iff #18 #166)
+#164 := (iff #17 #16)
+#165 := [rewrite]: #164
+#168 := [monotonicity #165]: #167
+#171 := [monotonicity #168]: #170
+#612 := [monotonicity #171 #609]: #611
+#618 := [trans #612 #616]: #617
+#621 := [monotonicity #618]: #620
+#627 := [trans #621 #625]: #626
+#630 := [monotonicity #627]: #629
+#634 := [trans #630 #632]: #633
+#637 := [monotonicity #634]: #636
+#643 := [trans #637 #641]: #642
+#646 := [monotonicity #643]: #645
+#650 := [trans #646 #648]: #649
+#653 := [monotonicity #650]: #652
+#1025 := [trans #653 #1023]: #1024
+#163 := [asserted]: #143
+#1026 := [mp #163 #1025]: #1021
+#1028 := [not-or-elim #1026]: #13
+#1933 := [trans #1028 #1926]: #1999
+#1821 := (not #1033)
+#1812 := (or #1405 #1821)
+#1823 := [def-axiom]: #1812
+#1935 := [unit-resolution #1823 #1998]: #1821
+#1936 := (not #1999)
+#1964 := (or #1936 #1033)
+#1937 := [th-lemma arith triangle-eq]: #1964
+#1939 := [unit-resolution #1937 #1935 #1933]: false
+#1940 := [lemma #1939]: #1405
+#2250 := (or #1410 #2247)
+#1592 := (forall (vars (?v0 Int)) #1587)
+#1598 := (not #1592)
+#1599 := (or #1598 #94)
+#1600 := (not #1599)
+#1605 := (or #1570 #1600)
+#1613 := (not #1605)
+#1623 := (or #1542 #844 #462 #453 #841 #832 #1611 #1612 #1613)
+#1624 := (not #1623)
+#1614 := (or #1542 #845 #420 #395 #386 #841 #832 #1611 #1612 #1613)
+#1615 := (not #1614)
+#1629 := (or #1615 #1624)
+#1635 := (not #1629)
+#1636 := (or #1542 #841 #755 #1635)
+#1637 := (not #1636)
+#1480 := (forall (vars (?v0 Int)) #1477)
+#1536 := (or #1480 #1531)
+#1543 := (not #1536)
+#1544 := (or #1542 #841 #758 #253 #244 #235 #1543)
+#1545 := (not #1544)
+#1642 := (or #1545 #1637)
+#1649 := (not #1642)
+#1455 := (forall (vars (?v0 Int)) #1450)
+#1648 := (not #1455)
+#1650 := (or #589 #1542 #841 #1648 #556 #1649)
+#1651 := (not #1650)
+#1656 := (or #589 #1651)
+#1663 := (not #1656)
+#1433 := (forall (vars (?v0 Int)) #1428)
+#1662 := (not #1433)
+#1664 := (or #1662 #1663)
+#1665 := (not #1664)
+#1670 := (or #1410 #1665)
+#2251 := (iff #1670 #2250)
+#2248 := (iff #1665 #2247)
+#2245 := (iff #1664 #2244)
+#2242 := (iff #1663 #2241)
+#2239 := (iff #1656 #2238)
+#2236 := (iff #1651 #2235)
+#2233 := (iff #1650 #2232)
+#2230 := (iff #1649 #2229)
+#2227 := (iff #1642 #2226)
+#2224 := (iff #1637 #2223)
+#2221 := (iff #1636 #2220)
+#2218 := (iff #1635 #2217)
+#2215 := (iff #1629 #2214)
+#2212 := (iff #1624 #2211)
+#2209 := (iff #1623 #2208)
+#2200 := (iff #1613 #2199)
+#2197 := (iff #1605 #2196)
+#2194 := (iff #1600 #2193)
+#2191 := (iff #1599 #2190)
+#2188 := (iff #1598 #2187)
+#2185 := (iff #1592 #2182)
+#2183 := (iff #1587 #1587)
+#2184 := [refl]: #2183
+#2186 := [quant-intro #2184]: #2185
+#2189 := [monotonicity #2186]: #2188
+#2192 := [monotonicity #2189]: #2191
+#2195 := [monotonicity #2192]: #2194
+#2198 := [monotonicity #2195]: #2197
+#2201 := [monotonicity #2198]: #2200
+#2210 := [monotonicity #2201]: #2209
+#2213 := [monotonicity #2210]: #2212
+#2206 := (iff #1615 #2205)
+#2203 := (iff #1614 #2202)
+#2204 := [monotonicity #2201]: #2203
+#2207 := [monotonicity #2204]: #2206
+#2216 := [monotonicity #2207 #2213]: #2215
+#2219 := [monotonicity #2216]: #2218
+#2222 := [monotonicity #2219]: #2221
+#2225 := [monotonicity #2222]: #2224
+#2180 := (iff #1545 #2179)
+#2177 := (iff #1544 #2176)
+#2174 := (iff #1543 #2173)
+#2171 := (iff #1536 #2170)
+#2168 := (iff #1480 #2165)
+#2166 := (iff #1477 #1477)
+#2167 := [refl]: #2166
+#2169 := [quant-intro #2167]: #2168
+#2172 := [monotonicity #2169]: #2171
+#2175 := [monotonicity #2172]: #2174
+#2178 := [monotonicity #2175]: #2177
+#2181 := [monotonicity #2178]: #2180
+#2228 := [monotonicity #2181 #2225]: #2227
+#2231 := [monotonicity #2228]: #2230
+#2163 := (iff #1648 #2162)
+#2160 := (iff #1455 #2157)
+#2158 := (iff #1450 #1450)
+#2159 := [refl]: #2158
+#2161 := [quant-intro #2159]: #2160
+#2164 := [monotonicity #2161]: #2163
+#2234 := [monotonicity #2164 #2231]: #2233
+#2237 := [monotonicity #2234]: #2236
+#2240 := [monotonicity #2237]: #2239
+#2243 := [monotonicity #2240]: #2242
+#2155 := (iff #1662 #2154)
+#2152 := (iff #1433 #2149)
+#2150 := (iff #1428 #1428)
+#2151 := [refl]: #2150
+#2153 := [quant-intro #2151]: #2152
+#2156 := [monotonicity #2153]: #2155
+#2246 := [monotonicity #2156 #2243]: #2245
+#2249 := [monotonicity #2246]: #2248
+#2252 := [monotonicity #2249]: #2251
+#1188 := (not #94)
+#1191 := (and #816 #1188)
+#1317 := (not #1312)
+#1320 := (and #1166 #1317)
+#1323 := (not #1320)
+#1339 := (or #1323 #1334)
+#1342 := (not #1339)
+#1345 := (or #1342 #1191)
+#1363 := (and #704 #845 #114 #115 #707 #828 #785 #788 #1345)
+#1351 := (and #704 #844 #74 #77 #79 #707 #828 #785 #788 #1345)
+#1368 := (or #1351 #1363)
+#1374 := (and #704 #707 #758 #1368)
+#1115 := (not #1114)
+#1116 := (and #1111 #1115)
+#1117 := (not #1116)
+#1270 := (or #1117 #1265)
+#1273 := (not #1270)
+#1097 := (not #1096)
+#1098 := (and #1093 #1097)
+#1099 := (not #1098)
+#1102 := (or #1099 #1101)
+#1276 := (and #1102 #1273)
+#1086 := (not #727)
+#1089 := (forall (vars (?v0 Int)) #1086)
+#1279 := (or #1089 #1276)
+#1285 := (and #704 #707 #755 #44 #46 #48 #1279)
+#1379 := (or #1285 #1374)
+#1385 := (and #28 #704 #707 #961 #41 #1379)
+#1390 := (or #589 #1385)
+#1393 := (and #697 #1390)
+#1036 := (and #1041 #1043)
+#1037 := (not #1036)
+#1044 := (or #1037 #1033)
+#1045 := (not #1044)
+#1396 := (or #1045 #1393)
+#1671 := (iff #1396 #1670)
+#1668 := (iff #1393 #1665)
+#1659 := (and #1433 #1656)
+#1666 := (iff #1659 #1665)
+#1667 := [rewrite]: #1666
+#1660 := (iff #1393 #1659)
+#1657 := (iff #1390 #1656)
+#1654 := (iff #1385 #1651)
+#1645 := (and #28 #704 #707 #1455 #41 #1642)
+#1652 := (iff #1645 #1651)
+#1653 := [rewrite]: #1652
+#1646 := (iff #1385 #1645)
+#1643 := (iff #1379 #1642)
+#1640 := (iff #1374 #1637)
+#1632 := (and #704 #707 #758 #1629)
+#1638 := (iff #1632 #1637)
+#1639 := [rewrite]: #1638
+#1633 := (iff #1374 #1632)
+#1630 := (iff #1368 #1629)
+#1627 := (iff #1363 #1624)
+#1620 := (and #704 #845 #114 #115 #707 #828 #785 #788 #1605)
+#1625 := (iff #1620 #1624)
+#1626 := [rewrite]: #1625
+#1621 := (iff #1363 #1620)
+#1606 := (iff #1345 #1605)
+#1603 := (iff #1191 #1600)
+#1595 := (and #1592 #1188)
+#1601 := (iff #1595 #1600)
+#1602 := [rewrite]: #1601
+#1596 := (iff #1191 #1595)
+#1593 := (iff #816 #1592)
+#1590 := (iff #813 #1587)
+#1573 := (or #1413 #796)
+#1584 := (or #1573 #810)
+#1588 := (iff #1584 #1587)
+#1589 := [rewrite]: #1588
+#1585 := (iff #813 #1584)
+#1582 := (iff #805 #1573)
+#1574 := (not #1573)
+#1577 := (not #1574)
+#1580 := (iff #1577 #1573)
+#1581 := [rewrite]: #1580
+#1578 := (iff #805 #1577)
+#1575 := (iff #802 #1574)
+#1576 := [rewrite]: #1575
+#1579 := [monotonicity #1576]: #1578
+#1583 := [trans #1579 #1581]: #1582
+#1586 := [monotonicity #1583]: #1585
+#1591 := [trans #1586 #1589]: #1590
+#1594 := [quant-intro #1591]: #1593
+#1597 := [monotonicity #1594]: #1596
+#1604 := [trans #1597 #1602]: #1603
+#1571 := (iff #1342 #1570)
+#1568 := (iff #1339 #1565)
+#1551 := (or #1550 #1312)
+#1562 := (or #1551 #1334)
+#1566 := (iff #1562 #1565)
+#1567 := [rewrite]: #1566
+#1563 := (iff #1339 #1562)
+#1560 := (iff #1323 #1551)
+#1552 := (not #1551)
+#1555 := (not #1552)
+#1558 := (iff #1555 #1551)
+#1559 := [rewrite]: #1558
+#1556 := (iff #1323 #1555)
+#1553 := (iff #1320 #1552)
+#1554 := [rewrite]: #1553
+#1557 := [monotonicity #1554]: #1556
+#1561 := [trans #1557 #1559]: #1560
+#1564 := [monotonicity #1561]: #1563
+#1569 := [trans #1564 #1567]: #1568
+#1572 := [monotonicity #1569]: #1571
+#1607 := [monotonicity #1572 #1604]: #1606
+#1622 := [monotonicity #1607]: #1621
+#1628 := [trans #1622 #1626]: #1627
+#1618 := (iff #1351 #1615)
+#1608 := (and #704 #844 #74 #77 #79 #707 #828 #785 #788 #1605)
+#1616 := (iff #1608 #1615)
+#1617 := [rewrite]: #1616
+#1609 := (iff #1351 #1608)
+#1610 := [monotonicity #1607]: #1609
+#1619 := [trans #1610 #1617]: #1618
+#1631 := [monotonicity #1619 #1628]: #1630
+#1634 := [monotonicity #1631]: #1633
+#1641 := [trans #1634 #1639]: #1640
+#1548 := (iff #1285 #1545)
+#1539 := (and #704 #707 #755 #44 #46 #48 #1536)
+#1546 := (iff #1539 #1545)
+#1547 := [rewrite]: #1546
+#1540 := (iff #1285 #1539)
+#1537 := (iff #1279 #1536)
+#1534 := (iff #1276 #1531)
+#1518 := (or #1503 #1114 #1265)
+#1523 := (not #1518)
+#1526 := (and #1498 #1523)
+#1532 := (iff #1526 #1531)
+#1533 := [rewrite]: #1532
+#1527 := (iff #1276 #1526)
+#1524 := (iff #1273 #1523)
+#1521 := (iff #1270 #1518)
+#1504 := (or #1503 #1114)
+#1515 := (or #1504 #1265)
+#1519 := (iff #1515 #1518)
+#1520 := [rewrite]: #1519
+#1516 := (iff #1270 #1515)
+#1513 := (iff #1117 #1504)
+#1505 := (not #1504)
+#1508 := (not #1505)
+#1511 := (iff #1508 #1504)
+#1512 := [rewrite]: #1511
+#1509 := (iff #1117 #1508)
+#1506 := (iff #1116 #1505)
+#1507 := [rewrite]: #1506
+#1510 := [monotonicity #1507]: #1509
+#1514 := [trans #1510 #1512]: #1513
+#1517 := [monotonicity #1514]: #1516
+#1522 := [trans #1517 #1520]: #1521
+#1525 := [monotonicity #1522]: #1524
+#1501 := (iff #1102 #1498)
+#1484 := (or #1483 #1096)
+#1495 := (or #1484 #1101)
+#1499 := (iff #1495 #1498)
+#1500 := [rewrite]: #1499
+#1496 := (iff #1102 #1495)
+#1493 := (iff #1099 #1484)
+#1485 := (not #1484)
+#1488 := (not #1485)
+#1491 := (iff #1488 #1484)
+#1492 := [rewrite]: #1491
+#1489 := (iff #1099 #1488)
+#1486 := (iff #1098 #1485)
+#1487 := [rewrite]: #1486
+#1490 := [monotonicity #1487]: #1489
+#1494 := [trans #1490 #1492]: #1493
+#1497 := [monotonicity #1494]: #1496
+#1502 := [trans #1497 #1500]: #1501
+#1528 := [monotonicity #1502 #1525]: #1527
+#1535 := [trans #1528 #1533]: #1534
+#1481 := (iff #1089 #1480)
+#1478 := (iff #1086 #1477)
+#1475 := (iff #727 #1472)
+#1458 := (or #1413 #717)
+#1469 := (or #1458 #51)
+#1473 := (iff #1469 #1472)
+#1474 := [rewrite]: #1473
+#1470 := (iff #727 #1469)
+#1467 := (iff #724 #1458)
+#1459 := (not #1458)
+#1462 := (not #1459)
+#1465 := (iff #1462 #1458)
+#1466 := [rewrite]: #1465
+#1463 := (iff #724 #1462)
+#1460 := (iff #721 #1459)
+#1461 := [rewrite]: #1460
+#1464 := [monotonicity #1461]: #1463
+#1468 := [trans #1464 #1466]: #1467
+#1471 := [monotonicity #1468]: #1470
+#1476 := [trans #1471 #1474]: #1475
+#1479 := [monotonicity #1476]: #1478
+#1482 := [quant-intro #1479]: #1481
+#1538 := [monotonicity #1482 #1535]: #1537
+#1541 := [monotonicity #1538]: #1540
+#1549 := [trans #1541 #1547]: #1548
+#1644 := [monotonicity #1549 #1641]: #1643
+#1456 := (iff #961 #1455)
+#1453 := (iff #958 #1450)
+#1436 := (or #1413 #942)
+#1447 := (or #1436 #955)
+#1451 := (iff #1447 #1450)
+#1452 := [rewrite]: #1451
+#1448 := (iff #958 #1447)
+#1445 := (iff #950 #1436)
+#1437 := (not #1436)
+#1440 := (not #1437)
+#1443 := (iff #1440 #1436)
+#1444 := [rewrite]: #1443
+#1441 := (iff #950 #1440)
+#1438 := (iff #947 #1437)
+#1439 := [rewrite]: #1438
+#1442 := [monotonicity #1439]: #1441
+#1446 := [trans #1442 #1444]: #1445
+#1449 := [monotonicity #1446]: #1448
+#1454 := [trans #1449 #1452]: #1453
+#1457 := [quant-intro #1454]: #1456
+#1647 := [monotonicity #1457 #1644]: #1646
+#1655 := [trans #1647 #1653]: #1654
+#1658 := [monotonicity #1655]: #1657
+#1434 := (iff #697 #1433)
+#1431 := (iff #694 #1428)
+#1414 := (or #1413 #680)
+#1425 := (or #1414 #690)
+#1429 := (iff #1425 #1428)
+#1430 := [rewrite]: #1429
+#1426 := (iff #694 #1425)
+#1423 := (iff #685 #1414)
+#1415 := (not #1414)
+#1418 := (not #1415)
+#1421 := (iff #1418 #1414)
+#1422 := [rewrite]: #1421
+#1419 := (iff #685 #1418)
+#1416 := (iff #682 #1415)
+#1417 := [rewrite]: #1416
+#1420 := [monotonicity #1417]: #1419
+#1424 := [trans #1420 #1422]: #1423
+#1427 := [monotonicity #1424]: #1426
+#1432 := [trans #1427 #1430]: #1431
+#1435 := [quant-intro #1432]: #1434
+#1661 := [monotonicity #1435 #1658]: #1660
+#1669 := [trans #1661 #1667]: #1668
+#1411 := (iff #1045 #1410)
+#1408 := (iff #1044 #1405)
+#1180 := (or #1179 #1042)
+#1402 := (or #1180 #1033)
+#1406 := (iff #1402 #1405)
+#1407 := [rewrite]: #1406
+#1403 := (iff #1044 #1402)
+#1400 := (iff #1037 #1180)
+#1126 := (not #1180)
+#1049 := (not #1126)
+#1244 := (iff #1049 #1180)
+#1399 := [rewrite]: #1244
+#1105 := (iff #1037 #1049)
+#1127 := (iff #1036 #1126)
+#1048 := [rewrite]: #1127
+#1106 := [monotonicity #1048]: #1105
+#1401 := [trans #1106 #1399]: #1400
+#1404 := [monotonicity #1401]: #1403
+#1409 := [trans #1404 #1407]: #1408
+#1412 := [monotonicity #1409]: #1411
+#1672 := [monotonicity #1412 #1669]: #1671
+#1173 := (+ #1172 #808)
+#1174 := (<= #1173 0::Int)
+#1167 := (+ ?v0!3 #797)
+#1168 := (>= #1167 0::Int)
+#1169 := (not #1168)
+#1170 := (and #1166 #1169)
+#1171 := (not #1170)
+#1175 := (or #1171 #1174)
+#1176 := (not #1175)
+#1195 := (or #1176 #1191)
+#1162 := (not #793)
+#1159 := (not #832)
+#1156 := (not #838)
+#1208 := (not #453)
+#1205 := (not #462)
+#1062 := (not #712)
+#1211 := (and #1062 #850 #1205 #1208 #1156 #1159 #1162 #1195)
+#1153 := (not #386)
+#1150 := (not #395)
+#1147 := (not #841)
+#1144 := (not #420)
+#1199 := (and #1062 #844 #1144 #1147 #1150 #1153 #1156 #1159 #1162 #1195)
+#1215 := (or #1199 #1211)
+#1219 := (and #1062 #758 #1215)
+#1119 := (+ #1118 #736)
+#1120 := (<= #1119 0::Int)
+#1121 := (or #1117 #1120)
+#1122 := (not #1121)
+#1128 := (and #1102 #1122)
+#1132 := (or #1089 #1128)
+#1083 := (not #235)
+#1080 := (not #244)
+#1077 := (not #253)
+#1136 := (and #1062 #918 #1077 #1080 #1083 #1132)
+#1223 := (or #1136 #1219)
+#1072 := (not #556)
+#1059 := (not #589)
+#1227 := (and #1059 #1062 #961 #1072 #1223)
+#1231 := (or #589 #1227)
+#1235 := (and #697 #1231)
+#1239 := (or #1045 #1235)
+#1397 := (iff #1239 #1396)
+#1394 := (iff #1235 #1393)
+#1391 := (iff #1231 #1390)
+#1388 := (iff #1227 #1385)
+#1382 := (and #28 #709 #961 #41 #1379)
+#1386 := (iff #1382 #1385)
+#1387 := [rewrite]: #1386
+#1383 := (iff #1227 #1382)
+#1380 := (iff #1223 #1379)
+#1377 := (iff #1219 #1374)
+#1371 := (and #709 #758 #1368)
+#1375 := (iff #1371 #1374)
+#1376 := [rewrite]: #1375
+#1372 := (iff #1219 #1371)
+#1369 := (iff #1215 #1368)
+#1366 := (iff #1211 #1363)
+#1360 := (and #709 #845 #114 #115 #835 #828 #790 #1345)
+#1364 := (iff #1360 #1363)
+#1365 := [rewrite]: #1364
+#1361 := (iff #1211 #1360)
+#1346 := (iff #1195 #1345)
+#1343 := (iff #1176 #1342)
+#1340 := (iff #1175 #1339)
+#1337 := (iff #1174 #1334)
+#1326 := (+ #808 #1172)
+#1329 := (<= #1326 0::Int)
+#1335 := (iff #1329 #1334)
+#1336 := [rewrite]: #1335
+#1330 := (iff #1174 #1329)
+#1327 := (= #1173 #1326)
+#1328 := [rewrite]: #1327
+#1331 := [monotonicity #1328]: #1330
+#1338 := [trans #1331 #1336]: #1337
+#1324 := (iff #1171 #1323)
+#1321 := (iff #1170 #1320)
+#1318 := (iff #1169 #1317)
+#1315 := (iff #1168 #1312)
+#1304 := (+ #797 ?v0!3)
+#1307 := (>= #1304 0::Int)
+#1313 := (iff #1307 #1312)
+#1314 := [rewrite]: #1313
+#1308 := (iff #1168 #1307)
+#1305 := (= #1167 #1304)
+#1306 := [rewrite]: #1305
+#1309 := [monotonicity #1306]: #1308
+#1316 := [trans #1309 #1314]: #1315
+#1319 := [monotonicity #1316]: #1318
+#1322 := [monotonicity #1319]: #1321
+#1325 := [monotonicity #1322]: #1324
+#1341 := [monotonicity #1325 #1338]: #1340
+#1344 := [monotonicity #1341]: #1343
+#1347 := [monotonicity #1344]: #1346
+#1302 := (iff #1162 #790)
+#1303 := [rewrite]: #1302
+#1300 := (iff #1159 #828)
+#1301 := [rewrite]: #1300
+#1298 := (iff #1156 #835)
+#1299 := [rewrite]: #1298
+#1358 := (iff #1208 #115)
+#1359 := [rewrite]: #1358
+#1356 := (iff #1205 #114)
+#1357 := [rewrite]: #1356
+#1247 := (iff #1062 #709)
+#1248 := [rewrite]: #1247
+#1362 := [monotonicity #1248 #854 #1357 #1359 #1299 #1301 #1303 #1347]: #1361
+#1367 := [trans #1362 #1365]: #1366
+#1354 := (iff #1199 #1351)
+#1348 := (and #709 #844 #74 #707 #77 #79 #835 #828 #790 #1345)
+#1352 := (iff #1348 #1351)
+#1353 := [rewrite]: #1352
+#1349 := (iff #1199 #1348)
+#1296 := (iff #1153 #79)
+#1297 := [rewrite]: #1296
+#1294 := (iff #1150 #77)
+#1295 := [rewrite]: #1294
+#1292 := (iff #1147 #707)
+#1293 := [rewrite]: #1292
+#1290 := (iff #1144 #74)
+#1291 := [rewrite]: #1290
+#1350 := [monotonicity #1248 #1291 #1293 #1295 #1297 #1299 #1301 #1303 #1347]: #1349
+#1355 := [trans #1350 #1353]: #1354
+#1370 := [monotonicity #1355 #1367]: #1369
+#1373 := [monotonicity #1248 #1370]: #1372
+#1378 := [trans #1373 #1376]: #1377
+#1288 := (iff #1136 #1285)
+#1282 := (and #709 #755 #44 #46 #48 #1279)
+#1286 := (iff #1282 #1285)
+#1287 := [rewrite]: #1286
+#1283 := (iff #1136 #1282)
+#1280 := (iff #1132 #1279)
+#1277 := (iff #1128 #1276)
+#1274 := (iff #1122 #1273)
+#1271 := (iff #1121 #1270)
+#1268 := (iff #1120 #1265)
+#1257 := (+ #736 #1118)
+#1260 := (<= #1257 0::Int)
+#1266 := (iff #1260 #1265)
+#1267 := [rewrite]: #1266
+#1261 := (iff #1120 #1260)
+#1258 := (= #1119 #1257)
+#1259 := [rewrite]: #1258
+#1262 := [monotonicity #1259]: #1261
+#1269 := [trans #1262 #1267]: #1268
+#1272 := [monotonicity #1269]: #1271
+#1275 := [monotonicity #1272]: #1274
+#1278 := [monotonicity #1275]: #1277
+#1281 := [monotonicity #1278]: #1280
+#1255 := (iff #1083 #48)
+#1256 := [rewrite]: #1255
+#1253 := (iff #1080 #46)
+#1254 := [rewrite]: #1253
+#1251 := (iff #1077 #44)
+#1252 := [rewrite]: #1251
+#1284 := [monotonicity #1248 #922 #1252 #1254 #1256 #1281]: #1283
+#1289 := [trans #1284 #1287]: #1288
+#1381 := [monotonicity #1289 #1378]: #1380
+#1249 := (iff #1072 #41)
+#1250 := [rewrite]: #1249
+#1245 := (iff #1059 #28)
+#1246 := [rewrite]: #1245
+#1384 := [monotonicity #1246 #1248 #1250 #1381]: #1383
+#1389 := [trans #1384 #1387]: #1388
+#1392 := [monotonicity #1389]: #1391
+#1395 := [monotonicity #1392]: #1394
+#1398 := [monotonicity #1395]: #1397
+#1029 := (not #993)
+#1240 := (~ #1029 #1239)
+#1236 := (not #990)
+#1237 := (~ #1236 #1235)
+#1232 := (not #987)
+#1233 := (~ #1232 #1231)
+#1228 := (not #982)
+#1229 := (~ #1228 #1227)
+#1224 := (not #939)
+#1225 := (~ #1224 #1223)
+#1220 := (not #934)
+#1221 := (~ #1220 #1219)
+#1216 := (not #913)
+#1217 := (~ #1216 #1215)
+#1212 := (not #908)
+#1213 := (~ #1212 #1211)
+#1196 := (not #825)
+#1197 := (~ #1196 #1195)
+#1192 := (not #822)
+#1193 := (~ #1192 #1191)
+#1189 := (~ #1188 #1188)
+#1190 := [refl]: #1189
+#1185 := (not #819)
+#1186 := (~ #1185 #816)
+#1183 := (~ #816 #816)
+#1181 := (~ #813 #813)
+#1182 := [refl]: #1181
+#1184 := [nnf-pos #1182]: #1183
+#1187 := [nnf-neg #1184]: #1186
+#1194 := [nnf-neg #1187 #1190]: #1193
+#1177 := (~ #819 #1176)
+#1178 := [sk]: #1177
+#1198 := [nnf-neg #1178 #1194]: #1197
+#1163 := (~ #1162 #1162)
+#1164 := [refl]: #1163
+#1160 := (~ #1159 #1159)
+#1161 := [refl]: #1160
+#1157 := (~ #1156 #1156)
+#1158 := [refl]: #1157
+#1209 := (~ #1208 #1208)
+#1210 := [refl]: #1209
+#1206 := (~ #1205 #1205)
+#1207 := [refl]: #1206
+#1203 := (~ #850 #850)
+#1204 := [refl]: #1203
+#1063 := (~ #1062 #1062)
+#1064 := [refl]: #1063
+#1214 := [nnf-neg #1064 #1204 #1207 #1210 #1158 #1161 #1164 #1198]: #1213
+#1200 := (not #884)
+#1201 := (~ #1200 #1199)
+#1154 := (~ #1153 #1153)
+#1155 := [refl]: #1154
+#1151 := (~ #1150 #1150)
+#1152 := [refl]: #1151
+#1148 := (~ #1147 #1147)
+#1149 := [refl]: #1148
+#1145 := (~ #1144 #1144)
+#1146 := [refl]: #1145
+#1142 := (~ #844 #844)
+#1143 := [refl]: #1142
+#1202 := [nnf-neg #1064 #1143 #1146 #1149 #1152 #1155 #1158 #1161 #1164 #1198]: #1201
+#1218 := [nnf-neg #1202 #1214]: #1217
+#1140 := (~ #758 #758)
+#1141 := [refl]: #1140
+#1222 := [nnf-neg #1064 #1141 #1218]: #1221
+#1137 := (not #779)
+#1138 := (~ #1137 #1136)
+#1133 := (not #750)
+#1134 := (~ #1133 #1132)
+#1129 := (not #747)
+#1130 := (~ #1129 #1128)
+#1123 := (not #744)
+#1124 := (~ #1123 #1122)
+#1125 := [sk]: #1124
+#1107 := (not #733)
+#1108 := (~ #1107 #1102)
+#1103 := (~ #730 #1102)
+#1104 := [sk]: #1103
+#1109 := [nnf-neg #1104]: #1108
+#1131 := [nnf-neg #1109 #1125]: #1130
+#1090 := (~ #733 #1089)
+#1087 := (~ #1086 #1086)
+#1088 := [refl]: #1087
+#1091 := [nnf-neg #1088]: #1090
+#1135 := [nnf-neg #1091 #1131]: #1134
+#1084 := (~ #1083 #1083)
+#1085 := [refl]: #1084
+#1081 := (~ #1080 #1080)
+#1082 := [refl]: #1081
+#1078 := (~ #1077 #1077)
+#1079 := [refl]: #1078
+#1075 := (~ #918 #918)
+#1076 := [refl]: #1075
+#1139 := [nnf-neg #1064 #1076 #1079 #1082 #1085 #1135]: #1138
+#1226 := [nnf-neg #1139 #1222]: #1225
+#1073 := (~ #1072 #1072)
+#1074 := [refl]: #1073
+#1069 := (not #964)
+#1070 := (~ #1069 #961)
+#1067 := (~ #961 #961)
+#1065 := (~ #958 #958)
+#1066 := [refl]: #1065
+#1068 := [nnf-pos #1066]: #1067
+#1071 := [nnf-neg #1068]: #1070
+#1060 := (~ #1059 #1059)
+#1061 := [refl]: #1060
+#1230 := [nnf-neg #1061 #1064 #1071 #1074 #1226]: #1229
+#1057 := (~ #589 #589)
+#1058 := [refl]: #1057
+#1234 := [nnf-neg #1058 #1230]: #1233
+#1054 := (not #700)
+#1055 := (~ #1054 #697)
+#1052 := (~ #697 #697)
+#1050 := (~ #694 #694)
+#1051 := [refl]: #1050
+#1053 := [nnf-pos #1051]: #1052
+#1056 := [nnf-neg #1053]: #1055
+#1238 := [nnf-neg #1056 #1234]: #1237
+#1046 := (~ #700 #1045)
+#1047 := [sk]: #1046
+#1241 := [nnf-neg #1047 #1238]: #1240
+#1030 := [not-or-elim #1026]: #1029
+#1242 := [mp~ #1030 #1241]: #1239
+#1243 := [mp #1242 #1398]: #1396
+#1673 := [mp #1243 #1672]: #1670
+#2253 := [mp #1673 #2252]: #2250
+#1748 := [unit-resolution #2253 #1940]: #2247
+#2018 := (or #2244 #2238)
+#2019 := [def-axiom]: #2018
+#1744 := [unit-resolution #2019 #1748]: #2238
+#1740 := (or #2241 #2235)
+#1738 := (iff #13 #28)
+#1749 := (iff #28 #13)
+#1736 := [commutativity]: #1749
+#1739 := [symm #1736]: #1738
+#1737 := [mp #1028 #1739]: #28
+#2017 := (or #2241 #589 #2235)
+#2013 := [def-axiom]: #2017
+#2254 := [unit-resolution #2013 #1737]: #1740
+#2255 := [unit-resolution #2254 #1744]: #2235
+#2024 := (or #2232 #2226)
+#2026 := [def-axiom]: #2024
+#2335 := [unit-resolution #2026 #2255]: #2226
+#2293 := [hypothesis]: #2179
+#1841 := (or #2176 #2170)
+#2115 := [def-axiom]: #1841
+#2294 := [unit-resolution #2115 #2293]: #2170
+#2130 := (not #2165)
+#1741 := (or #2176 #46)
+#2117 := [def-axiom]: #1741
+#2295 := [unit-resolution #2117 #2293]: #46
+#2328 := (or #2130 #244)
+#2303 := (= #40 f11)
+#2300 := (* -1::Int f7)
+#2301 := (+ f3 #2300)
+#2302 := (<= #2301 0::Int)
+#2304 := (or #1542 #2302 #2303)
+#1857 := (= f9 f11)
+#2306 := [hypothesis]: #46
+#2323 := [symm #2306]: #1857
+#2045 := (or #2232 #41)
+#2023 := [def-axiom]: #2045
+#2307 := [unit-resolution #2023 #2255]: #41
+#2324 := [trans #2307 #2323]: #2303
+#2318 := (not #2303)
+#2319 := (or #2304 #2318)
+#2320 := [def-axiom]: #2319
+#2325 := [unit-resolution #2320 #2324]: #2304
+#2326 := [hypothesis]: #2165
+#2305 := (not #2304)
+#2308 := (or #2130 #2305)
+#2309 := [quant-inst #29]: #2308
+#2327 := [unit-resolution #2309 #2326 #2325]: false
+#2329 := [lemma #2327]: #2328
+#2296 := [unit-resolution #2329 #2295]: #2130
+#1774 := (or #2173 #2165 #1531)
+#2134 := [def-axiom]: #1774
+#2297 := [unit-resolution #2134 #2296 #2294]: #1531
+#1787 := (or #1530 #1111)
+#1788 := [def-axiom]: #1787
+#2298 := [unit-resolution #1788 #2297]: #1111
+#1816 := (+ f8 #1112)
+#1817 := (<= #1816 0::Int)
+#1767 := (not #1817)
+#1844 := (or #2176 #755)
+#1845 := [def-axiom]: #1844
+#2299 := [unit-resolution #1845 #2293]: #755
+#1789 := (or #1530 #1115)
+#2125 := [def-axiom]: #1789
+#2282 := [unit-resolution #2125 #2297]: #1115
+#1750 := (or #1767 #758 #1114)
+#1763 := [hypothesis]: #1115
+#1764 := [hypothesis]: #1817
+#1766 := [hypothesis]: #755
+#1762 := [th-lemma arith farkas -1 -1 1 #1766 #1764 #1763]: false
+#1753 := [lemma #1762]: #1750
+#2283 := [unit-resolution #1753 #2282 #2299]: #1767
+#1793 := (+ f9 #1263)
+#1794 := (>= #1793 0::Int)
+#1752 := (not #1794)
+#1859 := (+ f9 #736)
+#1848 := (<= #1859 0::Int)
+#2281 := [symm #2295]: #1857
+#2284 := (not #1857)
+#2285 := (or #2284 #1848)
+#2286 := [th-lemma arith triangle-eq]: #2285
+#2287 := [unit-resolution #2286 #2281]: #1848
+#2126 := (not #1265)
+#2127 := (or #1530 #2126)
+#2128 := [def-axiom]: #2127
+#2288 := [unit-resolution #2128 #2297]: #2126
+#1760 := (not #1848)
+#1745 := (or #1752 #1265 #1760)
+#1754 := [hypothesis]: #1848
+#1757 := [hypothesis]: #2126
+#1758 := [hypothesis]: #1794
+#1759 := [th-lemma arith farkas -1 1 1 #1758 #1757 #1754]: false
+#1742 := [lemma #1759]: #1745
+#2289 := [unit-resolution #1742 #2288 #2287]: #1752
+#1779 := (or #1503 #1817 #1794)
+#1743 := [hypothesis]: #1752
+#1746 := [hypothesis]: #1767
+#1747 := [hypothesis]: #1111
+#2044 := (or #2232 #2157)
+#2034 := [def-axiom]: #2044
+#2256 := [unit-resolution #2034 #2255]: #2157
+#1835 := (or #2162 #1503 #1817 #1794)
+#1829 := (+ #1118 #953)
+#1839 := (<= #1829 0::Int)
+#1843 := (+ ?v0!2 #753)
+#1853 := (>= #1843 0::Int)
+#1806 := (or #1503 #1853 #1839)
+#1836 := (or #2162 #1806)
+#1755 := (iff #1836 #1835)
+#1775 := (or #2162 #1779)
+#1776 := (iff #1775 #1835)
+#1751 := [rewrite]: #1776
+#1770 := (iff #1836 #1775)
+#1780 := (iff #1806 #1779)
+#1796 := (iff #1839 #1794)
+#1804 := (+ #953 #1118)
+#1790 := (<= #1804 0::Int)
+#1795 := (iff #1790 #1794)
+#1784 := [rewrite]: #1795
+#1791 := (iff #1839 #1790)
+#1783 := (= #1829 #1804)
+#1785 := [rewrite]: #1783
+#1792 := [monotonicity #1785]: #1791
+#1777 := [trans #1792 #1784]: #1796
+#1801 := (iff #1853 #1817)
+#1808 := (+ #753 ?v0!2)
+#1813 := (>= #1808 0::Int)
+#1807 := (iff #1813 #1817)
+#1818 := [rewrite]: #1807
+#1814 := (iff #1853 #1813)
+#1809 := (= #1843 #1808)
+#1800 := [rewrite]: #1809
+#1815 := [monotonicity #1800]: #1814
+#1803 := [trans #1815 #1818]: #1801
+#1778 := [monotonicity #1803 #1777]: #1780
+#1781 := [monotonicity #1778]: #1770
+#1756 := [trans #1781 #1751]: #1755
+#1772 := [quant-inst #1110]: #1836
+#1761 := [mp #1772 #1756]: #1835
+#2257 := [unit-resolution #1761 #2256 #1747 #1746 #1743]: false
+#2258 := [lemma #2257]: #1779
+#2313 := [unit-resolution #2258 #2289 #2283 #2298]: false
+#2314 := [lemma #2313]: #2176
+#2036 := (or #2229 #2179 #2223)
+#2037 := [def-axiom]: #2036
+#2333 := [unit-resolution #2037 #2314 #2335]: #2223
+#2049 := (or #2220 #2214)
+#2050 := [def-axiom]: #2049
+#2499 := [unit-resolution #2050 #2333]: #2214
+#2420 := (= #71 #1172)
+#2434 := (not #2420)
+#2421 := (+ #71 #1332)
+#2423 := (>= #2421 0::Int)
+#2428 := (not #2423)
+#2322 := (+ #71 #808)
+#2290 := (<= #2322 0::Int)
+#2321 := (= #71 f15)
+#2455 := (= f13 f15)
+#2450 := [hypothesis]: #2205
+#1931 := (or #2202 #79)
+#2084 := [def-axiom]: #1931
+#2451 := [unit-resolution #2084 #2450]: #79
+#2456 := [symm #2451]: #2455
+#2453 := (= #71 f13)
+#2092 := (or #2202 #74)
+#2099 := [def-axiom]: #2092
+#2452 := [unit-resolution #2099 #2450]: #74
+#2454 := [symm #2452]: #2453
+#2457 := [trans #2454 #2456]: #2321
+#2458 := (not #2321)
+#2459 := (or #2458 #2290)
+#2460 := [th-lemma arith triangle-eq]: #2459
+#2461 := [unit-resolution #2460 #2457]: #2290
+#2112 := (not #1334)
+#1932 := (or #2202 #2196)
+#2080 := [def-axiom]: #1932
+#2462 := [unit-resolution #2080 #2450]: #2196
+#2466 := (= #93 f13)
+#2464 := (= #93 #71)
+#1928 := (or #2202 #77)
+#1930 := [def-axiom]: #1928
+#2463 := [unit-resolution #1930 #2450]: #77
+#2465 := [monotonicity #2463]: #2464
+#2467 := [trans #2465 #2454]: #2466
+#2468 := [trans #2467 #2456]: #94
+#2104 := (or #2190 #1188)
+#2100 := [def-axiom]: #2104
+#2469 := [unit-resolution #2100 #2468]: #2190
+#2095 := (or #2199 #1570 #2193)
+#2096 := [def-axiom]: #2095
+#2470 := [unit-resolution #2096 #2469 #2462]: #1570
+#1827 := (or #1565 #2112)
+#2109 := [def-axiom]: #1827
+#2471 := [unit-resolution #2109 #2470]: #2112
+#2429 := (not #2290)
+#2430 := (or #2428 #1334 #2429)
+#2424 := [hypothesis]: #2423
+#2425 := [hypothesis]: #2290
+#2426 := [hypothesis]: #2112
+#2427 := [th-lemma arith farkas -1 -1 1 #2426 #2425 #2424]: false
+#2431 := [lemma #2427]: #2430
+#2472 := [unit-resolution #2431 #2471 #2461]: #2428
+#2435 := (or #2434 #2423)
+#2436 := [th-lemma arith triangle-eq]: #2435
+#2473 := [unit-resolution #2436 #2472]: #2434
+#2422 := (= f8 ?v0!3)
+#2088 := (or #2202 #828)
+#2086 := [def-axiom]: #2088
+#2474 := [unit-resolution #2086 #2450]: #828
+#2475 := (or #832 #1830)
+#2476 := [th-lemma arith triangle-eq]: #2475
+#2477 := [unit-resolution #2476 #2474]: #1830
+#1833 := (or #1565 #1317)
+#2111 := [def-axiom]: #1833
+#2478 := [unit-resolution #2111 #2470]: #1317
+#2479 := (not #1830)
+#2480 := (or #2419 #1312 #2479)
+#2481 := [th-lemma arith assign-bounds 1 1]: #2480
+#2482 := [unit-resolution #2481 #2478 #2477]: #2419
+#2398 := (+ f9 #1332)
+#2399 := (>= #2398 0::Int)
+#2484 := (not #2399)
+#2097 := (or #2202 #844)
+#2098 := [def-axiom]: #2097
+#2483 := [unit-resolution #2098 #2450]: #844
+#2485 := (or #2484 #1334 #2429 #845)
+#2486 := [th-lemma arith assign-bounds 1 1 1]: #2485
+#2487 := [unit-resolution #2486 #2471 #2461 #2483]: #2484
+#2489 := (or #2387 #2399)
+#1831 := (or #1565 #1166)
+#1832 := [def-axiom]: #1831
+#2488 := [unit-resolution #1832 #2470]: #1166
+#2407 := (or #2162 #1550 #2387 #2399)
+#2377 := (+ #1172 #953)
+#2378 := (<= #2377 0::Int)
+#2369 := (+ ?v0!3 #753)
+#2370 := (>= #2369 0::Int)
+#2379 := (or #1550 #2370 #2378)
+#2408 := (or #2162 #2379)
+#2415 := (iff #2408 #2407)
+#2404 := (or #1550 #2387 #2399)
+#2410 := (or #2162 #2404)
+#2413 := (iff #2410 #2407)
+#2414 := [rewrite]: #2413
+#2411 := (iff #2408 #2410)
+#2405 := (iff #2379 #2404)
+#2402 := (iff #2378 #2399)
+#2392 := (+ #953 #1172)
+#2395 := (<= #2392 0::Int)
+#2400 := (iff #2395 #2399)
+#2401 := [rewrite]: #2400
+#2396 := (iff #2378 #2395)
+#2393 := (= #2377 #2392)
+#2394 := [rewrite]: #2393
+#2397 := [monotonicity #2394]: #2396
+#2403 := [trans #2397 #2401]: #2402
+#2390 := (iff #2370 #2387)
+#2380 := (+ #753 ?v0!3)
+#2383 := (>= #2380 0::Int)
+#2388 := (iff #2383 #2387)
+#2389 := [rewrite]: #2388
+#2384 := (iff #2370 #2383)
+#2381 := (= #2369 #2380)
+#2382 := [rewrite]: #2381
+#2385 := [monotonicity #2382]: #2384
+#2391 := [trans #2385 #2389]: #2390
+#2406 := [monotonicity #2391 #2403]: #2405
+#2412 := [monotonicity #2406]: #2411
+#2416 := [trans #2412 #2414]: #2415
+#2409 := [quant-inst #1165]: #2408
+#2417 := [mp #2409 #2416]: #2407
+#2490 := [unit-resolution #2417 #2256 #2488]: #2489
+#2491 := [unit-resolution #2490 #2487]: #2387
+#2493 := (not #2419)
+#2494 := (or #2422 #2492 #2493)
+#2495 := [th-lemma arith triangle-eq]: #2494
+#2496 := [unit-resolution #2495 #2491 #2482]: #2422
+#2447 := (not #2422)
+#2448 := (or #2447 #2420)
+#2443 := (= #1172 #71)
+#2441 := (= ?v0!3 f8)
+#2440 := [hypothesis]: #2422
+#2442 := [symm #2440]: #2441
+#2444 := [monotonicity #2442]: #2443
+#2445 := [symm #2444]: #2420
+#2439 := [hypothesis]: #2434
+#2446 := [unit-resolution #2439 #2445]: false
+#2449 := [lemma #2446]: #2448
+#2497 := [unit-resolution #2449 #2496 #2473]: false
+#2498 := [lemma #2497]: #2202
+#2056 := (or #2217 #2205 #2211)
+#2057 := [def-axiom]: #2056
+#2500 := [unit-resolution #2057 #2498 #2499]: #2211
+#2063 := (or #2208 #828)
+#2073 := [def-axiom]: #2063
+#2501 := [unit-resolution #2073 #2500]: #828
+#2502 := [unit-resolution #2476 #2501]: #1830
+#2065 := (or #2208 #2196)
+#2066 := [def-axiom]: #2065
+#2503 := [unit-resolution #2066 #2500]: #2196
+#1984 := (= f9 f15)
+#2070 := (or #2208 #115)
+#2072 := [def-axiom]: #2070
+#2504 := [unit-resolution #2072 #2500]: #115
+#2508 := [symm #2504]: #1984
+#2509 := (= #93 f9)
+#2506 := (= #93 #40)
+#2079 := (or #2208 #114)
+#2083 := [def-axiom]: #2079
+#2505 := [unit-resolution #2083 #2500]: #114
+#2507 := [monotonicity #2505]: #2506
+#2510 := [trans #2507 #2307]: #2509
+#2511 := [trans #2510 #2508]: #94
+#2512 := [unit-resolution #2100 #2511]: #2190
+#2513 := [unit-resolution #2096 #2512 #2503]: #1570
+#2514 := [unit-resolution #2111 #2513]: #1317
+#2515 := [unit-resolution #2481 #2514 #2502]: #2419
+#1989 := (or #2208 #845)
+#2082 := [def-axiom]: #1989
+#2516 := [unit-resolution #2082 #2500]: #845
+#1977 := (+ f9 #808)
+#1985 := (<= #1977 0::Int)
+#1880 := (or #453 #1984)
+#1876 := (iff #115 #1984)
+#1874 := (iff #1984 #115)
+#1875 := [commutativity]: #1874
+#1877 := [symm #1875]: #1876
+#1873 := [hypothesis]: #115
+#1878 := [mp #1873 #1877]: #1984
+#1870 := (not #1984)
+#1872 := [hypothesis]: #1870
+#1879 := [unit-resolution #1872 #1878]: false
+#1881 := [lemma #1879]: #1880
+#2517 := [unit-resolution #1881 #2504]: #1984
+#2518 := (or #1870 #1985)
+#2519 := [th-lemma arith triangle-eq]: #2518
+#2520 := [unit-resolution #2519 #2517]: #1985
+#2521 := (not #1985)
+#2522 := (or #2290 #844 #2521)
+#2523 := [th-lemma arith assign-bounds 1 -1]: #2522
+#2524 := [unit-resolution #2523 #2520 #2516]: #2290
+#2525 := [unit-resolution #2109 #2513]: #2112
+#2526 := [unit-resolution #2431 #2525 #2524]: #2428
+#2527 := [unit-resolution #2436 #2526]: #2434
+#2528 := [unit-resolution #2449 #2527]: #2447
+#2529 := [unit-resolution #2495 #2528 #2515]: #2492
+#2530 := (or #2484 #1334 #2521)
+#2531 := [th-lemma arith assign-bounds -1 -1]: #2530
+#2532 := [unit-resolution #2531 #2520 #2525]: #2484
+#2533 := [unit-resolution #1832 #2513]: #1166
+[unit-resolution #2417 #2256 #2533 #2532 #2529]: false
+unsat
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/SMT_Examples/VCC_Max.b2i	Tue Jul 23 18:36:23 2013 +0200
@@ -0,0 +1,19619 @@
+type-decl $ctype 0 0
+type-decl $ptr 0 0
+type-decl $field 0 0
+type-decl $kind 0 0
+type-decl $type_state 0 0
+type-decl $status 0 0
+type-decl $primitive 0 0
+type-decl $struct 0 0
+type-decl $token 0 0
+type-decl $state 0 0
+type-decl $pure_function 0 0
+type-decl $label 0 0
+type-decl $memory_t 0 0
+type-decl $typemap_t 0 0
+type-decl $statusmap_t 0 0
+type-decl $record 0 0
+type-decl $version 0 0
+type-decl $vol_version 0 0
+type-decl $ptrset 0 0
+fun-decl $kind_of 2 0
+    type-con $ctype 0
+    type-con $kind 0
+fun-decl $kind_composite 1 1
+    type-con $kind 0
+  attribute unique 0
+fun-decl $kind_primitive 1 1
+    type-con $kind 0
+  attribute unique 0
+fun-decl $kind_array 1 1
+    type-con $kind 0
+  attribute unique 0
+fun-decl $kind_thread 1 1
+    type-con $kind 0
+  attribute unique 0
+fun-decl $sizeof 2 0
+    type-con $ctype 0
+    int
+fun-decl ^^i1 1 1
+    type-con $ctype 0
+  attribute unique 0
+fun-decl ^^i2 1 1
+    type-con $ctype 0
+  attribute unique 0
+fun-decl ^^i4 1 1
+    type-con $ctype 0
+  attribute unique 0
+fun-decl ^^i8 1 1
+    type-con $ctype 0
+  attribute unique 0
+fun-decl ^^u1 1 1
+    type-con $ctype 0
+  attribute unique 0
+fun-decl ^^u2 1 1
+    type-con $ctype 0
+  attribute unique 0
+fun-decl ^^u4 1 1
+    type-con $ctype 0
+  attribute unique 0
+fun-decl ^^u8 1 1
+    type-con $ctype 0
+  attribute unique 0
+fun-decl ^^void 1 1
+    type-con $ctype 0
+  attribute unique 0
+fun-decl ^^bool 1 1
+    type-con $ctype 0
+  attribute unique 0
+fun-decl ^^f4 1 1
+    type-con $ctype 0
+  attribute unique 0
+fun-decl ^^f8 1 1
+    type-con $ctype 0
+  attribute unique 0
+fun-decl ^^claim 1 1
+    type-con $ctype 0
+  attribute unique 0
+fun-decl ^^root_emb 1 1
+    type-con $ctype 0
+  attribute unique 0
+fun-decl ^^mathint 1 1
+    type-con $ctype 0
+  attribute unique 0
+fun-decl ^$#thread_id_t 1 1
+    type-con $ctype 0
+  attribute unique 0
+fun-decl ^$#ptrset 1 1
+    type-con $ctype 0
+  attribute unique 0
+fun-decl ^$#state_t 1 1
+    type-con $ctype 0
+  attribute unique 0
+fun-decl ^$#struct 1 1
+    type-con $ctype 0
+  attribute unique 0
+fun-decl $ptr_to 2 0
+    type-con $ctype 0
+    type-con $ctype 0
+fun-decl $unptr_to 2 0
+    type-con $ctype 0
+    type-con $ctype 0
+fun-decl $ptr_level 2 0
+    type-con $ctype 0
+    int
+fun-decl $map_t 3 0
+    type-con $ctype 0
+    type-con $ctype 0
+    type-con $ctype 0
+fun-decl $map_domain 2 0
+    type-con $ctype 0
+    type-con $ctype 0
+fun-decl $map_range 2 0
+    type-con $ctype 0
+    type-con $ctype 0
+fun-decl $is_primitive 2 1
+    type-con $ctype 0
+    bool
+  attribute weight 1
+    expr-attr
+      int-num 0
+fun-decl $is_primitive_ch 2 1
+    type-con $ctype 0
+    bool
+  attribute inline 1
+    expr-attr
+      true
+fun-decl $is_composite 2 1
+    type-con $ctype 0
+    bool
+  attribute weight 1
+    expr-attr
+      int-num 0
+fun-decl $is_composite_ch 2 1
+    type-con $ctype 0
+    bool
+  attribute inline 1
+    expr-attr
+      true
+fun-decl $is_arraytype 2 1
+    type-con $ctype 0
+    bool
+  attribute weight 1
+    expr-attr
+      int-num 0
+fun-decl $is_arraytype_ch 2 1
+    type-con $ctype 0
+    bool
+  attribute inline 1
+    expr-attr
+      true
+fun-decl $is_threadtype 2 1
+    type-con $ctype 0
+    bool
+  attribute weight 1
+    expr-attr
+      int-num 0
+fun-decl $is_thread 2 1
+    type-con $ptr 0
+    bool
+  attribute inline 1
+    expr-attr
+      true
+fun-decl $is_ptr_to_composite 2 1
+    type-con $ptr 0
+    bool
+  attribute inline 1
+    expr-attr
+      true
+fun-decl $field_offset 2 0
+    type-con $field 0
+    int
+fun-decl $field_parent_type 2 0
+    type-con $field 0
+    type-con $ctype 0
+fun-decl $is_non_primitive 2 0
+    type-con $ctype 0
+    bool
+fun-decl $is_non_primitive_ch 2 1
+    type-con $ctype 0
+    bool
+  attribute inline 1
+    expr-attr
+      true
+fun-decl $is_non_primitive_ptr 2 1
+    type-con $ptr 0
+    bool
+  attribute inline 1
+    expr-attr
+      true
+fun-decl $me_ref 1 0
+    int
+fun-decl $me 1 0
+    type-con $ptr 0
+fun-decl $current_state 2 1
+    type-con $state 0
+    type-con $state 0
+  attribute inline 1
+    expr-attr
+      true
+fun-decl $select.mem 3 0
+    type-con $memory_t 0
+    type-con $ptr 0
+    int
+fun-decl $store.mem 4 0
+    type-con $memory_t 0
+    type-con $ptr 0
+    int
+    type-con $memory_t 0
+fun-decl $select.tm 3 0
+    type-con $typemap_t 0
+    type-con $ptr 0
+    type-con $type_state 0
+fun-decl $store.tm 4 0
+    type-con $typemap_t 0
+    type-con $ptr 0
+    type-con $type_state 0
+    type-con $typemap_t 0
+fun-decl $select.sm 3 0
+    type-con $statusmap_t 0
+    type-con $ptr 0
+    type-con $status 0
+fun-decl $store.sm 4 0
+    type-con $statusmap_t 0
+    type-con $ptr 0
+    type-con $status 0
+    type-con $statusmap_t 0
+fun-decl $memory 2 0
+    type-con $state 0
+    type-con $memory_t 0
+fun-decl $typemap 2 0
+    type-con $state 0
+    type-con $typemap_t 0
+fun-decl $statusmap 2 0
+    type-con $state 0
+    type-con $statusmap_t 0
+fun-decl $mem 3 1
+    type-con $state 0
+    type-con $ptr 0
+    int
+  attribute inline 1
+    expr-attr
+      true
+fun-decl $read_any 3 1
+    type-con $state 0
+    type-con $ptr 0
+    int
+  attribute inline 1
+    expr-attr
+      true
+fun-decl $mem_eq 4 1
+    type-con $state 0
+    type-con $state 0
+    type-con $ptr 0
+    bool
+  attribute inline 1
+    expr-attr
+      true
+fun-decl $st_eq 4 1
+    type-con $state 0
+    type-con $state 0
+    type-con $ptr 0
+    bool
+  attribute inline 1
+    expr-attr
+      true
+fun-decl $ts_eq 4 1
+    type-con $state 0
+    type-con $state 0
+    type-con $ptr 0
+    bool
+  attribute inline 1
+    expr-attr
+      true
+fun-decl $extent_hint 3 0
+    type-con $ptr 0
+    type-con $ptr 0
+    bool
+fun-decl $nesting_level 2 0
+    type-con $ctype 0
+    int
+fun-decl $is_nested 3 0
+    type-con $ctype 0
+    type-con $ctype 0
+    bool
+fun-decl $nesting_min 3 0
+    type-con $ctype 0
+    type-con $ctype 0
+    int
+fun-decl $nesting_max 3 0
+    type-con $ctype 0
+    type-con $ctype 0
+    int
+fun-decl $is_nested_range 5 0
+    type-con $ctype 0
+    type-con $ctype 0
+    int
+    int
+    bool
+fun-decl $typ 2 0
+    type-con $ptr 0
+    type-con $ctype 0
+fun-decl $ref 2 0
+    type-con $ptr 0
+    int
+fun-decl $ptr 3 0
+    type-con $ctype 0
+    int
+    type-con $ptr 0
+fun-decl $ghost_ref 3 0
+    type-con $ptr 0
+    type-con $field 0
+    int
+fun-decl $ghost_emb 2 0
+    int
+    type-con $ptr 0
+fun-decl $ghost_path 2 0
+    int
+    type-con $field 0
+fun-decl $physical_ref 3 0
+    type-con $ptr 0
+    type-con $field 0
+    int
+fun-decl $array_path 3 0
+    type-con $field 0
+    int
+    type-con $field 0
+fun-decl $is_base_field 2 0
+    type-con $field 0
+    bool
+fun-decl $array_path_1 2 0
+    type-con $field 0
+    type-con $field 0
+fun-decl $array_path_2 2 0
+    type-con $field 0
+    int
+fun-decl $null 1 0
+    type-con $ptr 0
+fun-decl $is 3 0
+    type-con $ptr 0
+    type-con $ctype 0
+    bool
+fun-decl $ptr_cast 3 1
+    type-con $ptr 0
+    type-con $ctype 0
+    type-con $ptr 0
+  attribute inline 1
+    expr-attr
+      true
+fun-decl $read_ptr 4 1
+    type-con $state 0
+    type-con $ptr 0
+    type-con $ctype 0
+    type-con $ptr 0
+  attribute inline 1
+    expr-attr
+      true
+fun-decl $dot 3 0
+    type-con $ptr 0
+    type-con $field 0
+    type-con $ptr 0
+fun-decl $emb 3 1
+    type-con $state 0
+    type-con $ptr 0
+    type-con $ptr 0
+  attribute inline 1
+    expr-attr
+      true
+fun-decl $path 3 1
+    type-con $state 0
+    type-con $ptr 0
+    type-con $field 0
+  attribute inline 1
+    expr-attr
+      true
+fun-decl $containing_struct 3 0
+    type-con $ptr 0
+    type-con $field 0
+    type-con $ptr 0
+fun-decl $containing_struct_ref 3 0
+    type-con $ptr 0
+    type-con $field 0
+    int
+fun-decl $is_primitive_non_volatile_field 2 0
+    type-con $field 0
+    bool
+fun-decl $is_primitive_volatile_field 2 0
+    type-con $field 0
+    bool
+fun-decl $is_primitive_embedded_array 3 0
+    type-con $field 0
+    int
+    bool
+fun-decl $is_primitive_embedded_volatile_array 4 0
+    type-con $field 0
+    int
+    type-con $ctype 0
+    bool
+fun-decl $static_field_properties 3 1
+    type-con $field 0
+    type-con $ctype 0
+    bool
+  attribute inline 1
+    expr-attr
+      true
+fun-decl $field_properties 6 1
+    type-con $state 0
+    type-con $ptr 0
+    type-con $field 0
+    type-con $ctype 0
+    bool
+    bool
+  attribute inline 1
+    expr-attr
+      true
+fun-decl $ts_typed 2 0
+    type-con $type_state 0
+    bool
+fun-decl $ts_emb 2 0
+    type-con $type_state 0
+    type-con $ptr 0
+fun-decl $ts_path 2 0
+    type-con $type_state 0
+    type-con $field 0
+fun-decl $ts_is_array_elt 2 0
+    type-con $type_state 0
+    bool
+fun-decl $ts_is_volatile 2 0
+    type-con $type_state 0
+    bool
+fun-decl $is_object_root 3 1
+    type-con $state 0
+    type-con $ptr 0
+    bool
+  attribute inline 1
+    expr-attr
+      true
+fun-decl $is_volatile 3 1
+    type-con $state 0
+    type-con $ptr 0
+    bool
+  attribute inline 1
+    expr-attr
+      true
+fun-decl $is_malloc_root 3 1
+    type-con $state 0
+    type-con $ptr 0
+    bool
+  attribute inline 1
+    expr-attr
+      true
+fun-decl $current_timestamp 2 0
+    type-con $state 0
+    int
+fun-decl $is_fresh 4 1
+    type-con $state 0
+    type-con $state 0
+    type-con $ptr 0
+    bool
+  attribute inline 1
+    expr-attr
+      true
+fun-decl $in_writes_at 3 0
+    int
+    type-con $ptr 0
+    bool
+fun-decl $writable 4 1
+    type-con $state 0
+    int
+    type-con $ptr 0
+    bool
+  attribute inline 1
+    expr-attr
+      true
+fun-decl $top_writable 4 1
+    type-con $state 0
+    int
+    type-con $ptr 0
+    bool
+  attribute inline 1
+    expr-attr
+      true
+fun-decl $struct_zero 1 0
+    type-con $struct 0
+fun-decl $vs_base 3 1
+    type-con $struct 0
+    type-con $ctype 0
+    type-con $ptr 0
+  attribute inline 1
+    expr-attr
+      true
+fun-decl $vs_base_ref 2 0
+    type-con $struct 0
+    int
+fun-decl $vs_state 2 0
+    type-con $struct 0
+    type-con $state 0
+fun-decl $vs_ctor 3 0
+    type-con $state 0
+    type-con $ptr 0
+    type-con $struct 0
+fun-decl $rec_zero 1 0
+    type-con $record 0
+fun-decl $rec_update 4 0
+    type-con $record 0
+    type-con $field 0
+    int
+    type-con $record 0
+fun-decl $rec_fetch 3 0
+    type-con $record 0
+    type-con $field 0
+    int
+fun-decl $rec_update_bv 7 0
+    type-con $record 0
+    type-con $field 0
+    int
+    int
+    int
+    int
+    type-con $record 0
+fun-decl $is_record_type 2 0
+    type-con $ctype 0
+    bool
+fun-decl $is_record_field 4 0
+    type-con $ctype 0
+    type-con $field 0
+    type-con $ctype 0
+    bool
+fun-decl $as_record_record_field 2 0
+    type-con $field 0
+    type-con $field 0
+fun-decl $rec_eq 3 0
+    type-con $record 0
+    type-con $record 0
+    bool
+fun-decl $rec_base_eq 3 0
+    int
+    int
+    bool
+fun-decl $int_to_record 2 0
+    int
+    type-con $record 0
+fun-decl $record_to_int 2 0
+    type-con $record 0
+    int
+fun-decl $good_state 2 0
+    type-con $state 0
+    bool
+fun-decl $invok_state 2 0
+    type-con $state 0
+    bool
+fun-decl $has_volatile_owns_set 2 0
+    type-con $ctype 0
+    bool
+fun-decl $owns_set_field 2 0
+    type-con $ctype 0
+    type-con $field 0
+fun-decl $st_owner 2 0
+    type-con $status 0
+    type-con $ptr 0
+fun-decl $st_closed 2 0
+    type-con $status 0
+    bool
+fun-decl $st_timestamp 2 0
+    type-con $status 0
+    int
+fun-decl $st_ref_cnt 2 0
+    type-con $status 0
+    int
+fun-decl $owner 3 0
+    type-con $state 0
+    type-con $ptr 0
+    type-con $ptr 0
+fun-decl $closed 3 0
+    type-con $state 0
+    type-con $ptr 0
+    bool
+fun-decl $timestamp 3 0
+    type-con $state 0
+    type-con $ptr 0
+    int
+fun-decl $position_marker 1 0
+    bool
+fun-decl $st 3 1
+    type-con $state 0
+    type-con $ptr 0
+    type-con $status 0
+  attribute inline 1
+    expr-attr
+      true
+fun-decl $ts 3 1
+    type-con $state 0
+    type-con $ptr 0
+    type-con $type_state 0
+  attribute inline 1
+    expr-attr
+      true
+fun-decl $owns 3 1
+    type-con $state 0
+    type-con $ptr 0
+    type-con $ptrset 0
+  attribute weight 1
+    expr-attr
+      int-num 0
+fun-decl $nested 3 1
+    type-con $state 0
+    type-con $ptr 0
+    bool
+  attribute inline 1
+    expr-attr
+      true
+fun-decl $nested_in 4 1
+    type-con $state 0
+    type-con $ptr 0
+    type-con $ptr 0
+    bool
+  attribute inline 1
+    expr-attr
+      true
+fun-decl $wrapped 4 1
+    type-con $state 0
+    type-con $ptr 0
+    type-con $ctype 0
+    bool
+  attribute inline 1
+    expr-attr
+      true
+fun-decl $irrelevant 3 1
+    type-con $state 0
+    type-con $ptr 0
+    bool
+  attribute inline 1
+    expr-attr
+      true
+fun-decl $mutable 3 1
+    type-con $state 0
+    type-con $ptr 0
+    bool
+  attribute weight 1
+    expr-attr
+      int-num 0
+fun-decl $thread_owned 3 1
+    type-con $state 0
+    type-con $ptr 0
+    bool
+  attribute inline 1
+    expr-attr
+      true
+fun-decl $thread_owned_or_even_mutable 3 1
+    type-con $state 0
+    type-con $ptr 0
+    bool
+  attribute inline 1
+    expr-attr
+      true
+fun-decl $typed 3 0
+    type-con $state 0
+    type-con $ptr 0
+    bool
+fun-decl $typed2 4 1
+    type-con $state 0
+    type-con $ptr 0
+    type-con $ctype 0
+    bool
+  attribute inline 1
+    expr-attr
+      true
+fun-decl $ptr_eq 3 1
+    type-con $ptr 0
+    type-con $ptr 0
+    bool
+  attribute inline 1
+    expr-attr
+      true
+fun-decl $ptr_neq 3 1
+    type-con $ptr 0
+    type-con $ptr 0
+    bool
+  attribute inline 1
+    expr-attr
+      true
+fun-decl $is_primitive_field_of 4 1
+    type-con $state 0
+    type-con $ptr 0
+    type-con $ptr 0
+    bool
+  attribute inline 1
+    expr-attr
+      true
+fun-decl $instantiate_st 2 0
+    type-con $status 0
+    bool
+fun-decl $is_domain_root 3 0
+    type-con $state 0
+    type-con $ptr 0
+    bool
+fun-decl $in_wrapped_domain 3 0
+    type-con $state 0
+    type-con $ptr 0
+    bool
+fun-decl $thread_local_np 3 1
+    type-con $state 0
+    type-con $ptr 0
+    bool
+  attribute inline 1
+    expr-attr
+      true
+fun-decl $thread_local 3 0
+    type-con $state 0
+    type-con $ptr 0
+    bool
+fun-decl $thread_local2 4 1
+    type-con $state 0
+    type-con $ptr 0
+    type-con $ctype 0
+    bool
+  attribute inline 1
+    expr-attr
+      true
+fun-decl $dont_instantiate 2 0
+    type-con $ptr 0
+    bool
+fun-decl $dont_instantiate_int 2 0
+    int
+    bool
+fun-decl $dont_instantiate_state 2 0
+    type-con $state 0
+    bool
+fun-decl $instantiate_int 2 0
+    int
+    bool
+fun-decl $instantiate_bool 2 0
+    bool
+    bool
+fun-decl $instantiate_ptr 2 0
+    type-con $ptr 0
+    bool
+fun-decl $instantiate_ptrset 2 0
+    type-con $ptrset 0
+    bool
+fun-decl $inv 4 1
+    type-con $state 0
+    type-con $ptr 0
+    type-con $ctype 0
+    bool
+  attribute inline 1
+    expr-attr
+      true
+fun-decl $inv2nt 4 1
+    type-con $state 0
+    type-con $state 0
+    type-con $ptr 0
+    bool
+  attribute inline 1
+    expr-attr
+      true
+fun-decl $imply_inv 4 0
+    type-con $state 0
+    type-con $ptr 0
+    type-con $ctype 0
+    bool
+fun-decl $inv2 5 0
+    type-con $state 0
+    type-con $state 0
+    type-con $ptr 0
+    type-con $ctype 0
+    bool
+fun-decl $inv2_when_closed 5 1
+    type-con $state 0
+    type-con $state 0
+    type-con $ptr 0
+    type-con $ctype 0
+    bool
+  attribute inline 1
+    expr-attr
+      true
+fun-decl $sequential 5 1
+    type-con $state 0
+    type-con $state 0
+    type-con $ptr 0
+    type-con $ctype 0
+    bool
+  attribute weight 1
+    expr-attr
+      int-num 0
+fun-decl $depends 5 1
+    type-con $state 0
+    type-con $state 0
+    type-con $ptr 0
+    type-con $ptr 0
+    bool
+  attribute weight 1
+    expr-attr
+      int-num 0
+fun-decl $spans_the_same 5 1
+    type-con $state 0
+    type-con $state 0
+    type-con $ptr 0
+    type-con $ctype 0
+    bool
+  attribute weight 1
+    expr-attr
+      int-num 0
+fun-decl $state_spans_the_same 5 0
+    type-con $state 0
+    type-con $state 0
+    type-con $ptr 0
+    type-con $ctype 0
+    bool
+fun-decl $nonvolatile_spans_the_same 5 1
+    type-con $state 0
+    type-con $state 0
+    type-con $ptr 0
+    type-con $ctype 0
+    bool
+  attribute weight 1
+    expr-attr
+      int-num 0
+fun-decl $state_nonvolatile_spans_the_same 5 0
+    type-con $state 0
+    type-con $state 0
+    type-con $ptr 0
+    type-con $ctype 0
+    bool
+fun-decl $in_extent_of 4 1
+    type-con $state 0
+    type-con $ptr 0
+    type-con $ptr 0
+    bool
+  attribute inline 1
+    expr-attr
+      true
+fun-decl $in_full_extent_of 3 1
+    type-con $ptr 0
+    type-con $ptr 0
+    bool
+  attribute inline 1
+    expr-attr
+      true
+fun-decl $extent_mutable 3 0
+    type-con $state 0
+    type-con $ptr 0
+    bool
+fun-decl $extent_is_fresh 4 0
+    type-con $state 0
+    type-con $state 0
+    type-con $ptr 0
+    bool
+fun-decl $forall_inv2_when_closed 3 1
+    type-con $state 0
+    type-con $state 0
+    bool
+  attribute inline 1
+    expr-attr
+      true
+fun-decl $function_entry 2 0
+    type-con $state 0
+    bool
+fun-decl $full_stop 2 0
+    type-con $state 0
+    bool
+fun-decl $full_stop_ext 3 1
+    type-con $token 0
+    type-con $state 0
+    bool
+  attribute inline 1
+    expr-attr
+      true
+fun-decl $file_name_is 3 0
+    int
+    type-con $token 0
+    bool
+fun-decl $closed_is_transitive 2 1
+    type-con $state 0
+    bool
+  attribute inline 1
+    expr-attr
+      true
+fun-decl $call_transition 3 0
+    type-con $state 0
+    type-con $state 0
+    bool
+fun-decl $good_state_ext 3 0
+    type-con $token 0
+    type-con $state 0
+    bool
+fun-decl $local_value_is 6 0
+    type-con $state 0
+    type-con $token 0
+    type-con $token 0
+    int
+    type-con $ctype 0
+    bool
+fun-decl $local_value_is_ptr 6 0
+    type-con $state 0
+    type-con $token 0
+    type-con $token 0
+    type-con $ptr 0
+    type-con $ctype 0
+    bool
+fun-decl $read_ptr_m 4 0
+    type-con $state 0
+    type-con $ptr 0
+    type-con $ctype 0
+    type-con $ptr 0
+fun-decl $type_code_is 3 0
+    int
+    type-con $ctype 0
+    bool
+fun-decl $function_arg_type 4 0
+    type-con $pure_function 0
+    int
+    type-con $ctype 0
+    bool
+fun-decl $ver_domain 2 0
+    type-con $version 0
+    type-con $ptrset 0
+fun-decl $read_version 3 1
+    type-con $state 0
+    type-con $ptr 0
+    type-con $version 0
+  attribute weight 1
+    expr-attr
+      int-num 0
+fun-decl $domain 3 1
+    type-con $state 0
+    type-con $ptr 0
+    type-con $ptrset 0
+  attribute weight 1
+    expr-attr
+      int-num 0
+fun-decl $in_domain 4 0
+    type-con $state 0
+    type-con $ptr 0
+    type-con $ptr 0
+    bool
+fun-decl $in_vdomain 4 0
+    type-con $state 0
+    type-con $ptr 0
+    type-con $ptr 0
+    bool
+fun-decl $in_domain_lab 5 0
+    type-con $state 0
+    type-con $ptr 0
+    type-con $ptr 0
+    type-con $label 0
+    bool
+fun-decl $in_vdomain_lab 5 0
+    type-con $state 0
+    type-con $ptr 0
+    type-con $ptr 0
+    type-con $label 0
+    bool
+fun-decl $inv_lab 4 0
+    type-con $state 0
+    type-con $ptr 0
+    type-con $label 0
+    bool
+fun-decl $dom_thread_local 3 1
+    type-con $state 0
+    type-con $ptr 0
+    bool
+  attribute inline 1
+    expr-attr
+      true
+fun-decl $fetch_from_domain 3 0
+    type-con $version 0
+    type-con $ptr 0
+    int
+fun-decl $in_claim_domain 3 0
+    type-con $ptr 0
+    type-con $ptr 0
+    bool
+fun-decl $by_claim 5 1
+    type-con $state 0
+    type-con $ptr 0
+    type-con $ptr 0
+    type-con $ptr 0
+    type-con $ptr 0
+  attribute weight 1
+    expr-attr
+      int-num 0
+fun-decl $claim_version 2 0
+    type-con $ptr 0
+    type-con $version 0
+fun-decl $read_vol_version 3 1
+    type-con $state 0
+    type-con $ptr 0
+    type-con $vol_version 0
+  attribute weight 1
+    expr-attr
+      int-num 0
+fun-decl $fetch_from_vv 3 0
+    type-con $vol_version 0
+    type-con $ptr 0
+    int
+fun-decl $fetch_vol_field 4 1
+    type-con $state 0
+    type-con $ptr 0
+    type-con $field 0
+    int
+  attribute inline 1
+    expr-attr
+      true
+fun-decl $is_approved_by 4 0
+    type-con $ctype 0
+    type-con $field 0
+    type-con $field 0
+    bool
+fun-decl $inv_is_approved_by_ptr 6 1
+    type-con $state 0
+    type-con $state 0
+    type-con $ptr 0
+    type-con $ptr 0
+    type-con $field 0
+    bool
+  attribute inline 1
+    expr-attr
+      true
+fun-decl $inv_is_approved_by 6 1
+    type-con $state 0
+    type-con $state 0
+    type-con $ptr 0
+    type-con $field 0
+    type-con $field 0
+    bool
+  attribute inline 1
+    expr-attr
+      true
+fun-decl $is_owner_approved 3 0
+    type-con $ctype 0
+    type-con $field 0
+    bool
+fun-decl $inv_is_owner_approved 5 1
+    type-con $state 0
+    type-con $state 0
+    type-con $ptr 0
+    type-con $field 0
+    bool
+  attribute inline 1
+    expr-attr
+      true
+fun-decl $good_for_admissibility 2 0
+    type-con $state 0
+    bool
+fun-decl $good_for_post_admissibility 2 0
+    type-con $state 0
+    bool
+fun-decl $stuttering_pre 3 1
+    type-con $state 0
+    type-con $ptr 0
+    bool
+  attribute inline 1
+    expr-attr
+      true
+fun-decl $admissibility_pre 3 1
+    type-con $state 0
+    type-con $ptr 0
+    bool
+  attribute inline 1
+    expr-attr
+      true
+fun-decl $mutable_increases 3 1
+    type-con $state 0
+    type-con $state 0
+    bool
+  attribute inline 1
+    expr-attr
+      true
+fun-decl $meta_eq 3 1
+    type-con $state 0
+    type-con $state 0
+    bool
+  attribute inline 1
+    expr-attr
+      true
+fun-decl $is_stuttering_check 1 0
+    bool
+fun-decl $is_unwrap_check 1 0
+    bool
+fun-decl $is_admissibility_check 1 1
+    bool
+  attribute inline 1
+    expr-attr
+      true
+fun-decl $good_for_pre_can_unwrap 2 0
+    type-con $state 0
+    bool
+fun-decl $good_for_post_can_unwrap 2 0
+    type-con $state 0
+    bool
+fun-decl $unwrap_check_pre 3 1
+    type-con $state 0
+    type-con $ptr 0
+    bool
+  attribute inline 1
+    expr-attr
+      true
+fun-decl $update_int 4 0
+    type-con $state 0
+    type-con $ptr 0
+    int
+    type-con $state 0
+fun-decl $timestamp_is_now 3 1
+    type-con $state 0
+    type-con $ptr 0
+    bool
+  attribute inline 1
+    expr-attr
+      true
+fun-decl $now_writable 3 1
+    type-con $state 0
+    type-con $ptr 0
+    bool
+  attribute inline 1
+    expr-attr
+      true
+fun-decl $timestamp_post 3 1
+    type-con $state 0
+    type-con $state 0
+    bool
+  attribute inline 1
+    expr-attr
+      true
+fun-decl $timestamp_post_strict 3 1
+    type-con $state 0
+    type-con $state 0
+    bool
+  attribute inline 1
+    expr-attr
+      true
+fun-decl $pre_wrap 2 0
+    type-con $state 0
+    bool
+fun-decl $pre_unwrap 2 0
+    type-con $state 0
+    bool
+fun-decl $pre_static_wrap 2 0
+    type-con $state 0
+    bool
+fun-decl $pre_static_unwrap 2 0
+    type-con $state 0
+    bool
+fun-decl $unwrap_post 5 1
+    type-con $state 0
+    type-con $state 0
+    type-con $ptr 0
+    type-con $ptr 0
+    bool
+  attribute inline 1
+    expr-attr
+      true
+fun-decl $unwrap_post_claimable 5 1
+    type-con $state 0
+    type-con $state 0
+    type-con $ptr 0
+    type-con $ptr 0
+    bool
+  attribute inline 1
+    expr-attr
+      true
+fun-decl $keeps 4 2
+    type-con $state 0
+    type-con $ptr 0
+    type-con $ptr 0
+    bool
+  attribute inline 1
+    expr-attr
+      true
+  attribute expand 1
+    expr-attr
+      true
+fun-decl $expect_unreachable 1 0
+    bool
+fun-decl $taken_over 4 0
+    type-con $state 0
+    type-con $ptr 0
+    type-con $ptr 0
+    type-con $status 0
+fun-decl $take_over 4 0
+    type-con $state 0
+    type-con $ptr 0
+    type-con $ptr 0
+    type-con $state 0
+fun-decl $released 4 0
+    type-con $state 0
+    type-con $ptr 0
+    type-con $ptr 0
+    type-con $status 0
+fun-decl $release 5 0
+    type-con $state 0
+    type-con $state 0
+    type-con $ptr 0
+    type-con $ptr 0
+    type-con $state 0
+fun-decl $post_unwrap 3 0
+    type-con $state 0
+    type-con $state 0
+    bool
+fun-decl $new_ownees 4 1
+    type-con $state 0
+    type-con $ptr 0
+    type-con $ptrset 0
+    type-con $ptrset 0
+  attribute inline 1
+    expr-attr
+      true
+fun-decl $get_memory_allocator 1 0
+    type-con $ptr 0
+fun-decl $memory_allocator_type 1 1
+    type-con $ctype 0
+  attribute unique 0
+fun-decl $memory_allocator_ref 1 0
+    int
+fun-decl $program_entry_point 2 0
+    type-con $state 0
+    bool
+fun-decl $program_entry_point_ch 2 0
+    type-con $state 0
+    bool
+fun-decl $is_global 3 1
+    type-con $ptr 0
+    type-con $ctype 0
+    bool
+  attribute inline 1
+    expr-attr
+      true
+fun-decl $is_global_array 4 1
+    type-con $ptr 0
+    type-con $ctype 0
+    int
+    bool
+  attribute inline 1
+    expr-attr
+      true
+fun-decl $active_option 3 1
+    type-con $state 0
+    type-con $ptr 0
+    type-con $field 0
+  attribute inline 1
+    expr-attr
+      true
+fun-decl $ts_active_option 2 0
+    type-con $type_state 0
+    type-con $field 0
+fun-decl $union_active 4 1
+    type-con $state 0
+    type-con $ptr 0
+    type-con $field 0
+    bool
+  attribute inline 1
+    expr-attr
+      true
+fun-decl $is_union_field 3 0
+    type-con $ctype 0
+    type-con $field 0
+    bool
+fun-decl $union_havoced 4 1
+    type-con $state 0
+    type-con $state 0
+    type-con $ptr 0
+    bool
+  attribute inline 1
+    expr-attr
+      true
+fun-decl $full_extent 2 0
+    type-con $ptr 0
+    type-con $ptrset 0
+fun-decl $extent 3 0
+    type-con $state 0
+    type-con $ptr 0
+    type-con $ptrset 0
+fun-decl $span 2 0
+    type-con $ptr 0
+    type-con $ptrset 0
+fun-decl $in_span_of 3 1
+    type-con $ptr 0
+    type-con $ptr 0
+    bool
+  attribute inline 1
+    expr-attr
+      true
+fun-decl $first_option_typed 3 0
+    type-con $state 0
+    type-con $ptr 0
+    bool
+fun-decl $struct_extent 2 1
+    type-con $ptr 0
+    type-con $ptrset 0
+  attribute inline 1
+    expr-attr
+      true
+fun-decl $in_struct_extent_of 3 1
+    type-con $ptr 0
+    type-con $ptr 0
+    bool
+  attribute inline 1
+    expr-attr
+      true
+fun-decl $volatile_span 3 0
+    type-con $state 0
+    type-con $ptr 0
+    type-con $ptrset 0
+fun-decl $left_split 3 0
+    type-con $ptr 0
+    int
+    type-con $ptr 0
+fun-decl $right_split 3 0
+    type-con $ptr 0
+    int
+    type-con $ptr 0
+fun-decl $joined_array 3 0
+    type-con $ptr 0
+    type-con $ptr 0
+    type-con $ptr 0
+fun-decl $mutable_root 3 1
+    type-con $state 0
+    type-con $ptr 0
+    bool
+  attribute inline 1
+    expr-attr
+      true
+fun-decl $set_in 3 0
+    type-con $ptr 0
+    type-con $ptrset 0
+    bool
+fun-decl $set_empty 1 0
+    type-con $ptrset 0
+fun-decl $set_singleton 2 0
+    type-con $ptr 0
+    type-con $ptrset 0
+fun-decl $non_null_set_singleton 2 0
+    type-con $ptr 0
+    type-con $ptrset 0
+fun-decl $set_union 3 0
+    type-con $ptrset 0
+    type-con $ptrset 0
+    type-con $ptrset 0
+fun-decl $set_difference 3 0
+    type-con $ptrset 0
+    type-con $ptrset 0
+    type-con $ptrset 0
+fun-decl $set_intersection 3 0
+    type-con $ptrset 0
+    type-con $ptrset 0
+    type-con $ptrset 0
+fun-decl $set_subset 3 0
+    type-con $ptrset 0
+    type-con $ptrset 0
+    bool
+fun-decl $set_eq 3 0
+    type-con $ptrset 0
+    type-con $ptrset 0
+    bool
+fun-decl $set_cardinality 2 0
+    type-con $ptrset 0
+    int
+fun-decl $set_universe 1 0
+    type-con $ptrset 0
+fun-decl $set_disjoint 3 0
+    type-con $ptrset 0
+    type-con $ptrset 0
+    bool
+fun-decl $id_set_disjoint 4 0
+    type-con $ptr 0
+    type-con $ptrset 0
+    type-con $ptrset 0
+    int
+fun-decl $set_in3 3 0
+    type-con $ptr 0
+    type-con $ptrset 0
+    bool
+fun-decl $set_in2 3 0
+    type-con $ptr 0
+    type-con $ptrset 0
+    bool
+fun-decl $in_some_owns 2 0
+    type-con $ptr 0
+    bool
+fun-decl $set_in0 3 0
+    type-con $ptr 0
+    type-con $ptrset 0
+    bool
+fun-decl sk_hack 2 0
+    bool
+    bool
+fun-decl $writes_nothing 3 1
+    type-con $state 0
+    type-con $state 0
+    bool
+  attribute inline 1
+    expr-attr
+      true
+fun-decl $array 3 0
+    type-con $ctype 0
+    int
+    type-con $ctype 0
+fun-decl $element_type 2 0
+    type-con $ctype 0
+    type-con $ctype 0
+fun-decl $array_length 2 0
+    type-con $ctype 0
+    int
+fun-decl $is_array_elt 3 1
+    type-con $state 0
+    type-con $ptr 0
+    bool
+  attribute inline 1
+    expr-attr
+      true
+fun-decl $inlined_array 3 1
+    type-con $ptr 0
+    type-con $ctype 0
+    type-con $ptr 0
+  attribute weight 1
+    expr-attr
+      int-num 0
+fun-decl $idx 4 0
+    type-con $ptr 0
+    int
+    type-con $ctype 0
+    type-con $ptr 0
+fun-decl $add.mul 4 2
+    int
+    int
+    int
+    int
+  attribute inline 1
+    expr-attr
+      true
+  attribute expand 1
+    expr-attr
+      true
+fun-decl $add 3 2
+    int
+    int
+    int
+  attribute inline 1
+    expr-attr
+      true
+  attribute expand 1
+    expr-attr
+      true
+fun-decl $is_array_vol_or_nonvol 6 1
+    type-con $state 0
+    type-con $ptr 0
+    type-con $ctype 0
+    int
+    bool
+    bool
+  attribute weight 1
+    expr-attr
+      int-num 0
+fun-decl $is_array 5 1
+    type-con $state 0
+    type-con $ptr 0
+    type-con $ctype 0
+    int
+    bool
+  attribute weight 1
+    expr-attr
+      int-num 0
+fun-decl $is_thread_local_array 5 1
+    type-con $state 0
+    type-con $ptr 0
+    type-con $ctype 0
+    int
+    bool
+  attribute inline 1
+    expr-attr
+      true
+fun-decl $is_mutable_array 5 1
+    type-con $state 0
+    type-con $ptr 0
+    type-con $ctype 0
+    int
+    bool
+  attribute inline 1
+    expr-attr
+      true
+fun-decl $is_array_emb 6 1
+    type-con $state 0
+    type-con $ptr 0
+    type-con $ctype 0
+    int
+    type-con $ptr 0
+    bool
+  attribute inline 1
+    expr-attr
+      true
+fun-decl $is_array_emb_path 8 1
+    type-con $state 0
+    type-con $ptr 0
+    type-con $ctype 0
+    int
+    type-con $ptr 0
+    type-con $field 0
+    bool
+    bool
+  attribute inline 1
+    expr-attr
+      true
+fun-decl $array_field_properties 6 1
+    type-con $field 0
+    type-con $ctype 0
+    int
+    bool
+    bool
+    bool
+  attribute inline 1
+    expr-attr
+      true
+fun-decl $no_inline_array_field_properties 6 1
+    type-con $field 0
+    type-con $ctype 0
+    int
+    bool
+    bool
+    bool
+  attribute inline 1
+    expr-attr
+      true
+fun-decl $array_elt_emb 4 1
+    type-con $state 0
+    type-con $ptr 0
+    type-con $ptr 0
+    bool
+  attribute inline 1
+    expr-attr
+      true
+fun-decl $array_members 4 0
+    type-con $ptr 0
+    type-con $ctype 0
+    int
+    type-con $ptrset 0
+fun-decl $array_range 4 0
+    type-con $ptr 0
+    type-con $ctype 0
+    int
+    type-con $ptrset 0
+fun-decl $non_null_array_range 4 0
+    type-con $ptr 0
+    type-con $ctype 0
+    int
+    type-con $ptrset 0
+fun-decl $non_null_extent 3 0
+    type-con $state 0
+    type-con $ptr 0
+    type-con $ptrset 0
+fun-decl $as_array 4 1
+    type-con $ptr 0
+    type-con $ctype 0
+    int
+    type-con $ptr 0
+  attribute inline 1
+    expr-attr
+      true
+fun-decl $array_eq 6 1
+    type-con $state 0
+    type-con $state 0
+    type-con $ptr 0
+    type-con $ctype 0
+    int
+    bool
+  attribute inline 1
+    expr-attr
+      true
+fun-decl $index_within 3 0
+    type-con $ptr 0
+    type-con $ptr 0
+    int
+fun-decl $in_array 5 1
+    type-con $ptr 0
+    type-con $ptr 0
+    type-con $ctype 0
+    int
+    bool
+  attribute inline 1
+    expr-attr
+      true
+fun-decl $in_array_full_extent_of 5 1
+    type-con $ptr 0
+    type-con $ptr 0
+    type-con $ctype 0
+    int
+    bool
+  attribute inline 1
+    expr-attr
+      true
+fun-decl $in_array_extent_of 6 1
+    type-con $state 0
+    type-con $ptr 0
+    type-con $ptr 0
+    type-con $ctype 0
+    int
+    bool
+  attribute inline 1
+    expr-attr
+      true
+fun-decl $in_range 4 1
+    int
+    int
+    int
+    bool
+  attribute inline 1
+    expr-attr
+      true
+fun-decl $bool_to_int 2 1
+    bool
+    int
+  attribute inline 1
+    expr-attr
+      true
+fun-decl $int_to_bool 2 1
+    int
+    bool
+  attribute inline 1
+    expr-attr
+      true
+fun-decl $read_bool 3 1
+    type-con $state 0
+    type-con $ptr 0
+    bool
+  attribute inline 1
+    expr-attr
+      true
+fun-decl $ite.int 4 3
+    bool
+    int
+    int
+    int
+  attribute external 1
+    string-attr ITE
+  attribute bvz 1
+    string-attr ITE
+  attribute bvint 1
+    string-attr ITE
+fun-decl $ite.bool 4 3
+    bool
+    bool
+    bool
+    bool
+  attribute external 1
+    string-attr ITE
+  attribute bvz 1
+    string-attr ITE
+  attribute bvint 1
+    string-attr ITE
+fun-decl $ite.ptr 4 3
+    bool
+    type-con $ptr 0
+    type-con $ptr 0
+    type-con $ptr 0
+  attribute external 1
+    string-attr ITE
+  attribute bvz 1
+    string-attr ITE
+  attribute bvint 1
+    string-attr ITE
+fun-decl $ite.struct 4 3
+    bool
+    type-con $struct 0
+    type-con $struct 0
+    type-con $struct 0
+  attribute external 1
+    string-attr ITE
+  attribute bvz 1
+    string-attr ITE
+  attribute bvint 1
+    string-attr ITE
+fun-decl $ite.ptrset 4 3
+    bool
+    type-con $ptrset 0
+    type-con $ptrset 0
+    type-con $ptrset 0
+  attribute external 1
+    string-attr ITE
+  attribute bvz 1
+    string-attr ITE
+  attribute bvint 1
+    string-attr ITE
+fun-decl $ite.primitive 4 3
+    bool
+    type-con $primitive 0
+    type-con $primitive 0
+    type-con $primitive 0
+  attribute external 1
+    string-attr ITE
+  attribute bvz 1
+    string-attr ITE
+  attribute bvint 1
+    string-attr ITE
+fun-decl $ite.record 4 3
+    bool
+    type-con $record 0
+    type-con $record 0
+    type-con $record 0
+  attribute external 1
+    string-attr ITE
+  attribute bvz 1
+    string-attr ITE
+  attribute bvint 1
+    string-attr ITE
+fun-decl $bool_id 2 1
+    bool
+    bool
+  attribute weight 1
+    expr-attr
+      int-num 0
+fun-decl $min.i1 1 0
+    int
+fun-decl $max.i1 1 0
+    int
+fun-decl $min.i2 1 0
+    int
+fun-decl $max.i2 1 0
+    int
+fun-decl $min.i4 1 0
+    int
+fun-decl $max.i4 1 0
+    int
+fun-decl $min.i8 1 0
+    int
+fun-decl $max.i8 1 0
+    int
+fun-decl $max.u1 1 0
+    int
+fun-decl $max.u2 1 0
+    int
+fun-decl $max.u4 1 0
+    int
+fun-decl $max.u8 1 0
+    int
+fun-decl $in_range_i1 2 1
+    int
+    bool
+  attribute inline 1
+    expr-attr
+      true
+fun-decl $in_range_i2 2 1
+    int
+    bool
+  attribute inline 1
+    expr-attr
+      true
+fun-decl $in_range_i4 2 1
+    int
+    bool
+  attribute inline 1
+    expr-attr
+      true
+fun-decl $in_range_i8 2 1
+    int
+    bool
+  attribute inline 1
+    expr-attr
+      true
+fun-decl $in_range_u1 2 1
+    int
+    bool
+  attribute inline 1
+    expr-attr
+      true
+fun-decl $in_range_u2 2 1
+    int
+    bool
+  attribute inline 1
+    expr-attr
+      true
+fun-decl $in_range_u4 2 1
+    int
+    bool
+  attribute inline 1
+    expr-attr
+      true
+fun-decl $in_range_u8 2 1
+    int
+    bool
+  attribute inline 1
+    expr-attr
+      true
+fun-decl $in_range_div_i1 3 1
+    int
+    int
+    bool
+  attribute inline 1
+    expr-attr
+      true
+fun-decl $in_range_div_i2 3 1
+    int
+    int
+    bool
+  attribute inline 1
+    expr-attr
+      true
+fun-decl $in_range_div_i4 3 1
+    int
+    int
+    bool
+  attribute inline 1
+    expr-attr
+      true
+fun-decl $in_range_div_i8 3 1
+    int
+    int
+    bool
+  attribute inline 1
+    expr-attr
+      true
+fun-decl $read_i1 3 1
+    type-con $state 0
+    type-con $ptr 0
+    int
+  attribute weight 1
+    expr-attr
+      int-num 0
+fun-decl $read_i2 3 1
+    type-con $state 0
+    type-con $ptr 0
+    int
+  attribute weight 1
+    expr-attr
+      int-num 0
+fun-decl $read_i4 3 1
+    type-con $state 0
+    type-con $ptr 0
+    int
+  attribute weight 1
+    expr-attr
+      int-num 0
+fun-decl $read_i8 3 1
+    type-con $state 0
+    type-con $ptr 0
+    int
+  attribute weight 1
+    expr-attr
+      int-num 0
+fun-decl $read_u1 3 1
+    type-con $state 0
+    type-con $ptr 0
+    int
+  attribute weight 1
+    expr-attr
+      int-num 0
+fun-decl $read_u2 3 1
+    type-con $state 0
+    type-con $ptr 0
+    int
+  attribute weight 1
+    expr-attr
+      int-num 0
+fun-decl $read_u4 3 1
+    type-con $state 0
+    type-con $ptr 0
+    int
+  attribute weight 1
+    expr-attr
+      int-num 0
+fun-decl $read_u8 3 1
+    type-con $state 0
+    type-con $ptr 0
+    int
+  attribute weight 1
+    expr-attr
+      int-num 0
+fun-decl $ptr_to_u8 2 0
+    type-con $ptr 0
+    int
+fun-decl $ptr_to_i8 2 0
+    type-con $ptr 0
+    int
+fun-decl $ptr_to_u4 2 0
+    type-con $ptr 0
+    int
+fun-decl $ptr_to_i4 2 0
+    type-con $ptr 0
+    int
+fun-decl $u8_to_ptr 2 1
+    int
+    type-con $ptr 0
+  attribute inline 1
+    expr-attr
+      true
+fun-decl $i8_to_ptr 2 1
+    int
+    type-con $ptr 0
+  attribute inline 1
+    expr-attr
+      true
+fun-decl $u4_to_ptr 2 1
+    int
+    type-con $ptr 0
+  attribute inline 1
+    expr-attr
+      true
+fun-decl $i4_to_ptr 2 1
+    int
+    type-con $ptr 0
+  attribute inline 1
+    expr-attr
+      true
+fun-decl $byte_ptr_subtraction 3 1
+    type-con $ptr 0
+    type-con $ptr 0
+    int
+  attribute weight 1
+    expr-attr
+      int-num 0
+fun-decl $_pow2 2 0
+    int
+    int
+fun-decl $_or 4 0
+    type-con $ctype 0
+    int
+    int
+    int
+fun-decl $_xor 4 0
+    type-con $ctype 0
+    int
+    int
+    int
+fun-decl $_and 4 0
+    type-con $ctype 0
+    int
+    int
+    int
+fun-decl $_not 3 0
+    type-con $ctype 0
+    int
+    int
+fun-decl $unchk_add 4 1
+    type-con $ctype 0
+    int
+    int
+    int
+  attribute weight 1
+    expr-attr
+      int-num 0
+fun-decl $unchk_sub 4 1
+    type-con $ctype 0
+    int
+    int
+    int
+  attribute weight 1
+    expr-attr
+      int-num 0
+fun-decl $unchk_mul 4 1
+    type-con $ctype 0
+    int
+    int
+    int
+  attribute weight 1
+    expr-attr
+      int-num 0
+fun-decl $unchk_div 4 1
+    type-con $ctype 0
+    int
+    int
+    int
+  attribute weight 1
+    expr-attr
+      int-num 0
+fun-decl $unchk_mod 4 1
+    type-con $ctype 0
+    int
+    int
+    int
+  attribute weight 1
+    expr-attr
+      int-num 0
+fun-decl $_shl 4 1
+    type-con $ctype 0
+    int
+    int
+    int
+  attribute weight 1
+    expr-attr
+      int-num 0
+fun-decl $_shr 3 1
+    int
+    int
+    int
+  attribute weight 1
+    expr-attr
+      int-num 0
+fun-decl $bv_extract_signed 5 0
+    int
+    int
+    int
+    int
+    int
+fun-decl $bv_extract_unsigned 5 0
+    int
+    int
+    int
+    int
+    int
+fun-decl $bv_update 6 0
+    int
+    int
+    int
+    int
+    int
+    int
+fun-decl $unchecked 3 0
+    type-con $ctype 0
+    int
+    int
+fun-decl $in_range_t 3 0
+    type-con $ctype 0
+    int
+    bool
+fun-decl $_mul 3 1
+    int
+    int
+    int
+  attribute weight 1
+    expr-attr
+      int-num 0
+fun-decl $get_string_literal 3 0
+    int
+    int
+    type-con $ptr 0
+fun-decl $get_fnptr 3 0
+    int
+    type-con $ctype 0
+    type-con $ptr 0
+fun-decl $get_fnptr_ref 2 0
+    int
+    int
+fun-decl $get_fnptr_inv 2 0
+    int
+    int
+fun-decl $is_fnptr_type 2 0
+    type-con $ctype 0
+    bool
+fun-decl $is_math_type 2 0
+    type-con $ctype 0
+    bool
+fun-decl $claims_obj 3 0
+    type-con $ptr 0
+    type-con $ptr 0
+    bool
+fun-decl $valid_claim 3 0
+    type-con $state 0
+    type-con $ptr 0
+    bool
+fun-decl $claim_initial_assumptions 4 1
+    type-con $state 0
+    type-con $ptr 0
+    type-con $token 0
+    bool
+  attribute inline 1
+    expr-attr
+      true
+fun-decl $claim_transitivity_assumptions 5 1
+    type-con $state 0
+    type-con $state 0
+    type-con $ptr 0
+    type-con $token 0
+    bool
+  attribute inline 1
+    expr-attr
+      true
+fun-decl $valid_claim_impl 3 1
+    type-con $state 0
+    type-con $state 0
+    bool
+  attribute inline 1
+    expr-attr
+      true
+fun-decl $claims_claim 3 0
+    type-con $ptr 0
+    type-con $ptr 0
+    bool
+fun-decl $not_shared 3 1
+    type-con $state 0
+    type-con $ptr 0
+    bool
+  attribute weight 1
+    expr-attr
+      int-num 0
+fun-decl $claimed_closed 3 1
+    type-con $state 0
+    type-con $ptr 0
+    bool
+  attribute weight 1
+    expr-attr
+      int-num 0
+fun-decl $no_claim 1 1
+    type-con $ptr 0
+  attribute unique 0
+fun-decl $ref_cnt 3 1
+    type-con $state 0
+    type-con $ptr 0
+    int
+  attribute weight 1
+    expr-attr
+      int-num 0
+fun-decl $is_claimable 2 0
+    type-con $ctype 0
+    bool
+fun-decl $is_thread_local_storage 2 0
+    type-con $ctype 0
+    bool
+fun-decl $frame_level 2 0
+    type-con $pure_function 0
+    int
+fun-decl $current_frame_level 1 0
+    int
+fun-decl $can_use_all_frame_axioms 2 1
+    type-con $state 0
+    bool
+  attribute inline 1
+    expr-attr
+      true
+fun-decl $can_use_frame_axiom_of 2 1
+    type-con $pure_function 0
+    bool
+  attribute inline 1
+    expr-attr
+      true
+fun-decl $reads_check_pre 2 0
+    type-con $state 0
+    bool
+fun-decl $reads_check_post 2 0
+    type-con $state 0
+    bool
+fun-decl $start_here 1 0
+    bool
+fun-decl $ptrset_to_int 2 0
+    type-con $ptrset 0
+    int
+fun-decl $int_to_ptrset 2 0
+    int
+    type-con $ptrset 0
+fun-decl $version_to_int 2 0
+    type-con $version 0
+    int
+fun-decl $int_to_version 2 0
+    int
+    type-con $version 0
+fun-decl $vol_version_to_int 2 0
+    type-con $vol_version 0
+    int
+fun-decl $int_to_vol_version 2 0
+    int
+    type-con $vol_version 0
+fun-decl $ptr_to_int 2 0
+    type-con $ptr 0
+    int
+fun-decl $int_to_ptr 2 0
+    int
+    type-con $ptr 0
+fun-decl $precise_test 2 0
+    type-con $ptr 0
+    bool
+fun-decl $updated_only_values 4 0
+    type-con $state 0
+    type-con $state 0
+    type-con $ptrset 0
+    bool
+fun-decl $updated_only_domains 4 0
+    type-con $state 0
+    type-con $state 0
+    type-con $ptrset 0
+    bool
+fun-decl $domain_updated_at 5 0
+    type-con $state 0
+    type-con $state 0
+    type-con $ptr 0
+    type-con $ptrset 0
+    bool
+fun-decl l#public 1 1
+    type-con $label 0
+  attribute unique 0
+fun-decl #tok$1^16.24 1 1
+    type-con $token 0
+  attribute unique 0
+fun-decl #tok$1^24.47 1 1
+    type-con $token 0
+  attribute unique 0
+fun-decl #tok$1^23.7 1 1
+    type-con $token 0
+  attribute unique 0
+fun-decl #tok$1^16.3 1 1
+    type-con $token 0
+  attribute unique 0
+fun-decl #loc.p 1 1
+    type-con $token 0
+  attribute unique 0
+fun-decl #tok$1^16.8 1 1
+    type-con $token 0
+  attribute unique 0
+fun-decl #loc.witness 1 1
+    type-con $token 0
+  attribute unique 0
+fun-decl #tok$1^14.3 1 1
+    type-con $token 0
+  attribute unique 0
+fun-decl #loc.max 1 1
+    type-con $token 0
+  attribute unique 0
+fun-decl #tok$1^12.3 1 1
+    type-con $token 0
+  attribute unique 0
+fun-decl #loc.len 1 1
+    type-con $token 0
+  attribute unique 0
+fun-decl #distTp1 1 1
+    type-con $ctype 0
+  attribute unique 0
+fun-decl #loc.arr 1 1
+    type-con $token 0
+  attribute unique 0
+fun-decl #tok$1^6.1 1 1
+    type-con $token 0
+  attribute unique 0
+fun-decl #file^Z?3A?5CC?5Cmax.c 1 1
+    type-con $token 0
+  attribute unique 0
+axiom 0
+    =
+    fun $sizeof 1
+    fun ^^i1 0
+    int-num 1
+axiom 0
+    =
+    fun $sizeof 1
+    fun ^^i2 0
+    int-num 2
+axiom 0
+    =
+    fun $sizeof 1
+    fun ^^i4 0
+    int-num 4
+axiom 0
+    =
+    fun $sizeof 1
+    fun ^^i8 0
+    int-num 8
+axiom 0
+    =
+    fun $sizeof 1
+    fun ^^u1 0
+    int-num 1
+axiom 0
+    =
+    fun $sizeof 1
+    fun ^^u2 0
+    int-num 2
+axiom 0
+    =
+    fun $sizeof 1
+    fun ^^u4 0
+    int-num 4
+axiom 0
+    =
+    fun $sizeof 1
+    fun ^^u8 0
+    int-num 8
+axiom 0
+    =
+    fun $sizeof 1
+    fun ^^f4 0
+    int-num 4
+axiom 0
+    =
+    fun $sizeof 1
+    fun ^^f8 0
+    int-num 8
+axiom 0
+    =
+    fun $sizeof 1
+    fun ^$#thread_id_t 0
+    int-num 1
+axiom 0
+    =
+    fun $sizeof 1
+    fun ^$#ptrset 0
+    int-num 1
+axiom 0
+    =
+    fun $ptr_level 1
+    fun ^^i1 0
+    int-num 0
+axiom 0
+    =
+    fun $ptr_level 1
+    fun ^^i2 0
+    int-num 0
+axiom 0
+    =
+    fun $ptr_level 1
+    fun ^^i4 0
+    int-num 0
+axiom 0
+    =
+    fun $ptr_level 1
+    fun ^^i8 0
+    int-num 0
+axiom 0
+    =
+    fun $ptr_level 1
+    fun ^^u1 0
+    int-num 0
+axiom 0
+    =
+    fun $ptr_level 1
+    fun ^^u2 0
+    int-num 0
+axiom 0
+    =
+    fun $ptr_level 1
+    fun ^^u4 0
+    int-num 0
+axiom 0
+    =
+    fun $ptr_level 1
+    fun ^^u8 0
+    int-num 0
+axiom 0
+    =
+    fun $ptr_level 1
+    fun ^^f4 0
+    int-num 0
+axiom 0
+    =
+    fun $ptr_level 1
+    fun ^^f8 0
+    int-num 0
+axiom 0
+    =
+    fun $ptr_level 1
+    fun ^^mathint 0
+    int-num 0
+axiom 0
+    =
+    fun $ptr_level 1
+    fun ^^bool 0
+    int-num 0
+axiom 0
+    =
+    fun $ptr_level 1
+    fun ^^void 0
+    int-num 0
+axiom 0
+    =
+    fun $ptr_level 1
+    fun ^^claim 0
+    int-num 0
+axiom 0
+    =
+    fun $ptr_level 1
+    fun ^^root_emb 0
+    int-num 0
+axiom 0
+    =
+    fun $ptr_level 1
+    fun ^$#ptrset 0
+    int-num 0
+axiom 0
+    =
+    fun $ptr_level 1
+    fun ^$#thread_id_t 0
+    int-num 0
+axiom 0
+    =
+    fun $ptr_level 1
+    fun ^$#state_t 0
+    int-num 0
+axiom 0
+    =
+    fun $ptr_level 1
+    fun ^$#struct 0
+    int-num 0
+axiom 0
+    fun $is_composite 1
+    fun ^^claim 0
+axiom 0
+    fun $is_composite 1
+    fun ^^root_emb 0
+axiom 0
+    forall 1 1 3
+      var #n
+        type-con $ctype 0
+      pat 1
+        fun $ptr_to 1
+        var #n
+          type-con $ctype 0
+      attribute qid 1
+        string-attr VccPrelu.145:15
+      attribute uniqueId 1
+        string-attr 4
+      attribute bvZ3Native 1
+        string-attr False
+    =
+    fun $unptr_to 1
+    fun $ptr_to 1
+    var #n
+      type-con $ctype 0
+    var #n
+      type-con $ctype 0
+axiom 0
+    forall 1 1 3
+      var #n
+        type-con $ctype 0
+      pat 1
+        fun $ptr_to 1
+        var #n
+          type-con $ctype 0
+      attribute qid 1
+        string-attr VccPrelu.146:15
+      attribute uniqueId 1
+        string-attr 5
+      attribute bvZ3Native 1
+        string-attr False
+    =
+    fun $sizeof 1
+    fun $ptr_to 1
+    var #n
+      type-con $ctype 0
+    int-num 8
+axiom 0
+    forall 2 1 3
+      var #r
+        type-con $ctype 0
+      var #d
+        type-con $ctype 0
+      pat 1
+        fun $map_t 2
+        var #r
+          type-con $ctype 0
+        var #d
+          type-con $ctype 0
+      attribute qid 1
+        string-attr VccPrelu.152:15
+      attribute uniqueId 1
+        string-attr 6
+      attribute bvZ3Native 1
+        string-attr False
+    =
+    fun $map_domain 1
+    fun $map_t 2
+    var #r
+      type-con $ctype 0
+    var #d
+      type-con $ctype 0
+    var #d
+      type-con $ctype 0
+axiom 0
+    forall 2 1 3
+      var #r
+        type-con $ctype 0
+      var #d
+        type-con $ctype 0
+      pat 1
+        fun $map_t 2
+        var #r
+          type-con $ctype 0
+        var #d
+          type-con $ctype 0
+      attribute qid 1
+        string-attr VccPrelu.153:15
+      attribute uniqueId 1
+        string-attr 7
+      attribute bvZ3Native 1
+        string-attr False
+    =
+    fun $map_range 1
+    fun $map_t 2
+    var #r
+      type-con $ctype 0
+    var #d
+      type-con $ctype 0
+    var #r
+      type-con $ctype 0
+axiom 0
+    forall 1 1 3
+      var #n
+        type-con $ctype 0
+      pat 1
+        fun $ptr_to 1
+        var #n
+          type-con $ctype 0
+      attribute qid 1
+        string-attr VccPrelu.158:15
+      attribute uniqueId 1
+        string-attr 8
+      attribute bvZ3Native 1
+        string-attr False
+    =
+    fun $ptr_level 1
+    fun $ptr_to 1
+    var #n
+      type-con $ctype 0
+    +
+    fun $ptr_level 1
+    var #n
+      type-con $ctype 0
+    int-num 17
+axiom 0
+    forall 2 1 3
+      var #r
+        type-con $ctype 0
+      var #d
+        type-con $ctype 0
+      pat 1
+        fun $map_t 2
+        var #r
+          type-con $ctype 0
+        var #d
+          type-con $ctype 0
+      attribute qid 1
+        string-attr VccPrelu.159:15
+      attribute uniqueId 1
+        string-attr 9
+      attribute bvZ3Native 1
+        string-attr False
+    =
+    fun $ptr_level 1
+    fun $map_t 2
+    var #r
+      type-con $ctype 0
+    var #d
+      type-con $ctype 0
+    +
+    fun $ptr_level 1
+    var #r
+      type-con $ctype 0
+    int-num 23
+axiom 0
+    forall 1 1 4
+      var t
+        type-con $ctype 0
+      pat 1
+        fun $is_primitive 1
+        var t
+          type-con $ctype 0
+      attribute qid 1
+        string-attr VccPrelu.167:36
+      attribute uniqueId 1
+        string-attr 10
+      attribute bvZ3Native 1
+        string-attr False
+      attribute weight 1
+        expr-attr
+          int-num 0
+    =
+    fun $is_primitive 1
+    var t
+      type-con $ctype 0
+    =
+    fun $kind_of 1
+    var t
+      type-con $ctype 0
+    fun $kind_primitive 0
+axiom 0
+    forall 1 1 4
+      var t
+        type-con $ctype 0
+      pat 1
+        fun $is_composite 1
+        var t
+          type-con $ctype 0
+      attribute qid 1
+        string-attr VccPrelu.173:36
+      attribute uniqueId 1
+        string-attr 11
+      attribute bvZ3Native 1
+        string-attr False
+      attribute weight 1
+        expr-attr
+          int-num 0
+    =
+    fun $is_composite 1
+    var t
+      type-con $ctype 0
+    =
+    fun $kind_of 1
+    var t
+      type-con $ctype 0
+    fun $kind_composite 0
+axiom 0
+    forall 1 1 4
+      var t
+        type-con $ctype 0
+      pat 1
+        fun $is_arraytype 1
+        var t
+          type-con $ctype 0
+      attribute qid 1
+        string-attr VccPrelu.179:36
+      attribute uniqueId 1
+        string-attr 12
+      attribute bvZ3Native 1
+        string-attr False
+      attribute weight 1
+        expr-attr
+          int-num 0
+    =
+    fun $is_arraytype 1
+    var t
+      type-con $ctype 0
+    =
+    fun $kind_of 1
+    var t
+      type-con $ctype 0
+    fun $kind_array 0
+axiom 0
+    forall 1 1 4
+      var t
+        type-con $ctype 0
+      pat 1
+        fun $is_threadtype 1
+        var t
+          type-con $ctype 0
+      attribute qid 1
+        string-attr VccPrelu.185:37
+      attribute uniqueId 1
+        string-attr 13
+      attribute bvZ3Native 1
+        string-attr False
+      attribute weight 1
+        expr-attr
+          int-num 0
+    =
+    fun $is_threadtype 1
+    var t
+      type-con $ctype 0
+    =
+    fun $kind_of 1
+    var t
+      type-con $ctype 0
+    fun $kind_thread 0
+axiom 0
+    forall 1 1 4
+      var t
+        type-con $ctype 0
+      pat 1
+        fun $is_composite 1
+        var t
+          type-con $ctype 0
+      attribute qid 1
+        string-attr VccPrelu.198:15
+      attribute uniqueId 1
+        string-attr 14
+      attribute bvZ3Native 1
+        string-attr False
+      attribute weight 1
+        expr-attr
+          int-num 0
+    implies
+    fun $is_composite 1
+    var t
+      type-con $ctype 0
+    fun $is_non_primitive 1
+    var t
+      type-con $ctype 0
+axiom 0
+    forall 1 1 4
+      var t
+        type-con $ctype 0
+      pat 1
+        fun $is_arraytype 1
+        var t
+          type-con $ctype 0
+      attribute qid 1
+        string-attr VccPrelu.199:15
+      attribute uniqueId 1
+        string-attr 15
+      attribute bvZ3Native 1
+        string-attr False
+      attribute weight 1
+        expr-attr
+          int-num 0
+    implies
+    fun $is_arraytype 1
+    var t
+      type-con $ctype 0
+    fun $is_non_primitive 1
+    var t
+      type-con $ctype 0
+axiom 0
+    forall 1 1 4
+      var t
+        type-con $ctype 0
+      pat 1
+        fun $is_threadtype 1
+        var t
+          type-con $ctype 0
+      attribute qid 1
+        string-attr VccPrelu.200:15
+      attribute uniqueId 1
+        string-attr 16
+      attribute bvZ3Native 1
+        string-attr False
+      attribute weight 1
+        expr-attr
+          int-num 0
+    implies
+    fun $is_threadtype 1
+    var t
+      type-con $ctype 0
+    fun $is_non_primitive 1
+    var t
+      type-con $ctype 0
+axiom 0
+    forall 2 1 3
+      var #r
+        type-con $ctype 0
+      var #d
+        type-con $ctype 0
+      pat 1
+        fun $map_t 2
+        var #r
+          type-con $ctype 0
+        var #d
+          type-con $ctype 0
+      attribute qid 1
+        string-attr VccPrelu.208:15
+      attribute uniqueId 1
+        string-attr 17
+      attribute bvZ3Native 1
+        string-attr False
+    fun $is_primitive 1
+    fun $map_t 2
+    var #r
+      type-con $ctype 0
+    var #d
+      type-con $ctype 0
+axiom 0
+    forall 1 1 3
+      var #n
+        type-con $ctype 0
+      pat 1
+        fun $ptr_to 1
+        var #n
+          type-con $ctype 0
+      attribute qid 1
+        string-attr VccPrelu.209:15
+      attribute uniqueId 1
+        string-attr 18
+      attribute bvZ3Native 1
+        string-attr False
+    fun $is_primitive 1
+    fun $ptr_to 1
+    var #n
+      type-con $ctype 0
+axiom 0
+    forall 1 1 3
+      var #n
+        type-con $ctype 0
+      pat 1
+        fun $is_primitive 1
+        var #n
+          type-con $ctype 0
+      attribute qid 1
+        string-attr VccPrelu.210:15
+      attribute uniqueId 1
+        string-attr 19
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    fun $is_primitive 1
+    var #n
+      type-con $ctype 0
+    not
+    fun $is_claimable 1
+    var #n
+      type-con $ctype 0
+axiom 0
+    fun $is_primitive 1
+    fun ^^void 0
+axiom 0
+    fun $is_primitive 1
+    fun ^^bool 0
+axiom 0
+    fun $is_primitive 1
+    fun ^^mathint 0
+axiom 0
+    fun $is_primitive 1
+    fun ^$#ptrset 0
+axiom 0
+    fun $is_primitive 1
+    fun ^$#state_t 0
+axiom 0
+    fun $is_threadtype 1
+    fun ^$#thread_id_t 0
+axiom 0
+    fun $is_primitive 1
+    fun ^^i1 0
+axiom 0
+    fun $is_primitive 1
+    fun ^^i2 0
+axiom 0
+    fun $is_primitive 1
+    fun ^^i4 0
+axiom 0
+    fun $is_primitive 1
+    fun ^^i8 0
+axiom 0
+    fun $is_primitive 1
+    fun ^^u1 0
+axiom 0
+    fun $is_primitive 1
+    fun ^^u2 0
+axiom 0
+    fun $is_primitive 1
+    fun ^^u4 0
+axiom 0
+    fun $is_primitive 1
+    fun ^^u8 0
+axiom 0
+    fun $is_primitive 1
+    fun ^^f4 0
+axiom 0
+    fun $is_primitive 1
+    fun ^^f8 0
+axiom 0
+    =
+    fun $me 0
+    fun $ptr 2
+    fun ^$#thread_id_t 0
+    fun $me_ref 0
+axiom 0
+    forall 3 0 4
+      var M
+        type-con $memory_t 0
+      var p
+        type-con $ptr 0
+      var v
+        int
+      attribute qid 1
+        string-attr VccPrelu.238:15
+      attribute uniqueId 1
+        string-attr 20
+      attribute bvZ3Native 1
+        string-attr False
+      attribute weight 1
+        expr-attr
+          int-num 0
+    =
+    fun $select.mem 2
+    fun $store.mem 3
+    var M
+      type-con $memory_t 0
+    var p
+      type-con $ptr 0
+    var v
+      int
+    var p
+      type-con $ptr 0
+    var v
+      int
+axiom 0
+    forall 4 0 4
+      var M
+        type-con $memory_t 0
+      var p
+        type-con $ptr 0
+      var q
+        type-con $ptr 0
+      var v
+        int
+      attribute qid 1
+        string-attr VccPrelu.240:15
+      attribute uniqueId 1
+        string-attr 21
+      attribute bvZ3Native 1
+        string-attr False
+      attribute weight 1
+        expr-attr
+          int-num 0
+    or 2
+    =
+    var p
+      type-con $ptr 0
+    var q
+      type-con $ptr 0
+    =
+    fun $select.mem 2
+    fun $store.mem 3
+    var M
+      type-con $memory_t 0
+    var p
+      type-con $ptr 0
+    var v
+      int
+    var q
+      type-con $ptr 0
+    fun $select.mem 2
+    var M
+      type-con $memory_t 0
+    var q
+      type-con $ptr 0
+axiom 0
+    forall 3 0 4
+      var M
+        type-con $typemap_t 0
+      var p
+        type-con $ptr 0
+      var v
+        type-con $type_state 0
+      attribute qid 1
+        string-attr VccPrelu.249:15
+      attribute uniqueId 1
+        string-attr 22
+      attribute bvZ3Native 1
+        string-attr False
+      attribute weight 1
+        expr-attr
+          int-num 0
+    =
+    fun $select.tm 2
+    fun $store.tm 3
+    var M
+      type-con $typemap_t 0
+    var p
+      type-con $ptr 0
+    var v
+      type-con $type_state 0
+    var p
+      type-con $ptr 0
+    var v
+      type-con $type_state 0
+axiom 0
+    forall 4 0 4
+      var M
+        type-con $typemap_t 0
+      var p
+        type-con $ptr 0
+      var q
+        type-con $ptr 0
+      var v
+        type-con $type_state 0
+      attribute qid 1
+        string-attr VccPrelu.251:15
+      attribute uniqueId 1
+        string-attr 23
+      attribute bvZ3Native 1
+        string-attr False
+      attribute weight 1
+        expr-attr
+          int-num 0
+    or 2
+    =
+    var p
+      type-con $ptr 0
+    var q
+      type-con $ptr 0
+    =
+    fun $select.tm 2
+    fun $store.tm 3
+    var M
+      type-con $typemap_t 0
+    var p
+      type-con $ptr 0
+    var v
+      type-con $type_state 0
+    var q
+      type-con $ptr 0
+    fun $select.tm 2
+    var M
+      type-con $typemap_t 0
+    var q
+      type-con $ptr 0
+axiom 0
+    forall 3 0 4
+      var M
+        type-con $statusmap_t 0
+      var p
+        type-con $ptr 0
+      var v
+        type-con $status 0
+      attribute qid 1
+        string-attr VccPrelu.260:15
+      attribute uniqueId 1
+        string-attr 24
+      attribute bvZ3Native 1
+        string-attr False
+      attribute weight 1
+        expr-attr
+          int-num 0
+    =
+    fun $select.sm 2
+    fun $store.sm 3
+    var M
+      type-con $statusmap_t 0
+    var p
+      type-con $ptr 0
+    var v
+      type-con $status 0
+    var p
+      type-con $ptr 0
+    var v
+      type-con $status 0
+axiom 0
+    forall 4 0 4
+      var M
+        type-con $statusmap_t 0
+      var p
+        type-con $ptr 0
+      var q
+        type-con $ptr 0
+      var v
+        type-con $status 0
+      attribute qid 1
+        string-attr VccPrelu.262:15
+      attribute uniqueId 1
+        string-attr 25
+      attribute bvZ3Native 1
+        string-attr False
+      attribute weight 1
+        expr-attr
+          int-num 0
+    or 2
+    =
+    var p
+      type-con $ptr 0
+    var q
+      type-con $ptr 0
+    =
+    fun $select.sm 2
+    fun $store.sm 3
+    var M
+      type-con $statusmap_t 0
+    var p
+      type-con $ptr 0
+    var v
+      type-con $status 0
+    var q
+      type-con $ptr 0
+    fun $select.sm 2
+    var M
+      type-con $statusmap_t 0
+    var q
+      type-con $ptr 0
+axiom 0
+    forall 3 1 3
+      var p
+        type-con $ptr 0
+      var q
+        type-con $ptr 0
+      var r
+        type-con $ptr 0
+      pat 2
+        fun $extent_hint 2
+        var p
+          type-con $ptr 0
+        var q
+          type-con $ptr 0
+        fun $extent_hint 2
+        var q
+          type-con $ptr 0
+        var r
+          type-con $ptr 0
+      attribute qid 1
+        string-attr VccPrelu.288:15
+      attribute uniqueId 1
+        string-attr 26
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 2
+    fun $extent_hint 2
+    var p
+      type-con $ptr 0
+    var q
+      type-con $ptr 0
+    fun $extent_hint 2
+    var q
+      type-con $ptr 0
+    var r
+      type-con $ptr 0
+    fun $extent_hint 2
+    var p
+      type-con $ptr 0
+    var r
+      type-con $ptr 0
+axiom 0
+    forall 1 1 3
+      var p
+        type-con $ptr 0
+      pat 1
+        fun $typ 1
+        var p
+          type-con $ptr 0
+      attribute qid 1
+        string-attr VccPrelu.290:15
+      attribute uniqueId 1
+        string-attr 27
+      attribute bvZ3Native 1
+        string-attr False
+    fun $extent_hint 2
+    var p
+      type-con $ptr 0
+    var p
+      type-con $ptr 0
+axiom 0
+    forall 4 1 3
+      var t
+        type-con $ctype 0
+      var s
+        type-con $ctype 0
+      var min
+        int
+      var max
+        int
+      pat 1
+        fun $is_nested_range 4
+        var t
+          type-con $ctype 0
+        var s
+          type-con $ctype 0
+        var min
+          int
+        var max
+          int
+      attribute qid 1
+        string-attr VccPrelu.297:27
+      attribute uniqueId 1
+        string-attr 28
+      attribute bvZ3Native 1
+        string-attr False
+    =
+    fun $is_nested_range 4
+    var t
+      type-con $ctype 0
+    var s
+      type-con $ctype 0
+    var min
+      int
+    var max
+      int
+    and 3
+    fun $is_nested 2
+    var t
+      type-con $ctype 0
+    var s
+      type-con $ctype 0
+    =
+    fun $nesting_min 2
+    var t
+      type-con $ctype 0
+    var s
+      type-con $ctype 0
+    var min
+      int
+    =
+    fun $nesting_max 2
+    var t
+      type-con $ctype 0
+    var s
+      type-con $ctype 0
+    var max
+      int
+axiom 0
+    forall 2 0 4
+      var #t
+        type-con $ctype 0
+      var #b
+        int
+      attribute qid 1
+        string-attr VccPrelu.334:15
+      attribute uniqueId 1
+        string-attr 29
+      attribute bvZ3Native 1
+        string-attr False
+      attribute weight 1
+        expr-attr
+          int-num 0
+    =
+    fun $typ 1
+    fun $ptr 2
+    var #t
+      type-con $ctype 0
+    var #b
+      int
+    var #t
+      type-con $ctype 0
+axiom 0
+    forall 2 0 4
+      var #t
+        type-con $ctype 0
+      var #b
+        int
+      attribute qid 1
+        string-attr VccPrelu.335:15
+      attribute uniqueId 1
+        string-attr 30
+      attribute bvZ3Native 1
+        string-attr False
+      attribute weight 1
+        expr-attr
+          int-num 0
+    =
+    fun $ref 1
+    fun $ptr 2
+    var #t
+      type-con $ctype 0
+    var #b
+      int
+    var #b
+      int
+axiom 0
+    forall 2 1 4
+      var p
+        type-con $ptr 0
+      var f
+        type-con $field 0
+      pat 1
+        fun $ghost_ref 2
+        var p
+          type-con $ptr 0
+        var f
+          type-con $field 0
+      attribute qid 1
+        string-attr VccPrelu.344:15
+      attribute uniqueId 1
+        string-attr 31
+      attribute bvZ3Native 1
+        string-attr False
+      attribute weight 1
+        expr-attr
+          int-num 0
+    and 2
+    =
+    fun $ghost_emb 1
+    fun $ghost_ref 2
+    var p
+      type-con $ptr 0
+    var f
+      type-con $field 0
+    var p
+      type-con $ptr 0
+    =
+    fun $ghost_path 1
+    fun $ghost_ref 2
+    var p
+      type-con $ptr 0
+    var f
+      type-con $field 0
+    var f
+      type-con $field 0
+axiom 0
+    forall 2 1 4
+      var fld
+        type-con $field 0
+      var off
+        int
+      pat 1
+        fun $array_path 2
+        var fld
+          type-con $field 0
+        var off
+          int
+      attribute qid 1
+        string-attr VccPrelu.355:15
+      attribute uniqueId 1
+        string-attr 32
+      attribute bvZ3Native 1
+        string-attr False
+      attribute weight 1
+        expr-attr
+          int-num 0
+    and 3
+    not
+    fun $is_base_field 1
+    fun $array_path 2
+    var fld
+      type-con $field 0
+    var off
+      int
+    =
+    fun $array_path_1 1
+    fun $array_path 2
+    var fld
+      type-con $field 0
+    var off
+      int
+    var fld
+      type-con $field 0
+    =
+    fun $array_path_2 1
+    fun $array_path 2
+    var fld
+      type-con $field 0
+    var off
+      int
+    var off
+      int
+axiom 0
+    =
+    fun $null 0
+    fun $ptr 2
+    fun ^^void 0
+    int-num 0
+axiom 0
+    forall 2 0 4
+      var #p
+        type-con $ptr 0
+      var #t
+        type-con $ctype 0
+      attribute qid 1
+        string-attr VccPrelu.368:15
+      attribute uniqueId 1
+        string-attr 33
+      attribute bvZ3Native 1
+        string-attr False
+      attribute weight 1
+        expr-attr
+          int-num 0
+    =
+    fun $is 2
+    var #p
+      type-con $ptr 0
+    var #t
+      type-con $ctype 0
+    =
+    fun $typ 1
+    var #p
+      type-con $ptr 0
+    var #t
+      type-con $ctype 0
+axiom 0
+    forall 2 1 3
+      var #p
+        type-con $ptr 0
+      var #t
+        type-con $ctype 0
+      pat 1
+        fun $is 2
+        var #p
+          type-con $ptr 0
+        var #t
+          type-con $ctype 0
+      attribute qid 1
+        string-attr VccPrelu.370:15
+      attribute uniqueId 1
+        string-attr 34
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    fun $is 2
+    var #p
+      type-con $ptr 0
+    var #t
+      type-con $ctype 0
+    =
+    var #p
+      type-con $ptr 0
+    fun $ptr 2
+    var #t
+      type-con $ctype 0
+    fun $ref 1
+    var #p
+      type-con $ptr 0
+axiom 0
+    forall 2 1 3
+      var r
+        int
+      var f
+        type-con $field 0
+      pat 1
+        fun $containing_struct 2
+        fun $dot 2
+        fun $ptr 2
+        fun $field_parent_type 1
+        var f
+          type-con $field 0
+        var r
+          int
+        var f
+          type-con $field 0
+        var f
+          type-con $field 0
+      attribute qid 1
+        string-attr VccPrelu.388:15
+      attribute uniqueId 1
+        string-attr 35
+      attribute bvZ3Native 1
+        string-attr False
+    =
+    fun $containing_struct 2
+    fun $dot 2
+    fun $ptr 2
+    fun $field_parent_type 1
+    var f
+      type-con $field 0
+    var r
+      int
+    var f
+      type-con $field 0
+    var f
+      type-con $field 0
+    fun $ptr 2
+    fun $field_parent_type 1
+    var f
+      type-con $field 0
+    var r
+      int
+axiom 0
+    forall 2 1 3
+      var p
+        type-con $ptr 0
+      var f
+        type-con $field 0
+      pat 1
+        fun $containing_struct 2
+        var p
+          type-con $ptr 0
+        var f
+          type-con $field 0
+      attribute qid 1
+        string-attr VccPrelu.392:15
+      attribute uniqueId 1
+        string-attr 36
+      attribute bvZ3Native 1
+        string-attr False
+    =
+    fun $containing_struct 2
+    var p
+      type-con $ptr 0
+    var f
+      type-con $field 0
+    fun $ptr 2
+    fun $field_parent_type 1
+    var f
+      type-con $field 0
+    fun $containing_struct_ref 2
+    var p
+      type-con $ptr 0
+    var f
+      type-con $field 0
+axiom 0
+    forall 2 1 3
+      var p
+        type-con $ptr 0
+      var f
+        type-con $field 0
+      pat 1
+        fun $dot 2
+        fun $containing_struct 2
+        var p
+          type-con $ptr 0
+        var f
+          type-con $field 0
+        var f
+          type-con $field 0
+      attribute qid 1
+        string-attr VccPrelu.396:15
+      attribute uniqueId 1
+        string-attr 37
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    >=
+    fun $field_offset 1
+    var f
+      type-con $field 0
+    int-num 0
+    =
+    fun $ref 1
+    fun $dot 2
+    fun $containing_struct 2
+    var p
+      type-con $ptr 0
+    var f
+      type-con $field 0
+    var f
+      type-con $field 0
+    fun $ref 1
+    var p
+      type-con $ptr 0
+axiom 0
+    forall 1 1 3
+      var ts
+        type-con $type_state 0
+      pat 1
+        fun $ts_emb 1
+        var ts
+          type-con $type_state 0
+      attribute qid 1
+        string-attr VccPrelu.427:15
+      attribute uniqueId 1
+        string-attr 38
+      attribute bvZ3Native 1
+        string-attr False
+    and 2
+    not
+    =
+    fun $kind_of 1
+    fun $typ 1
+    fun $ts_emb 1
+    var ts
+      type-con $type_state 0
+    fun $kind_primitive 0
+    fun $is_non_primitive 1
+    fun $typ 1
+    fun $ts_emb 1
+    var ts
+      type-con $type_state 0
+axiom 0
+    forall 2 1 3
+      var S
+        type-con $state 0
+      var p
+        type-con $ptr 0
+      pat 2
+        fun $typed 2
+        var S
+          type-con $state 0
+        var p
+          type-con $ptr 0
+        fun $select.tm 2
+        fun $typemap 1
+        var S
+          type-con $state 0
+        fun $ts_emb 1
+        fun $select.tm 2
+        fun $typemap 1
+        var S
+          type-con $state 0
+        var p
+          type-con $ptr 0
+      attribute qid 1
+        string-attr VccPrelu.430:15
+      attribute uniqueId 1
+        string-attr 39
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    fun $typed 2
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    fun $typed 2
+    var S
+      type-con $state 0
+    fun $ts_emb 1
+    fun $select.tm 2
+    fun $typemap 1
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+axiom 0
+    forall 2 1 3
+      var S
+        type-con $state 0
+      var p
+        type-con $ptr 0
+      pat 1
+        fun $ts_is_volatile 1
+        fun $select.tm 2
+        fun $typemap 1
+        var S
+          type-con $state 0
+        var p
+          type-con $ptr 0
+      attribute qid 1
+        string-attr VccPrelu.440:15
+      attribute uniqueId 1
+        string-attr 40
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 2
+    fun $good_state 1
+    var S
+      type-con $state 0
+    fun $ts_is_volatile 1
+    fun $select.tm 2
+    fun $typemap 1
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    =
+    fun $kind_of 1
+    fun $typ 1
+    var p
+      type-con $ptr 0
+    fun $kind_primitive 0
+axiom 0
+    forall 2 1 4
+      var S
+        type-con $state 0
+      var p
+        type-con $ptr 0
+      pat 1
+        fun $select.sm 2
+        fun $statusmap 1
+        var S
+          type-con $state 0
+        var p
+          type-con $ptr 0
+      attribute qid 1
+        string-attr VccPrelu.456:15
+      attribute uniqueId 1
+        string-attr 41
+      attribute bvZ3Native 1
+        string-attr False
+      attribute weight 1
+        expr-attr
+          int-num 0
+    or 2
+    <=
+    fun $timestamp 2
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    fun $current_timestamp 1
+    var S
+      type-con $state 0
+    not
+    fun $ts_typed 1
+    fun $select.tm 2
+    fun $typemap 1
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+axiom 0
+    fun $good_state 1
+    fun $vs_state 1
+    fun $struct_zero 0
+axiom 0
+    forall 1 0 3
+      var s
+        type-con $struct 0
+      attribute qid 1
+        string-attr VccPrelu.486:15
+      attribute uniqueId 1
+        string-attr 42
+      attribute bvZ3Native 1
+        string-attr False
+    fun $good_state 1
+    fun $vs_state 1
+    var s
+      type-con $struct 0
+axiom 0
+    forall 2 1 3
+      var S
+        type-con $state 0
+      var p
+        type-con $ptr 0
+      pat 1
+        fun $vs_ctor 2
+        var S
+          type-con $state 0
+        var p
+          type-con $ptr 0
+      attribute qid 1
+        string-attr VccPrelu.489:15
+      attribute uniqueId 1
+        string-attr 43
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    fun $good_state 1
+    var S
+      type-con $state 0
+    and 2
+    =
+    fun $vs_base_ref 1
+    fun $vs_ctor 2
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    fun $ref 1
+    var p
+      type-con $ptr 0
+    =
+    fun $vs_state 1
+    fun $vs_ctor 2
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    var S
+      type-con $state 0
+axiom 0
+    forall 6 1 3
+      var r
+        type-con $record 0
+      var f
+        type-con $field 0
+      var val_bitsize
+        int
+      var from
+        int
+      var to
+        int
+      var repl
+        int
+      pat 1
+        fun $rec_update_bv 6
+        var r
+          type-con $record 0
+        var f
+          type-con $field 0
+        var val_bitsize
+          int
+        var from
+          int
+        var to
+          int
+        var repl
+          int
+      attribute qid 1
+        string-attr VccPrelu.502:25
+      attribute uniqueId 1
+        string-attr 44
+      attribute bvZ3Native 1
+        string-attr False
+    =
+    fun $rec_update_bv 6
+    var r
+      type-con $record 0
+    var f
+      type-con $field 0
+    var val_bitsize
+      int
+    var from
+      int
+    var to
+      int
+    var repl
+      int
+    fun $rec_update 3
+    var r
+      type-con $record 0
+    var f
+      type-con $field 0
+    fun $bv_update 5
+    fun $rec_fetch 2
+    var r
+      type-con $record 0
+    var f
+      type-con $field 0
+    var val_bitsize
+      int
+    var from
+      int
+    var to
+      int
+    var repl
+      int
+axiom 0
+    forall 1 0 3
+      var f
+        type-con $field 0
+      attribute qid 1
+        string-attr VccPrelu.505:15
+      attribute uniqueId 1
+        string-attr 45
+      attribute bvZ3Native 1
+        string-attr False
+    =
+    fun $rec_fetch 2
+    fun $rec_zero 0
+    var f
+      type-con $field 0
+    int-num 0
+axiom 0
+    forall 3 1 3
+      var r
+        type-con $record 0
+      var f
+        type-con $field 0
+      var v
+        int
+      pat 1
+        fun $rec_fetch 2
+        fun $rec_update 3
+        var r
+          type-con $record 0
+        var f
+          type-con $field 0
+        var v
+          int
+        var f
+          type-con $field 0
+      attribute qid 1
+        string-attr VccPrelu.507:15
+      attribute uniqueId 1
+        string-attr 46
+      attribute bvZ3Native 1
+        string-attr False
+    =
+    fun $rec_fetch 2
+    fun $rec_update 3
+    var r
+      type-con $record 0
+    var f
+      type-con $field 0
+    var v
+      int
+    var f
+      type-con $field 0
+    var v
+      int
+axiom 0
+    forall 4 1 3
+      var r
+        type-con $record 0
+      var f1
+        type-con $field 0
+      var f2
+        type-con $field 0
+      var v
+        int
+      pat 1
+        fun $rec_fetch 2
+        fun $rec_update 3
+        var r
+          type-con $record 0
+        var f1
+          type-con $field 0
+        var v
+          int
+        var f2
+          type-con $field 0
+      attribute qid 1
+        string-attr VccPrelu.510:15
+      attribute uniqueId 1
+        string-attr 47
+      attribute bvZ3Native 1
+        string-attr False
+    or 2
+    =
+    fun $rec_fetch 2
+    fun $rec_update 3
+    var r
+      type-con $record 0
+    var f1
+      type-con $field 0
+    var v
+      int
+    var f2
+      type-con $field 0
+    fun $rec_fetch 2
+    var r
+      type-con $record 0
+    var f2
+      type-con $field 0
+    =
+    var f1
+      type-con $field 0
+    var f2
+      type-con $field 0
+axiom 0
+    forall 1 1 3
+      var t
+        type-con $ctype 0
+      pat 1
+        fun $is_record_type 1
+        var t
+          type-con $ctype 0
+      attribute qid 1
+        string-attr VccPrelu.516:15
+      attribute uniqueId 1
+        string-attr 48
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    fun $is_record_type 1
+    var t
+      type-con $ctype 0
+    fun $is_primitive 1
+    var t
+      type-con $ctype 0
+axiom 0
+    forall 3 1 3
+      var p
+        type-con $ctype 0
+      var f
+        type-con $field 0
+      var ft
+        type-con $ctype 0
+      pat 2
+        fun $is_record_field 3
+        var p
+          type-con $ctype 0
+        var f
+          type-con $field 0
+        var ft
+          type-con $ctype 0
+        fun $is_record_type 1
+        var ft
+          type-con $ctype 0
+      attribute qid 1
+        string-attr VccPrelu.519:15
+      attribute uniqueId 1
+        string-attr 49
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 2
+    fun $is_record_field 3
+    var p
+      type-con $ctype 0
+    var f
+      type-con $field 0
+    var ft
+      type-con $ctype 0
+    fun $is_record_type 1
+    var ft
+      type-con $ctype 0
+    =
+    fun $as_record_record_field 1
+    var f
+      type-con $field 0
+    var f
+      type-con $field 0
+axiom 0
+    forall 2 1 3
+      var r1
+        type-con $record 0
+      var r2
+        type-con $record 0
+      pat 1
+        fun $rec_eq 2
+        var r1
+          type-con $record 0
+        var r2
+          type-con $record 0
+      attribute qid 1
+        string-attr VccPrelu.522:18
+      attribute uniqueId 1
+        string-attr 50
+      attribute bvZ3Native 1
+        string-attr False
+    =
+    fun $rec_eq 2
+    var r1
+      type-con $record 0
+    var r2
+      type-con $record 0
+    =
+    var r1
+      type-con $record 0
+    var r2
+      type-con $record 0
+axiom 0
+    forall 2 1 3
+      var x
+        int
+      var y
+        int
+      pat 1
+        fun $rec_base_eq 2
+        var x
+          int
+        var y
+          int
+      attribute qid 1
+        string-attr VccPrelu.524:23
+      attribute uniqueId 1
+        string-attr 51
+      attribute bvZ3Native 1
+        string-attr False
+    =
+    fun $rec_base_eq 2
+    var x
+      int
+    var y
+      int
+    =
+    var x
+      int
+    var y
+      int
+axiom 0
+    forall 1 0 3
+      var r
+        type-con $record 0
+      attribute qid 1
+        string-attr VccPrelu.530:15
+      attribute uniqueId 1
+        string-attr 52
+      attribute bvZ3Native 1
+        string-attr False
+    =
+    fun $int_to_record 1
+    fun $record_to_int 1
+    var r
+      type-con $record 0
+    var r
+      type-con $record 0
+axiom 0
+    forall 2 1 3
+      var r1
+        type-con $record 0
+      var r2
+        type-con $record 0
+      pat 1
+        fun $rec_eq 2
+        var r1
+          type-con $record 0
+        var r2
+          type-con $record 0
+      attribute qid 1
+        string-attr VccPrelu.532:15
+      attribute uniqueId 1
+        string-attr 54
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    forall 1 0 3
+      var f
+        type-con $field 0
+      attribute qid 1
+        string-attr VccPrelu.534:11
+      attribute uniqueId 1
+        string-attr 53
+      attribute bvZ3Native 1
+        string-attr False
+    fun $rec_base_eq 2
+    fun $rec_fetch 2
+    var r1
+      type-con $record 0
+    var f
+      type-con $field 0
+    fun $rec_fetch 2
+    var r2
+      type-con $record 0
+    var f
+      type-con $field 0
+    fun $rec_eq 2
+    var r1
+      type-con $record 0
+    var r2
+      type-con $record 0
+axiom 0
+    forall 3 1 3
+      var r1
+        type-con $record 0
+      var r2
+        type-con $record 0
+      var f
+        type-con $field 0
+      pat 1
+        fun $rec_base_eq 2
+        fun $rec_fetch 2
+        var r1
+          type-con $record 0
+        var f
+          type-con $field 0
+        fun $rec_fetch 2
+        var r2
+          type-con $record 0
+        fun $as_record_record_field 1
+        var f
+          type-con $field 0
+      attribute qid 1
+        string-attr VccPrelu.536:15
+      attribute uniqueId 1
+        string-attr 55
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    fun $rec_eq 2
+    fun $int_to_record 1
+    fun $rec_fetch 2
+    var r1
+      type-con $record 0
+    var f
+      type-con $field 0
+    fun $int_to_record 1
+    fun $rec_fetch 2
+    var r2
+      type-con $record 0
+    var f
+      type-con $field 0
+    fun $rec_base_eq 2
+    fun $rec_fetch 2
+    var r1
+      type-con $record 0
+    var f
+      type-con $field 0
+    fun $rec_fetch 2
+    var r2
+      type-con $record 0
+    var f
+      type-con $field 0
+axiom 0
+    fun $has_volatile_owns_set 1
+    fun ^^claim 0
+axiom 0
+    forall 2 1 3
+      var #p
+        type-con $ptr 0
+      var t
+        type-con $ctype 0
+      pat 1
+        fun $dot 2
+        var #p
+          type-con $ptr 0
+        fun $owns_set_field 1
+        var t
+          type-con $ctype 0
+      attribute qid 1
+        string-attr VccPrelu.555:15
+      attribute uniqueId 1
+        string-attr 56
+      attribute bvZ3Native 1
+        string-attr False
+    =
+    fun $dot 2
+    var #p
+      type-con $ptr 0
+    fun $owns_set_field 1
+    var t
+      type-con $ctype 0
+    fun $ptr 2
+    fun ^$#ptrset 0
+    fun $ghost_ref 2
+    var #p
+      type-con $ptr 0
+    fun $owns_set_field 1
+    var t
+      type-con $ctype 0
+axiom 0
+    forall 2 1 4
+      var S
+        type-con $state 0
+      var p
+        type-con $ptr 0
+      pat 2
+        fun $is_primitive 1
+        fun $typ 1
+        var p
+          type-con $ptr 0
+        fun $owner 2
+        var S
+          type-con $state 0
+        var p
+          type-con $ptr 0
+      attribute qid 1
+        string-attr VccPrelu.567:15
+      attribute uniqueId 1
+        string-attr 57
+      attribute bvZ3Native 1
+        string-attr False
+      attribute weight 1
+        expr-attr
+          int-num 0
+    implies
+    fun $is_primitive 1
+    fun $typ 1
+    var p
+      type-con $ptr 0
+    =
+    fun $owner 2
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    fun $owner 2
+    var S
+      type-con $state 0
+    fun $ts_emb 1
+    fun $select.tm 2
+    fun $typemap 1
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+axiom 0
+    forall 2 1 4
+      var S
+        type-con $state 0
+      var p
+        type-con $ptr 0
+      pat 2
+        fun $is_non_primitive 1
+        fun $typ 1
+        var p
+          type-con $ptr 0
+        fun $owner 2
+        var S
+          type-con $state 0
+        var p
+          type-con $ptr 0
+      attribute qid 1
+        string-attr VccPrelu.569:15
+      attribute uniqueId 1
+        string-attr 58
+      attribute bvZ3Native 1
+        string-attr False
+      attribute weight 1
+        expr-attr
+          int-num 0
+    implies
+    fun $is_non_primitive 1
+    fun $typ 1
+    var p
+      type-con $ptr 0
+    =
+    fun $owner 2
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    fun $st_owner 1
+    fun $select.sm 2
+    fun $statusmap 1
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+axiom 0
+    forall 2 1 4
+      var S
+        type-con $state 0
+      var p
+        type-con $ptr 0
+      pat 2
+        fun $is_primitive 1
+        fun $typ 1
+        var p
+          type-con $ptr 0
+        fun $closed 2
+        var S
+          type-con $state 0
+        var p
+          type-con $ptr 0
+      attribute qid 1
+        string-attr VccPrelu.572:15
+      attribute uniqueId 1
+        string-attr 59
+      attribute bvZ3Native 1
+        string-attr False
+      attribute weight 1
+        expr-attr
+          int-num 0
+    implies
+    fun $is_primitive 1
+    fun $typ 1
+    var p
+      type-con $ptr 0
+    =
+    fun $closed 2
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    fun $st_closed 1
+    fun $select.sm 2
+    fun $statusmap 1
+    var S
+      type-con $state 0
+    fun $ts_emb 1
+    fun $select.tm 2
+    fun $typemap 1
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+axiom 0
+    forall 2 1 4
+      var S
+        type-con $state 0
+      var p
+        type-con $ptr 0
+      pat 2
+        fun $is_non_primitive 1
+        fun $typ 1
+        var p
+          type-con $ptr 0
+        fun $closed 2
+        var S
+          type-con $state 0
+        var p
+          type-con $ptr 0
+      attribute qid 1
+        string-attr VccPrelu.574:15
+      attribute uniqueId 1
+        string-attr 60
+      attribute bvZ3Native 1
+        string-attr False
+      attribute weight 1
+        expr-attr
+          int-num 0
+    implies
+    fun $is_non_primitive 1
+    fun $typ 1
+    var p
+      type-con $ptr 0
+    =
+    fun $closed 2
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    fun $st_closed 1
+    fun $select.sm 2
+    fun $statusmap 1
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+axiom 0
+    forall 2 1 4
+      var S
+        type-con $state 0
+      var p
+        type-con $ptr 0
+      pat 2
+        fun $is_primitive 1
+        fun $typ 1
+        var p
+          type-con $ptr 0
+        fun $timestamp 2
+        var S
+          type-con $state 0
+        var p
+          type-con $ptr 0
+      attribute qid 1
+        string-attr VccPrelu.577:15
+      attribute uniqueId 1
+        string-attr 61
+      attribute bvZ3Native 1
+        string-attr False
+      attribute weight 1
+        expr-attr
+          int-num 0
+    implies
+    fun $is_primitive 1
+    fun $typ 1
+    var p
+      type-con $ptr 0
+    =
+    fun $timestamp 2
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    fun $st_timestamp 1
+    fun $select.sm 2
+    fun $statusmap 1
+    var S
+      type-con $state 0
+    fun $ts_emb 1
+    fun $select.tm 2
+    fun $typemap 1
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+axiom 0
+    forall 2 1 4
+      var S
+        type-con $state 0
+      var p
+        type-con $ptr 0
+      pat 2
+        fun $is_non_primitive 1
+        fun $typ 1
+        var p
+          type-con $ptr 0
+        fun $timestamp 2
+        var S
+          type-con $state 0
+        var p
+          type-con $ptr 0
+      attribute qid 1
+        string-attr VccPrelu.579:15
+      attribute uniqueId 1
+        string-attr 62
+      attribute bvZ3Native 1
+        string-attr False
+      attribute weight 1
+        expr-attr
+          int-num 0
+    implies
+    fun $is_non_primitive 1
+    fun $typ 1
+    var p
+      type-con $ptr 0
+    =
+    fun $timestamp 2
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    fun $st_timestamp 1
+    fun $select.sm 2
+    fun $statusmap 1
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+axiom 0
+    fun $position_marker 0
+axiom 0
+    forall 1 1 3
+      var s
+        type-con $status 0
+      pat 1
+        fun $st_owner 1
+        var s
+          type-con $status 0
+      attribute qid 1
+        string-attr VccPrelu.585:15
+      attribute uniqueId 1
+        string-attr 63
+      attribute bvZ3Native 1
+        string-attr False
+    and 2
+    not
+    =
+    fun $kind_of 1
+    fun $typ 1
+    fun $st_owner 1
+    var s
+      type-con $status 0
+    fun $kind_primitive 0
+    fun $is_non_primitive 1
+    fun $typ 1
+    fun $st_owner 1
+    var s
+      type-con $status 0
+axiom 0
+    forall 2 1 4
+      var S
+        type-con $state 0
+      var #p
+        type-con $ptr 0
+      pat 1
+        fun $owns 2
+        var S
+          type-con $state 0
+        var #p
+          type-con $ptr 0
+      attribute qid 1
+        string-attr VccPrelu.593:28
+      attribute uniqueId 1
+        string-attr 64
+      attribute bvZ3Native 1
+        string-attr False
+      attribute weight 1
+        expr-attr
+          int-num 0
+    =
+    fun $owns 2
+    var S
+      type-con $state 0
+    var #p
+      type-con $ptr 0
+    fun $int_to_ptrset 1
+    fun $select.mem 2
+    fun $memory 1
+    var S
+      type-con $state 0
+    fun $dot 2
+    var #p
+      type-con $ptr 0
+    fun $owns_set_field 1
+    fun $typ 1
+    var #p
+      type-con $ptr 0
+axiom 0
+    forall 2 1 4
+      var S
+        type-con $state 0
+      var p
+        type-con $ptr 0
+      pat 1
+        fun $mutable 2
+        var S
+          type-con $state 0
+        var p
+          type-con $ptr 0
+      attribute qid 1
+        string-attr VccPrelu.608:31
+      attribute uniqueId 1
+        string-attr 65
+      attribute bvZ3Native 1
+        string-attr False
+      attribute weight 1
+        expr-attr
+          int-num 0
+    =
+    fun $mutable 2
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    and 3
+    fun $typed 2
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    =
+    fun $owner 2
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    fun $me 0
+    not
+    fun $closed 2
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+axiom 0
+    forall 2 1 4
+      var S
+        type-con $state 0
+      var #p
+        type-con $ptr 0
+      pat 1
+        fun $typed 2
+        var S
+          type-con $state 0
+        var #p
+          type-con $ptr 0
+      attribute qid 1
+        string-attr VccPrelu.619:11
+      attribute uniqueId 1
+        string-attr 66
+      attribute bvZ3Native 1
+        string-attr False
+      attribute weight 1
+        expr-attr
+          int-num 0
+    implies
+    fun $good_state 1
+    var S
+      type-con $state 0
+    =
+    fun $typed 2
+    var S
+      type-con $state 0
+    var #p
+      type-con $ptr 0
+    fun $ts_typed 1
+    fun $select.tm 2
+    fun $typemap 1
+    var S
+      type-con $state 0
+    var #p
+      type-con $ptr 0
+axiom 0
+    forall 2 1 3
+      var S
+        type-con $state 0
+      var #p
+        type-con $ptr 0
+      pat 1
+        fun $typed 2
+        var S
+          type-con $state 0
+        var #p
+          type-con $ptr 0
+      attribute qid 1
+        string-attr VccPrelu.621:11
+      attribute uniqueId 1
+        string-attr 67
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 2
+    fun $good_state 1
+    var S
+      type-con $state 0
+    fun $typed 2
+    var S
+      type-con $state 0
+    var #p
+      type-con $ptr 0
+    >
+    fun $ref 1
+    var #p
+      type-con $ptr 0
+    int-num 0
+axiom 0
+    forall 3 1 3
+      var S1
+        type-con $state 0
+      var S2
+        type-con $state 0
+      var p
+        type-con $ptr 0
+      pat 2
+        fun $select.sm 2
+        fun $statusmap 1
+        var S2
+          type-con $state 0
+        var p
+          type-con $ptr 0
+        fun $call_transition 2
+        var S1
+          type-con $state 0
+        var S2
+          type-con $state 0
+      attribute qid 1
+        string-attr VccPrelu.685:15
+      attribute uniqueId 1
+        string-attr 68
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    fun $call_transition 2
+    var S1
+      type-con $state 0
+    var S2
+      type-con $state 0
+    fun $instantiate_st 1
+    fun $select.sm 2
+    fun $statusmap 1
+    var S1
+      type-con $state 0
+    var p
+      type-con $ptr 0
+axiom 0
+    forall 2 1 3
+      var S
+        type-con $state 0
+      var p
+        type-con $ptr 0
+      pat 1
+        fun $is_domain_root 2
+        var S
+          type-con $state 0
+        var p
+          type-con $ptr 0
+      attribute qid 1
+        string-attr VccPrelu.711:26
+      attribute uniqueId 1
+        string-attr 69
+      attribute bvZ3Native 1
+        string-attr False
+    =
+    fun $is_domain_root 2
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    true
+axiom 0
+    forall 2 1 3
+      var S
+        type-con $state 0
+      var p
+        type-con $ptr 0
+      pat 1
+        fun $in_wrapped_domain 2
+        var S
+          type-con $state 0
+        var p
+          type-con $ptr 0
+      attribute qid 1
+        string-attr VccPrelu.714:29
+      attribute uniqueId 1
+        string-attr 71
+      attribute bvZ3Native 1
+        string-attr False
+    =
+    fun $in_wrapped_domain 2
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    exists 1 1 3
+      var q
+        type-con $ptr 0
+      pat 1
+        fun $set_in2 2
+        var p
+          type-con $ptr 0
+        fun $ver_domain 1
+        fun $read_version 2
+        var S
+          type-con $state 0
+        var q
+          type-con $ptr 0
+      attribute qid 1
+        string-attr VccPrelu.715:13
+      attribute uniqueId 1
+        string-attr 70
+      attribute bvZ3Native 1
+        string-attr False
+    and 8
+    fun $set_in 2
+    var p
+      type-con $ptr 0
+    fun $ver_domain 1
+    fun $read_version 2
+    var S
+      type-con $state 0
+    var q
+      type-con $ptr 0
+    fun $closed 2
+    var S
+      type-con $state 0
+    var q
+      type-con $ptr 0
+    =
+    fun $owner 2
+    var S
+      type-con $state 0
+    var q
+      type-con $ptr 0
+    fun $me 0
+    fun $is 2
+    var q
+      type-con $ptr 0
+    fun $typ 1
+    var q
+      type-con $ptr 0
+    fun $typed 2
+    var S
+      type-con $state 0
+    var q
+      type-con $ptr 0
+    not
+    =
+    fun $kind_of 1
+    fun $typ 1
+    var q
+      type-con $ptr 0
+    fun $kind_primitive 0
+    fun $is_non_primitive 1
+    fun $typ 1
+    var q
+      type-con $ptr 0
+    fun $is_domain_root 2
+    var S
+      type-con $state 0
+    var q
+      type-con $ptr 0
+axiom 0
+    forall 2 1 3
+      var S
+        type-con $state 0
+      var p
+        type-con $ptr 0
+      pat 1
+        fun $thread_local 2
+        var S
+          type-con $state 0
+        var p
+          type-con $ptr 0
+      attribute qid 1
+        string-attr VccPrelu.728:24
+      attribute uniqueId 1
+        string-attr 72
+      attribute bvZ3Native 1
+        string-attr False
+    =
+    fun $thread_local 2
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    and 2
+    fun $typed 2
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    or 2
+    and 4
+    =
+    fun $kind_of 1
+    fun $typ 1
+    var p
+      type-con $ptr 0
+    fun $kind_primitive 0
+    or 2
+    not
+    fun $ts_is_volatile 1
+    fun $select.tm 2
+    fun $typemap 1
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    not
+    fun $closed 2
+    var S
+      type-con $state 0
+    fun $ts_emb 1
+    fun $select.tm 2
+    fun $typemap 1
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    not
+    =
+    fun $kind_of 1
+    fun $typ 1
+    fun $ts_emb 1
+    fun $select.tm 2
+    fun $typemap 1
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    fun $kind_primitive 0
+    or 2
+    =
+    fun $owner 2
+    var S
+      type-con $state 0
+    fun $ts_emb 1
+    fun $select.tm 2
+    fun $typemap 1
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    fun $me 0
+    fun $in_wrapped_domain 2
+    var S
+      type-con $state 0
+    fun $ts_emb 1
+    fun $select.tm 2
+    fun $typemap 1
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    and 2
+    not
+    =
+    fun $kind_of 1
+    fun $typ 1
+    var p
+      type-con $ptr 0
+    fun $kind_primitive 0
+    or 2
+    =
+    fun $owner 2
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    fun $me 0
+    fun $in_wrapped_domain 2
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+axiom 0
+    forall 3 1 3
+      var #s1
+        type-con $state 0
+      var #p
+        type-con $ptr 0
+      var typ
+        type-con $ctype 0
+      pat 1
+        fun $inv2 4
+        var #s1
+          type-con $state 0
+        var #s1
+          type-con $state 0
+        var #p
+          type-con $ptr 0
+        var typ
+          type-con $ctype 0
+      attribute qid 1
+        string-attr VccPrelu.766:15
+      attribute uniqueId 1
+        string-attr 73
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    fun $imply_inv 3
+    var #s1
+      type-con $state 0
+    var #p
+      type-con $ptr 0
+    var typ
+      type-con $ctype 0
+    fun $inv2 4
+    var #s1
+      type-con $state 0
+    var #s1
+      type-con $state 0
+    var #p
+      type-con $ptr 0
+    var typ
+      type-con $ctype 0
+axiom 0
+    forall 4 1 4
+      var #s1
+        type-con $state 0
+      var #s2
+        type-con $state 0
+      var #p
+        type-con $ptr 0
+      var #t
+        type-con $ctype 0
+      pat 1
+        fun $sequential 4
+        var #s1
+          type-con $state 0
+        var #s2
+          type-con $state 0
+        var #p
+          type-con $ptr 0
+        var #t
+          type-con $ctype 0
+      attribute qid 1
+        string-attr VccPrelu.778:34
+      attribute uniqueId 1
+        string-attr 74
+      attribute bvZ3Native 1
+        string-attr False
+      attribute weight 1
+        expr-attr
+          int-num 0
+    =
+    fun $sequential 4
+    var #s1
+      type-con $state 0
+    var #s2
+      type-con $state 0
+    var #p
+      type-con $ptr 0
+    var #t
+      type-con $ctype 0
+    implies
+    and 2
+    fun $closed 2
+    var #s1
+      type-con $state 0
+    var #p
+      type-con $ptr 0
+    fun $closed 2
+    var #s2
+      type-con $state 0
+    var #p
+      type-con $ptr 0
+    fun $spans_the_same 4
+    var #s1
+      type-con $state 0
+    var #s2
+      type-con $state 0
+    var #p
+      type-con $ptr 0
+    var #t
+      type-con $ctype 0
+axiom 0
+    forall 4 1 4
+      var #s1
+        type-con $state 0
+      var #s2
+        type-con $state 0
+      var #dependant
+        type-con $ptr 0
+      var #this
+        type-con $ptr 0
+      pat 1
+        fun $depends 4
+        var #s1
+          type-con $state 0
+        var #s2
+          type-con $state 0
+        var #dependant
+          type-con $ptr 0
+        var #this
+          type-con $ptr 0
+      attribute qid 1
+        string-attr VccPrelu.781:31
+      attribute uniqueId 1
+        string-attr 75
+      attribute bvZ3Native 1
+        string-attr False
+      attribute weight 1
+        expr-attr
+          int-num 0
+    =
+    fun $depends 4
+    var #s1
+      type-con $state 0
+    var #s2
+      type-con $state 0
+    var #dependant
+      type-con $ptr 0
+    var #this
+      type-con $ptr 0
+    or 4
+    fun $spans_the_same 4
+    var #s1
+      type-con $state 0
+    var #s2
+      type-con $state 0
+    var #this
+      type-con $ptr 0
+    fun $typ 1
+    var #this
+      type-con $ptr 0
+    and 2
+    not
+    fun $closed 2
+    var #s1
+      type-con $state 0
+    var #dependant
+      type-con $ptr 0
+    not
+    fun $closed 2
+    var #s2
+      type-con $state 0
+    var #dependant
+      type-con $ptr 0
+    and 2
+    fun $inv2 4
+    var #s1
+      type-con $state 0
+    var #s2
+      type-con $state 0
+    var #dependant
+      type-con $ptr 0
+    fun $typ 1
+    var #dependant
+      type-con $ptr 0
+    fun $nonvolatile_spans_the_same 4
+    var #s1
+      type-con $state 0
+    var #s2
+      type-con $state 0
+    var #dependant
+      type-con $ptr 0
+    fun $typ 1
+    var #dependant
+      type-con $ptr 0
+    fun $is_threadtype 1
+    fun $typ 1
+    var #dependant
+      type-con $ptr 0
+axiom 0
+    forall 4 1 4
+      var #s1
+        type-con $state 0
+      var #s2
+        type-con $state 0
+      var #p
+        type-con $ptr 0
+      var #t
+        type-con $ctype 0
+      pat 1
+        fun $spans_the_same 4
+        var #s1
+          type-con $state 0
+        var #s2
+          type-con $state 0
+        var #p
+          type-con $ptr 0
+        var #t
+          type-con $ctype 0
+      attribute qid 1
+        string-attr VccPrelu.786:38
+      attribute uniqueId 1
+        string-attr 76
+      attribute bvZ3Native 1
+        string-attr False
+      attribute weight 1
+        expr-attr
+          int-num 0
+    =
+    fun $spans_the_same 4
+    var #s1
+      type-con $state 0
+    var #s2
+      type-con $state 0
+    var #p
+      type-con $ptr 0
+    var #t
+      type-con $ctype 0
+    and 4
+    =
+    fun $read_version 2
+    var #s1
+      type-con $state 0
+    var #p
+      type-con $ptr 0
+    fun $read_version 2
+    var #s2
+      type-con $state 0
+    var #p
+      type-con $ptr 0
+    =
+    fun $owns 2
+    var #s1
+      type-con $state 0
+    var #p
+      type-con $ptr 0
+    fun $owns 2
+    var #s2
+      type-con $state 0
+    var #p
+      type-con $ptr 0
+    =
+    fun $select.tm 2
+    fun $typemap 1
+    var #s1
+      type-con $state 0
+    var #p
+      type-con $ptr 0
+    fun $select.tm 2
+    fun $typemap 1
+    var #s2
+      type-con $state 0
+    var #p
+      type-con $ptr 0
+    fun $state_spans_the_same 4
+    var #s1
+      type-con $state 0
+    var #s2
+      type-con $state 0
+    var #p
+      type-con $ptr 0
+    var #t
+      type-con $ctype 0
+axiom 0
+    forall 4 1 4
+      var #s1
+        type-con $state 0
+      var #s2
+        type-con $state 0
+      var #p
+        type-con $ptr 0
+      var #t
+        type-con $ctype 0
+      pat 1
+        fun $nonvolatile_spans_the_same 4
+        var #s1
+          type-con $state 0
+        var #s2
+          type-con $state 0
+        var #p
+          type-con $ptr 0
+        var #t
+          type-con $ctype 0
+      attribute qid 1
+        string-attr VccPrelu.795:50
+      attribute uniqueId 1
+        string-attr 77
+      attribute bvZ3Native 1
+        string-attr False
+      attribute weight 1
+        expr-attr
+          int-num 0
+    =
+    fun $nonvolatile_spans_the_same 4
+    var #s1
+      type-con $state 0
+    var #s2
+      type-con $state 0
+    var #p
+      type-con $ptr 0
+    var #t
+      type-con $ctype 0
+    and 3
+    =
+    fun $read_version 2
+    var #s1
+      type-con $state 0
+    var #p
+      type-con $ptr 0
+    fun $read_version 2
+    var #s2
+      type-con $state 0
+    var #p
+      type-con $ptr 0
+    =
+    fun $select.tm 2
+    fun $typemap 1
+    var #s1
+      type-con $state 0
+    var #p
+      type-con $ptr 0
+    fun $select.tm 2
+    fun $typemap 1
+    var #s2
+      type-con $state 0
+    var #p
+      type-con $ptr 0
+    fun $state_nonvolatile_spans_the_same 4
+    var #s1
+      type-con $state 0
+    var #s2
+      type-con $state 0
+    var #p
+      type-con $ptr 0
+    var #t
+      type-con $ctype 0
+axiom 0
+    forall 1 1 3
+      var T
+        type-con $ctype 0
+      pat 1
+        fun $is_primitive 1
+        var T
+          type-con $ctype 0
+      attribute qid 1
+        string-attr VccPrelu.813:15
+      attribute uniqueId 1
+        string-attr 79
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    fun $is_primitive 1
+    var T
+      type-con $ctype 0
+    forall 2 1 3
+      var r
+        int
+      var p
+        type-con $ptr 0
+      pat 1
+        fun $set_in 2
+        var p
+          type-con $ptr 0
+        fun $full_extent 1
+        fun $ptr 2
+        var T
+          type-con $ctype 0
+        var r
+          int
+      attribute qid 1
+        string-attr VccPrelu.815:13
+      attribute uniqueId 1
+        string-attr 78
+      attribute bvZ3Native 1
+        string-attr False
+    =
+    fun $set_in 2
+    var p
+      type-con $ptr 0
+    fun $full_extent 1
+    fun $ptr 2
+    var T
+      type-con $ctype 0
+    var r
+      int
+    =
+    var p
+      type-con $ptr 0
+    fun $ptr 2
+    var T
+      type-con $ctype 0
+    var r
+      int
+axiom 0
+    forall 1 1 3
+      var T
+        type-con $ctype 0
+      pat 1
+        fun $is_primitive 1
+        var T
+          type-con $ctype 0
+      attribute qid 1
+        string-attr VccPrelu.818:15
+      attribute uniqueId 1
+        string-attr 81
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    fun $is_primitive 1
+    var T
+      type-con $ctype 0
+    forall 3 1 3
+      var S
+        type-con $state 0
+      var r
+        int
+      var p
+        type-con $ptr 0
+      pat 1
+        fun $set_in 2
+        var p
+          type-con $ptr 0
+        fun $extent 2
+        var S
+          type-con $state 0
+        fun $ptr 2
+        var T
+          type-con $ctype 0
+        var r
+          int
+      attribute qid 1
+        string-attr VccPrelu.820:13
+      attribute uniqueId 1
+        string-attr 80
+      attribute bvZ3Native 1
+        string-attr False
+    =
+    fun $set_in 2
+    var p
+      type-con $ptr 0
+    fun $extent 2
+    var S
+      type-con $state 0
+    fun $ptr 2
+    var T
+      type-con $ctype 0
+    var r
+      int
+    =
+    var p
+      type-con $ptr 0
+    fun $ptr 2
+    var T
+      type-con $ctype 0
+    var r
+      int
+axiom 0
+    forall 1 1 3
+      var S
+        type-con $state 0
+      pat 1
+        fun $function_entry 1
+        var S
+          type-con $state 0
+      attribute qid 1
+        string-attr VccPrelu.835:15
+      attribute uniqueId 1
+        string-attr 83
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    fun $function_entry 1
+    var S
+      type-con $state 0
+    and 2
+    fun $full_stop 1
+    var S
+      type-con $state 0
+    >=
+    fun $current_timestamp 1
+    var S
+      type-con $state 0
+    int-num 0
+axiom 0
+    forall 1 1 3
+      var S
+        type-con $state 0
+      pat 1
+        fun $full_stop 1
+        var S
+          type-con $state 0
+      attribute qid 1
+        string-attr VccPrelu.838:15
+      attribute uniqueId 1
+        string-attr 84
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    fun $full_stop 1
+    var S
+      type-con $state 0
+    and 2
+    fun $good_state 1
+    var S
+      type-con $state 0
+    fun $invok_state 1
+    var S
+      type-con $state 0
+axiom 0
+    forall 1 1 3
+      var S
+        type-con $state 0
+      pat 1
+        fun $invok_state 1
+        var S
+          type-con $state 0
+      attribute qid 1
+        string-attr VccPrelu.841:15
+      attribute uniqueId 1
+        string-attr 85
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    fun $invok_state 1
+    var S
+      type-con $state 0
+    fun $good_state 1
+    var S
+      type-con $state 0
+axiom 0
+    forall 2 1 3
+      var id
+        type-con $token 0
+      var S
+        type-con $state 0
+      pat 1
+        fun $good_state_ext 2
+        var id
+          type-con $token 0
+        var S
+          type-con $state 0
+      attribute qid 1
+        string-attr VccPrelu.860:15
+      attribute uniqueId 1
+        string-attr 87
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    fun $good_state_ext 2
+    var id
+      type-con $token 0
+    var S
+      type-con $state 0
+    fun $good_state 1
+    var S
+      type-con $state 0
+axiom 0
+    forall 3 1 3
+      var S
+        type-con $state 0
+      var r
+        int
+      var t
+        type-con $ctype 0
+      pat 1
+        fun $ptr 2
+        var t
+          type-con $ctype 0
+        fun $select.mem 2
+        fun $memory 1
+        var S
+          type-con $state 0
+        fun $ptr 2
+        fun $ptr_to 1
+        var t
+          type-con $ctype 0
+        var r
+          int
+      attribute qid 1
+        string-attr VccPrelu.872:15
+      attribute uniqueId 1
+        string-attr 88
+      attribute bvZ3Native 1
+        string-attr False
+    =
+    fun $ptr 2
+    var t
+      type-con $ctype 0
+    fun $select.mem 2
+    fun $memory 1
+    var S
+      type-con $state 0
+    fun $ptr 2
+    fun $ptr_to 1
+    var t
+      type-con $ctype 0
+    var r
+      int
+    fun $read_ptr_m 3
+    var S
+      type-con $state 0
+    fun $ptr 2
+    fun $ptr_to 1
+    var t
+      type-con $ctype 0
+    var r
+      int
+    var t
+      type-con $ctype 0
+axiom 0
+    forall 2 1 4
+      var S
+        type-con $state 0
+      var p
+        type-con $ptr 0
+      pat 1
+        fun $read_version 2
+        var S
+          type-con $state 0
+        var p
+          type-con $ptr 0
+      attribute qid 1
+        string-attr VccPrelu.886:36
+      attribute uniqueId 1
+        string-attr 89
+      attribute bvZ3Native 1
+        string-attr False
+      attribute weight 1
+        expr-attr
+          int-num 0
+    =
+    fun $read_version 2
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    fun $int_to_version 1
+    fun $select.mem 2
+    fun $memory 1
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+axiom 0
+    forall 2 1 4
+      var S
+        type-con $state 0
+      var p
+        type-con $ptr 0
+      pat 1
+        fun $domain 2
+        var S
+          type-con $state 0
+        var p
+          type-con $ptr 0
+      attribute qid 1
+        string-attr VccPrelu.889:30
+      attribute uniqueId 1
+        string-attr 90
+      attribute bvZ3Native 1
+        string-attr False
+      attribute weight 1
+        expr-attr
+          int-num 0
+    =
+    fun $domain 2
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    fun $ver_domain 1
+    fun $read_version 2
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+axiom 0
+    forall 4 1 4
+      var S
+        type-con $state 0
+      var p
+        type-con $ptr 0
+      var q
+        type-con $ptr 0
+      var l
+        type-con $label 0
+      pat 1
+        fun $in_domain_lab 4
+        var S
+          type-con $state 0
+        var p
+          type-con $ptr 0
+        var q
+          type-con $ptr 0
+        var l
+          type-con $label 0
+      attribute qid 1
+        string-attr VccPrelu.899:15
+      attribute uniqueId 1
+        string-attr 91
+      attribute bvZ3Native 1
+        string-attr False
+      attribute weight 1
+        expr-attr
+          int-num 0
+    implies
+    fun $in_domain_lab 4
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    var q
+      type-con $ptr 0
+    var l
+      type-con $label 0
+    fun $inv_lab 3
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    var l
+      type-con $label 0
+axiom 0
+    forall 4 1 4
+      var S
+        type-con $state 0
+      var p
+        type-con $ptr 0
+      var q
+        type-con $ptr 0
+      var l
+        type-con $label 0
+      pat 1
+        fun $in_domain_lab 4
+        var S
+          type-con $state 0
+        var p
+          type-con $ptr 0
+        var q
+          type-con $ptr 0
+        var l
+          type-con $label 0
+      attribute qid 1
+        string-attr VccPrelu.902:15
+      attribute uniqueId 1
+        string-attr 92
+      attribute bvZ3Native 1
+        string-attr False
+      attribute weight 1
+        expr-attr
+          int-num 0
+    =
+    fun $in_domain_lab 4
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    var q
+      type-con $ptr 0
+    var l
+      type-con $label 0
+    fun $in_domain 3
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    var q
+      type-con $ptr 0
+axiom 0
+    forall 4 1 4
+      var S
+        type-con $state 0
+      var p
+        type-con $ptr 0
+      var q
+        type-con $ptr 0
+      var l
+        type-con $label 0
+      pat 1
+        fun $in_vdomain_lab 4
+        var S
+          type-con $state 0
+        var p
+          type-con $ptr 0
+        var q
+          type-con $ptr 0
+        var l
+          type-con $label 0
+      attribute qid 1
+        string-attr VccPrelu.905:15
+      attribute uniqueId 1
+        string-attr 93
+      attribute bvZ3Native 1
+        string-attr False
+      attribute weight 1
+        expr-attr
+          int-num 0
+    implies
+    fun $in_vdomain_lab 4
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    var q
+      type-con $ptr 0
+    var l
+      type-con $label 0
+    fun $inv_lab 3
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    var l
+      type-con $label 0
+axiom 0
+    forall 4 1 4
+      var S
+        type-con $state 0
+      var p
+        type-con $ptr 0
+      var q
+        type-con $ptr 0
+      var l
+        type-con $label 0
+      pat 1
+        fun $in_vdomain_lab 4
+        var S
+          type-con $state 0
+        var p
+          type-con $ptr 0
+        var q
+          type-con $ptr 0
+        var l
+          type-con $label 0
+      attribute qid 1
+        string-attr VccPrelu.908:15
+      attribute uniqueId 1
+        string-attr 94
+      attribute bvZ3Native 1
+        string-attr False
+      attribute weight 1
+        expr-attr
+          int-num 0
+    =
+    fun $in_vdomain_lab 4
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    var q
+      type-con $ptr 0
+    var l
+      type-con $label 0
+    fun $in_vdomain 3
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    var q
+      type-con $ptr 0
+axiom 0
+    forall 3 1 4
+      var S
+        type-con $state 0
+      var p
+        type-con $ptr 0
+      var q
+        type-con $ptr 0
+      pat 1
+        fun $in_domain 3
+        var S
+          type-con $state 0
+        var p
+          type-con $ptr 0
+        var q
+          type-con $ptr 0
+      attribute qid 1
+        string-attr VccPrelu.914:15
+      attribute uniqueId 1
+        string-attr 96
+      attribute bvZ3Native 1
+        string-attr False
+      attribute weight 1
+        expr-attr
+          int-num 0
+    implies
+    fun $in_domain 3
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    var q
+      type-con $ptr 0
+    and 3
+    fun $set_in 2
+    var p
+      type-con $ptr 0
+    fun $domain 2
+    var S
+      type-con $state 0
+    var q
+      type-con $ptr 0
+    fun $closed 2
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    forall 1 1 3
+      var r
+        type-con $ptr 0
+      pat 1
+        fun $set_in 2
+        var r
+          type-con $ptr 0
+        fun $owns 2
+        var S
+          type-con $state 0
+        var p
+          type-con $ptr 0
+      attribute qid 1
+        string-attr VccPrelu.918:16
+      attribute uniqueId 1
+        string-attr 95
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 2
+    not
+    fun $has_volatile_owns_set 1
+    fun $typ 1
+    var p
+      type-con $ptr 0
+    fun $set_in 2
+    var r
+      type-con $ptr 0
+    fun $owns 2
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    fun $set_in2 2
+    var r
+      type-con $ptr 0
+    fun $ver_domain 1
+    fun $read_version 2
+    var S
+      type-con $state 0
+    var q
+      type-con $ptr 0
+axiom 0
+    forall 2 1 3
+      var S
+        type-con $state 0
+      var p
+        type-con $ptr 0
+      pat 1
+        fun $in_domain 3
+        var S
+          type-con $state 0
+        var p
+          type-con $ptr 0
+        var p
+          type-con $ptr 0
+      attribute qid 1
+        string-attr VccPrelu.923:15
+      attribute uniqueId 1
+        string-attr 97
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 7
+    fun $full_stop 1
+    var S
+      type-con $state 0
+    fun $closed 2
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    =
+    fun $owner 2
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    fun $me 0
+    fun $is 2
+    var p
+      type-con $ptr 0
+    fun $typ 1
+    var p
+      type-con $ptr 0
+    fun $typed 2
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    not
+    =
+    fun $kind_of 1
+    fun $typ 1
+    var p
+      type-con $ptr 0
+    fun $kind_primitive 0
+    fun $is_non_primitive 1
+    fun $typ 1
+    var p
+      type-con $ptr 0
+    fun $in_domain 3
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    var p
+      type-con $ptr 0
+axiom 0
+    forall 3 1 4
+      var S
+        type-con $state 0
+      var p
+        type-con $ptr 0
+      var q
+        type-con $ptr 0
+      pat 1
+        fun $in_domain 3
+        var S
+          type-con $state 0
+        var q
+          type-con $ptr 0
+        var p
+          type-con $ptr 0
+      attribute qid 1
+        string-attr VccPrelu.932:15
+      attribute uniqueId 1
+        string-attr 98
+      attribute bvZ3Native 1
+        string-attr False
+      attribute weight 1
+        expr-attr
+          int-num 0
+    implies
+    and 2
+    fun $full_stop 1
+    var S
+      type-con $state 0
+    fun $set_in 2
+    var q
+      type-con $ptr 0
+    fun $domain 2
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    fun $in_domain 3
+    var S
+      type-con $state 0
+    var q
+      type-con $ptr 0
+    var p
+      type-con $ptr 0
+axiom 0
+    forall 4 1 4
+      var S
+        type-con $state 0
+      var p
+        type-con $ptr 0
+      var q
+        type-con $ptr 0
+      var r
+        type-con $ptr 0
+      pat 2
+        fun $set_in 2
+        var q
+          type-con $ptr 0
+        fun $domain 2
+        var S
+          type-con $state 0
+        var p
+          type-con $ptr 0
+        fun $in_domain 3
+        var S
+          type-con $state 0
+        var r
+          type-con $ptr 0
+        var p
+          type-con $ptr 0
+      attribute qid 1
+        string-attr VccPrelu.936:15
+      attribute uniqueId 1
+        string-attr 99
+      attribute bvZ3Native 1
+        string-attr False
+      attribute weight 1
+        expr-attr
+          int-num 0
+    implies
+    and 3
+    not
+    fun $has_volatile_owns_set 1
+    fun $typ 1
+    var q
+      type-con $ptr 0
+    fun $set_in 2
+    var q
+      type-con $ptr 0
+    fun $domain 2
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    fun $set_in0 2
+    var r
+      type-con $ptr 0
+    fun $owns 2
+    var S
+      type-con $state 0
+    var q
+      type-con $ptr 0
+    and 2
+    fun $in_domain 3
+    var S
+      type-con $state 0
+    var r
+      type-con $ptr 0
+    var p
+      type-con $ptr 0
+    fun $set_in0 2
+    var r
+      type-con $ptr 0
+    fun $owns 2
+    var S
+      type-con $state 0
+    var q
+      type-con $ptr 0
+axiom 0
+    forall 4 1 4
+      var S
+        type-con $state 0
+      var p
+        type-con $ptr 0
+      var q
+        type-con $ptr 0
+      var r
+        type-con $ptr 0
+      pat 2
+        fun $set_in 2
+        var q
+          type-con $ptr 0
+        fun $domain 2
+        var S
+          type-con $state 0
+        var p
+          type-con $ptr 0
+        fun $in_vdomain 3
+        var S
+          type-con $state 0
+        var r
+          type-con $ptr 0
+        var p
+          type-con $ptr 0
+      attribute qid 1
+        string-attr VccPrelu.941:15
+      attribute uniqueId 1
+        string-attr 101
+      attribute bvZ3Native 1
+        string-attr False
+      attribute weight 1
+        expr-attr
+          int-num 0
+    implies
+    and 3
+    fun $has_volatile_owns_set 1
+    fun $typ 1
+    var q
+      type-con $ptr 0
+    fun $set_in 2
+    var q
+      type-con $ptr 0
+    fun $domain 2
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    forall 1 0 3
+      var S1
+        type-con $state 0
+      attribute qid 1
+        string-attr VccPrelu.945:11
+      attribute uniqueId 1
+        string-attr 100
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 3
+    fun $inv2 4
+    var S1
+      type-con $state 0
+    var S1
+      type-con $state 0
+    var q
+      type-con $ptr 0
+    fun $typ 1
+    var q
+      type-con $ptr 0
+    =
+    fun $read_version 2
+    var S1
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    fun $read_version 2
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    =
+    fun $domain 2
+    var S1
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    fun $domain 2
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    fun $set_in0 2
+    var r
+      type-con $ptr 0
+    fun $owns 2
+    var S1
+      type-con $state 0
+    var q
+      type-con $ptr 0
+    and 2
+    fun $in_vdomain 3
+    var S
+      type-con $state 0
+    var r
+      type-con $ptr 0
+    var p
+      type-con $ptr 0
+    fun $set_in0 2
+    var r
+      type-con $ptr 0
+    fun $owns 2
+    var S
+      type-con $state 0
+    var q
+      type-con $ptr 0
+axiom 0
+    forall 3 1 4
+      var S
+        type-con $state 0
+      var p
+        type-con $ptr 0
+      var q
+        type-con $ptr 0
+      pat 1
+        fun $in_vdomain 3
+        var S
+          type-con $state 0
+        var p
+          type-con $ptr 0
+        var q
+          type-con $ptr 0
+      attribute qid 1
+        string-attr VccPrelu.952:15
+      attribute uniqueId 1
+        string-attr 102
+      attribute bvZ3Native 1
+        string-attr False
+      attribute weight 1
+        expr-attr
+          int-num 0
+    implies
+    fun $in_vdomain 3
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    var q
+      type-con $ptr 0
+    fun $in_domain 3
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    var q
+      type-con $ptr 0
+axiom 0
+    forall 4 1 3
+      var S
+        type-con $state 0
+      var p
+        type-con $ptr 0
+      var d
+        type-con $ptr 0
+      var f
+        type-con $field 0
+      pat 3
+        fun $set_in 2
+        var p
+          type-con $ptr 0
+        fun $domain 2
+        var S
+          type-con $state 0
+        var d
+          type-con $ptr 0
+        fun $is_primitive_non_volatile_field 1
+        var f
+          type-con $field 0
+        fun $select.mem 2
+        fun $memory 1
+        var S
+          type-con $state 0
+        fun $dot 2
+        var p
+          type-con $ptr 0
+        var f
+          type-con $field 0
+      attribute qid 1
+        string-attr VccPrelu.957:15
+      attribute uniqueId 1
+        string-attr 103
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 2
+    fun $set_in 2
+    var p
+      type-con $ptr 0
+    fun $domain 2
+    var S
+      type-con $state 0
+    var d
+      type-con $ptr 0
+    fun $is_primitive_non_volatile_field 1
+    var f
+      type-con $field 0
+    =
+    fun $select.mem 2
+    fun $memory 1
+    var S
+      type-con $state 0
+    fun $dot 2
+    var p
+      type-con $ptr 0
+    var f
+      type-con $field 0
+    fun $fetch_from_domain 2
+    fun $read_version 2
+    var S
+      type-con $state 0
+    var d
+      type-con $ptr 0
+    fun $dot 2
+    var p
+      type-con $ptr 0
+    var f
+      type-con $field 0
+axiom 0
+    forall 3 2 3
+      var S
+        type-con $state 0
+      var p
+        type-con $ptr 0
+      var d
+        type-con $ptr 0
+      pat 3
+        fun $full_stop 1
+        var S
+          type-con $state 0
+        fun $set_in 2
+        var p
+          type-con $ptr 0
+        fun $domain 2
+        var S
+          type-con $state 0
+        var d
+          type-con $ptr 0
+        fun $select.sm 2
+        fun $statusmap 1
+        var S
+          type-con $state 0
+        var p
+          type-con $ptr 0
+      pat 3
+        fun $full_stop 1
+        var S
+          type-con $state 0
+        fun $set_in 2
+        var p
+          type-con $ptr 0
+        fun $domain 2
+        var S
+          type-con $state 0
+        var d
+          type-con $ptr 0
+        fun $select.tm 2
+        fun $typemap 1
+        var S
+          type-con $state 0
+        var p
+          type-con $ptr 0
+      attribute qid 1
+        string-attr VccPrelu.962:15
+      attribute uniqueId 1
+        string-attr 104
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 2
+    fun $full_stop 1
+    var S
+      type-con $state 0
+    fun $set_in 2
+    var p
+      type-con $ptr 0
+    fun $domain 2
+    var S
+      type-con $state 0
+    var d
+      type-con $ptr 0
+    and 2
+    fun $typed 2
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    not
+    fun $ts_is_volatile 1
+    fun $select.tm 2
+    fun $typemap 1
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+axiom 0
+    forall 4 2 3
+      var S
+        type-con $state 0
+      var p
+        type-con $ptr 0
+      var d
+        type-con $ptr 0
+      var f
+        type-con $field 0
+      pat 3
+        fun $set_in 2
+        var p
+          type-con $ptr 0
+        fun $domain 2
+        var S
+          type-con $state 0
+        var d
+          type-con $ptr 0
+        fun $is_primitive_non_volatile_field 1
+        var f
+          type-con $field 0
+        fun $owner 2
+        var S
+          type-con $state 0
+        fun $dot 2
+        var p
+          type-con $ptr 0
+        var f
+          type-con $field 0
+      pat 3
+        fun $set_in 2
+        var p
+          type-con $ptr 0
+        fun $domain 2
+        var S
+          type-con $state 0
+        var d
+          type-con $ptr 0
+        fun $is_primitive_non_volatile_field 1
+        var f
+          type-con $field 0
+        fun $select.tm 2
+        fun $typemap 1
+        var S
+          type-con $state 0
+        fun $dot 2
+        var p
+          type-con $ptr 0
+        var f
+          type-con $field 0
+      attribute qid 1
+        string-attr VccPrelu.968:15
+      attribute uniqueId 1
+        string-attr 105
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 3
+    fun $full_stop 1
+    var S
+      type-con $state 0
+    fun $set_in 2
+    var p
+      type-con $ptr 0
+    fun $domain 2
+    var S
+      type-con $state 0
+    var d
+      type-con $ptr 0
+    fun $is_primitive_non_volatile_field 1
+    var f
+      type-con $field 0
+    and 2
+    fun $typed 2
+    var S
+      type-con $state 0
+    fun $dot 2
+    var p
+      type-con $ptr 0
+    var f
+      type-con $field 0
+    not
+    fun $ts_is_volatile 1
+    fun $select.tm 2
+    fun $typemap 1
+    var S
+      type-con $state 0
+    fun $dot 2
+    var p
+      type-con $ptr 0
+    var f
+      type-con $field 0
+axiom 0
+    forall 7 1 3
+      var S
+        type-con $state 0
+      var p
+        type-con $ptr 0
+      var d
+        type-con $ptr 0
+      var f
+        type-con $field 0
+      var sz
+        int
+      var i
+        int
+      var t
+        type-con $ctype 0
+      pat 3
+        fun $set_in 2
+        var p
+          type-con $ptr 0
+        fun $domain 2
+        var S
+          type-con $state 0
+        var d
+          type-con $ptr 0
+        fun $is_primitive_embedded_array 2
+        var f
+          type-con $field 0
+        var sz
+          int
+        fun $select.mem 2
+        fun $memory 1
+        var S
+          type-con $state 0
+        fun $idx 3
+        fun $dot 2
+        var p
+          type-con $ptr 0
+        var f
+          type-con $field 0
+        var i
+          int
+        var t
+          type-con $ctype 0
+      attribute qid 1
+        string-attr VccPrelu.974:15
+      attribute uniqueId 1
+        string-attr 106
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 5
+    fun $full_stop 1
+    var S
+      type-con $state 0
+    fun $set_in 2
+    var p
+      type-con $ptr 0
+    fun $domain 2
+    var S
+      type-con $state 0
+    var d
+      type-con $ptr 0
+    fun $is_primitive_embedded_array 2
+    var f
+      type-con $field 0
+    var sz
+      int
+    <=
+    int-num 0
+    var i
+      int
+    <
+    var i
+      int
+    var sz
+      int
+    =
+    fun $select.mem 2
+    fun $memory 1
+    var S
+      type-con $state 0
+    fun $idx 3
+    fun $dot 2
+    var p
+      type-con $ptr 0
+    var f
+      type-con $field 0
+    var i
+      int
+    var t
+      type-con $ctype 0
+    fun $fetch_from_domain 2
+    fun $read_version 2
+    var S
+      type-con $state 0
+    var d
+      type-con $ptr 0
+    fun $idx 3
+    fun $dot 2
+    var p
+      type-con $ptr 0
+    var f
+      type-con $field 0
+    var i
+      int
+    var t
+      type-con $ctype 0
+axiom 0
+    forall 7 2 3
+      var S
+        type-con $state 0
+      var p
+        type-con $ptr 0
+      var d
+        type-con $ptr 0
+      var f
+        type-con $field 0
+      var sz
+        int
+      var i
+        int
+      var t
+        type-con $ctype 0
+      pat 3
+        fun $set_in 2
+        var p
+          type-con $ptr 0
+        fun $domain 2
+        var S
+          type-con $state 0
+        var d
+          type-con $ptr 0
+        fun $is_primitive_embedded_array 2
+        var f
+          type-con $field 0
+        var sz
+          int
+        fun $select.tm 2
+        fun $typemap 1
+        var S
+          type-con $state 0
+        fun $idx 3
+        fun $dot 2
+        var p
+          type-con $ptr 0
+        var f
+          type-con $field 0
+        var i
+          int
+        var t
+          type-con $ctype 0
+      pat 3
+        fun $set_in 2
+        var p
+          type-con $ptr 0
+        fun $domain 2
+        var S
+          type-con $state 0
+        var d
+          type-con $ptr 0
+        fun $is_primitive_embedded_array 2
+        var f
+          type-con $field 0
+        var sz
+          int
+        fun $owner 2
+        var S
+          type-con $state 0
+        fun $idx 3
+        fun $dot 2
+        var p
+          type-con $ptr 0
+        var f
+          type-con $field 0
+        var i
+          int
+        var t
+          type-con $ctype 0
+      attribute qid 1
+        string-attr VccPrelu.979:15
+      attribute uniqueId 1
+        string-attr 107
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 5
+    fun $full_stop 1
+    var S
+      type-con $state 0
+    fun $set_in 2
+    var p
+      type-con $ptr 0
+    fun $domain 2
+    var S
+      type-con $state 0
+    var d
+      type-con $ptr 0
+    fun $is_primitive_embedded_array 2
+    var f
+      type-con $field 0
+    var sz
+      int
+    <=
+    int-num 0
+    var i
+      int
+    <
+    var i
+      int
+    var sz
+      int
+    and 2
+    fun $typed 2
+    var S
+      type-con $state 0
+    fun $idx 3
+    fun $dot 2
+    var p
+      type-con $ptr 0
+    var f
+      type-con $field 0
+    var i
+      int
+    var t
+      type-con $ctype 0
+    not
+    fun $ts_is_volatile 1
+    fun $select.tm 2
+    fun $typemap 1
+    var S
+      type-con $state 0
+    fun $idx 3
+    fun $dot 2
+    var p
+      type-con $ptr 0
+    var f
+      type-con $field 0
+    var i
+      int
+    var t
+      type-con $ctype 0
+axiom 0
+    forall 6 2 3
+      var S
+        type-con $state 0
+      var r
+        int
+      var d
+        type-con $ptr 0
+      var sz
+        int
+      var i
+        int
+      var t
+        type-con $ctype 0
+      pat 3
+        fun $set_in 2
+        fun $ptr 2
+        fun $array 2
+        var t
+          type-con $ctype 0
+        var sz
+          int
+        var r
+          int
+        fun $domain 2
+        var S
+          type-con $state 0
+        var d
+          type-con $ptr 0
+        fun $select.tm 2
+        fun $typemap 1
+        var S
+          type-con $state 0
+        fun $idx 3
+        fun $ptr 2
+        var t
+          type-con $ctype 0
+        var r
+          int
+        var i
+          int
+        var t
+          type-con $ctype 0
+        fun $is_primitive 1
+        var t
+          type-con $ctype 0
+      pat 3
+        fun $set_in 2
+        fun $ptr 2
+        fun $array 2
+        var t
+          type-con $ctype 0
+        var sz
+          int
+        var r
+          int
+        fun $domain 2
+        var S
+          type-con $state 0
+        var d
+          type-con $ptr 0
+        fun $owner 2
+        var S
+          type-con $state 0
+        fun $idx 3
+        fun $ptr 2
+        var t
+          type-con $ctype 0
+        var r
+          int
+        var i
+          int
+        var t
+          type-con $ctype 0
+        fun $is_primitive 1
+        var t
+          type-con $ctype 0
+      attribute qid 1
+        string-attr VccPrelu.985:15
+      attribute uniqueId 1
+        string-attr 108
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 5
+    fun $full_stop 1
+    var S
+      type-con $state 0
+    fun $is_primitive 1
+    var t
+      type-con $ctype 0
+    fun $set_in 2
+    fun $ptr 2
+    fun $array 2
+    var t
+      type-con $ctype 0
+    var sz
+      int
+    var r
+      int
+    fun $domain 2
+    var S
+      type-con $state 0
+    var d
+      type-con $ptr 0
+    <=
+    int-num 0
+    var i
+      int
+    <
+    var i
+      int
+    var sz
+      int
+    and 2
+    fun $typed 2
+    var S
+      type-con $state 0
+    fun $idx 3
+    fun $ptr 2
+    var t
+      type-con $ctype 0
+    var r
+      int
+    var i
+      int
+    var t
+      type-con $ctype 0
+    not
+    fun $ts_is_volatile 1
+    fun $select.tm 2
+    fun $typemap 1
+    var S
+      type-con $state 0
+    fun $idx 3
+    fun $ptr 2
+    var t
+      type-con $ctype 0
+    var r
+      int
+    var i
+      int
+    var t
+      type-con $ctype 0
+axiom 0
+    forall 6 1 3
+      var S
+        type-con $state 0
+      var r
+        int
+      var d
+        type-con $ptr 0
+      var sz
+        int
+      var i
+        int
+      var t
+        type-con $ctype 0
+      pat 3
+        fun $set_in 2
+        fun $ptr 2
+        fun $array 2
+        var t
+          type-con $ctype 0
+        var sz
+          int
+        var r
+          int
+        fun $domain 2
+        var S
+          type-con $state 0
+        var d
+          type-con $ptr 0
+        fun $select.mem 2
+        fun $memory 1
+        var S
+          type-con $state 0
+        fun $idx 3
+        fun $ptr 2
+        var t
+          type-con $ctype 0
+        var r
+          int
+        var i
+          int
+        var t
+          type-con $ctype 0
+        fun $is_primitive 1
+        var t
+          type-con $ctype 0
+      attribute qid 1
+        string-attr VccPrelu.994:15
+      attribute uniqueId 1
+        string-attr 109
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 5
+    fun $full_stop 1
+    var S
+      type-con $state 0
+    fun $is_primitive 1
+    var t
+      type-con $ctype 0
+    fun $set_in 2
+    fun $ptr 2
+    fun $array 2
+    var t
+      type-con $ctype 0
+    var sz
+      int
+    var r
+      int
+    fun $domain 2
+    var S
+      type-con $state 0
+    var d
+      type-con $ptr 0
+    <=
+    int-num 0
+    var i
+      int
+    <
+    var i
+      int
+    var sz
+      int
+    =
+    fun $select.mem 2
+    fun $memory 1
+    var S
+      type-con $state 0
+    fun $idx 3
+    fun $ptr 2
+    var t
+      type-con $ctype 0
+    var r
+      int
+    var i
+      int
+    var t
+      type-con $ctype 0
+    fun $fetch_from_domain 2
+    fun $read_version 2
+    var S
+      type-con $state 0
+    var d
+      type-con $ptr 0
+    fun $idx 3
+    fun $ptr 2
+    var t
+      type-con $ctype 0
+    var r
+      int
+    var i
+      int
+    var t
+      type-con $ctype 0
+axiom 0
+    forall 6 1 3
+      var S
+        type-con $state 0
+      var p
+        type-con $ptr 0
+      var f
+        type-con $field 0
+      var sz
+        int
+      var i
+        int
+      var t
+        type-con $ctype 0
+      pat 2
+        fun $is_primitive_embedded_volatile_array 3
+        var f
+          type-con $field 0
+        var sz
+          int
+        var t
+          type-con $ctype 0
+        fun $ts_is_volatile 1
+        fun $select.tm 2
+        fun $typemap 1
+        var S
+          type-con $state 0
+        fun $idx 3
+        fun $dot 2
+        var p
+          type-con $ptr 0
+        var f
+          type-con $field 0
+        var i
+          int
+        var t
+          type-con $ctype 0
+      attribute qid 1
+        string-attr VccPrelu.1002:15
+      attribute uniqueId 1
+        string-attr 110
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 4
+    fun $good_state 1
+    var S
+      type-con $state 0
+    fun $is_primitive_embedded_volatile_array 3
+    var f
+      type-con $field 0
+    var sz
+      int
+    var t
+      type-con $ctype 0
+    <=
+    int-num 0
+    var i
+      int
+    <
+    var i
+      int
+    var sz
+      int
+    fun $ts_is_volatile 1
+    fun $select.tm 2
+    fun $typemap 1
+    var S
+      type-con $state 0
+    fun $idx 3
+    fun $dot 2
+    var p
+      type-con $ptr 0
+    var f
+      type-con $field 0
+    var i
+      int
+    var t
+      type-con $ctype 0
+axiom 0
+    forall 4 1 4
+      var p
+        type-con $ptr 0
+      var S1
+        type-con $state 0
+      var S2
+        type-con $state 0
+      var q
+        type-con $ptr 0
+      pat 2
+        fun $set_in 2
+        var q
+          type-con $ptr 0
+        fun $domain 2
+        var S1
+          type-con $state 0
+        var p
+          type-con $ptr 0
+        fun $call_transition 2
+        var S1
+          type-con $state 0
+        var S2
+          type-con $state 0
+      attribute qid 1
+        string-attr VccPrelu.1013:15
+      attribute uniqueId 1
+        string-attr 111
+      attribute bvZ3Native 1
+        string-attr False
+      attribute weight 1
+        expr-attr
+          int-num 0
+    fun $instantiate_bool 1
+    fun $set_in 2
+    var q
+      type-con $ptr 0
+    fun $domain 2
+    var S2
+      type-con $state 0
+    var p
+      type-con $ptr 0
+axiom 0
+    forall 4 1 4
+      var p
+        type-con $ptr 0
+      var S1
+        type-con $state 0
+      var S2
+        type-con $state 0
+      var q
+        type-con $ptr 0
+      pat 2
+        fun $set_in 2
+        var q
+          type-con $ptr 0
+        fun $ver_domain 1
+        fun $read_version 2
+        var S1
+          type-con $state 0
+        var p
+          type-con $ptr 0
+        fun $call_transition 2
+        var S1
+          type-con $state 0
+        var S2
+          type-con $state 0
+      attribute qid 1
+        string-attr VccPrelu.1017:15
+      attribute uniqueId 1
+        string-attr 112
+      attribute bvZ3Native 1
+        string-attr False
+      attribute weight 1
+        expr-attr
+          int-num 0
+    fun $instantiate_bool 1
+    fun $set_in 2
+    var q
+      type-con $ptr 0
+    fun $ver_domain 1
+    fun $read_version 2
+    var S2
+      type-con $state 0
+    var p
+      type-con $ptr 0
+axiom 0
+    forall 2 1 3
+      var p
+        type-con $ptr 0
+      var c
+        type-con $ptr 0
+      pat 1
+        fun $in_claim_domain 2
+        var p
+          type-con $ptr 0
+        var c
+          type-con $ptr 0
+      attribute qid 1
+        string-attr VccPrelu.1022:15
+      attribute uniqueId 1
+        string-attr 114
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    forall 1 1 3
+      var s
+        type-con $state 0
+      pat 1
+        fun $dont_instantiate_state 1
+        var s
+          type-con $state 0
+      attribute qid 1
+        string-attr VccPrelu.1023:11
+      attribute uniqueId 1
+        string-attr 113
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    fun $valid_claim 2
+    var s
+      type-con $state 0
+    var c
+      type-con $ptr 0
+    fun $closed 2
+    var s
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    fun $in_claim_domain 2
+    var p
+      type-con $ptr 0
+    var c
+      type-con $ptr 0
+axiom 0
+    forall 4 1 4
+      var S
+        type-con $state 0
+      var c
+        type-con $ptr 0
+      var obj
+        type-con $ptr 0
+      var ptr
+        type-con $ptr 0
+      pat 1
+        fun $by_claim 4
+        var S
+          type-con $state 0
+        var c
+          type-con $ptr 0
+        var obj
+          type-con $ptr 0
+        var ptr
+          type-con $ptr 0
+      attribute qid 1
+        string-attr VccPrelu.1026:32
+      attribute uniqueId 1
+        string-attr 115
+      attribute bvZ3Native 1
+        string-attr False
+      attribute weight 1
+        expr-attr
+          int-num 0
+    =
+    fun $by_claim 4
+    var S
+      type-con $state 0
+    var c
+      type-con $ptr 0
+    var obj
+      type-con $ptr 0
+    var ptr
+      type-con $ptr 0
+    var ptr
+      type-con $ptr 0
+axiom 0
+    forall 4 2 3
+      var S
+        type-con $state 0
+      var p
+        type-con $ptr 0
+      var c
+        type-con $ptr 0
+      var f
+        type-con $field 0
+      pat 2
+        fun $in_claim_domain 2
+        var p
+          type-con $ptr 0
+        var c
+          type-con $ptr 0
+        fun $select.mem 2
+        fun $memory 1
+        var S
+          type-con $state 0
+        fun $dot 2
+        var p
+          type-con $ptr 0
+        var f
+          type-con $field 0
+      pat 1
+        fun $by_claim 4
+        var S
+          type-con $state 0
+        var c
+          type-con $ptr 0
+        var p
+          type-con $ptr 0
+        fun $dot 2
+        var p
+          type-con $ptr 0
+        var f
+          type-con $field 0
+      attribute qid 1
+        string-attr VccPrelu.1031:15
+      attribute uniqueId 1
+        string-attr 116
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 4
+    fun $good_state 1
+    var S
+      type-con $state 0
+    fun $closed 2
+    var S
+      type-con $state 0
+    var c
+      type-con $ptr 0
+    fun $in_claim_domain 2
+    var p
+      type-con $ptr 0
+    var c
+      type-con $ptr 0
+    fun $is_primitive_non_volatile_field 1
+    var f
+      type-con $field 0
+    and 2
+    fun $in_claim_domain 2
+    var p
+      type-con $ptr 0
+    var c
+      type-con $ptr 0
+    =
+    fun $select.mem 2
+    fun $memory 1
+    var S
+      type-con $state 0
+    fun $dot 2
+    var p
+      type-con $ptr 0
+    var f
+      type-con $field 0
+    fun $fetch_from_domain 2
+    fun $claim_version 1
+    var c
+      type-con $ptr 0
+    fun $dot 2
+    var p
+      type-con $ptr 0
+    var f
+      type-con $field 0
+axiom 0
+    forall 7 2 3
+      var S
+        type-con $state 0
+      var p
+        type-con $ptr 0
+      var c
+        type-con $ptr 0
+      var f
+        type-con $field 0
+      var i
+        int
+      var sz
+        int
+      var t
+        type-con $ctype 0
+      pat 4
+        fun $valid_claim 2
+        var S
+          type-con $state 0
+        var c
+          type-con $ptr 0
+        fun $in_claim_domain 2
+        var p
+          type-con $ptr 0
+        var c
+          type-con $ptr 0
+        fun $select.mem 2
+        fun $memory 1
+        var S
+          type-con $state 0
+        fun $idx 3
+        fun $dot 2
+        var p
+          type-con $ptr 0
+        var f
+          type-con $field 0
+        var i
+          int
+        var t
+          type-con $ctype 0
+        fun $is_primitive_embedded_array 2
+        var f
+          type-con $field 0
+        var sz
+          int
+      pat 2
+        fun $by_claim 4
+        var S
+          type-con $state 0
+        var c
+          type-con $ptr 0
+        var p
+          type-con $ptr 0
+        fun $idx 3
+        fun $dot 2
+        var p
+          type-con $ptr 0
+        var f
+          type-con $field 0
+        var i
+          int
+        var t
+          type-con $ctype 0
+        fun $is_primitive_embedded_array 2
+        var f
+          type-con $field 0
+        var sz
+          int
+      attribute qid 1
+        string-attr VccPrelu.1040:15
+      attribute uniqueId 1
+        string-attr 117
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 6
+    fun $good_state 1
+    var S
+      type-con $state 0
+    fun $closed 2
+    var S
+      type-con $state 0
+    var c
+      type-con $ptr 0
+    fun $in_claim_domain 2
+    var p
+      type-con $ptr 0
+    var c
+      type-con $ptr 0
+    fun $is_primitive_embedded_array 2
+    var f
+      type-con $field 0
+    var sz
+      int
+    <=
+    int-num 0
+    var i
+      int
+    <
+    var i
+      int
+    var sz
+      int
+    =
+    fun $select.mem 2
+    fun $memory 1
+    var S
+      type-con $state 0
+    fun $idx 3
+    fun $dot 2
+    var p
+      type-con $ptr 0
+    var f
+      type-con $field 0
+    var i
+      int
+    var t
+      type-con $ctype 0
+    fun $fetch_from_domain 2
+    fun $claim_version 1
+    var c
+      type-con $ptr 0
+    fun $idx 3
+    fun $dot 2
+    var p
+      type-con $ptr 0
+    var f
+      type-con $field 0
+    var i
+      int
+    var t
+      type-con $ctype 0
+axiom 0
+    forall 2 1 4
+      var S
+        type-con $state 0
+      var p
+        type-con $ptr 0
+      pat 1
+        fun $read_vol_version 2
+        var S
+          type-con $state 0
+        var p
+          type-con $ptr 0
+      attribute qid 1
+        string-attr VccPrelu.1067:40
+      attribute uniqueId 1
+        string-attr 119
+      attribute bvZ3Native 1
+        string-attr False
+      attribute weight 1
+        expr-attr
+          int-num 0
+    =
+    fun $read_vol_version 2
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    fun $int_to_vol_version 1
+    fun $select.mem 2
+    fun $memory 1
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+axiom 0
+    forall 5 1 3
+      var S
+        type-con $state 0
+      var r
+        int
+      var t
+        type-con $ctype 0
+      var approver
+        type-con $field 0
+      var subject
+        type-con $field 0
+      pat 2
+        fun $is_approved_by 3
+        var t
+          type-con $ctype 0
+        var approver
+          type-con $field 0
+        var subject
+          type-con $field 0
+        fun $select.mem 2
+        fun $memory 1
+        var S
+          type-con $state 0
+        fun $dot 2
+        fun $ptr 2
+        var t
+          type-con $ctype 0
+        var r
+          int
+        var subject
+          type-con $field 0
+      attribute qid 1
+        string-attr VccPrelu.1078:15
+      attribute uniqueId 1
+        string-attr 120
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 4
+    fun $full_stop 1
+    var S
+      type-con $state 0
+    fun $is_approved_by 3
+    var t
+      type-con $ctype 0
+    var approver
+      type-con $field 0
+    var subject
+      type-con $field 0
+    fun $closed 2
+    var S
+      type-con $state 0
+    fun $ptr 2
+    var t
+      type-con $ctype 0
+    var r
+      int
+    or 2
+    =
+    fun $int_to_ptr 1
+    fun $select.mem 2
+    fun $memory 1
+    var S
+      type-con $state 0
+    fun $dot 2
+    fun $ptr 2
+    var t
+      type-con $ctype 0
+    var r
+      int
+    var approver
+      type-con $field 0
+    fun $me 0
+    =
+    fun $int_to_ptr 1
+    fun $fetch_from_vv 2
+    fun $read_vol_version 2
+    var S
+      type-con $state 0
+    fun $ptr 2
+    var t
+      type-con $ctype 0
+    var r
+      int
+    fun $dot 2
+    fun $ptr 2
+    var t
+      type-con $ctype 0
+    var r
+      int
+    var approver
+      type-con $field 0
+    fun $me 0
+    =
+    fun $select.mem 2
+    fun $memory 1
+    var S
+      type-con $state 0
+    fun $dot 2
+    fun $ptr 2
+    var t
+      type-con $ctype 0
+    var r
+      int
+    var subject
+      type-con $field 0
+    fun $fetch_from_vv 2
+    fun $read_vol_version 2
+    var S
+      type-con $state 0
+    fun $ptr 2
+    var t
+      type-con $ctype 0
+    var r
+      int
+    fun $dot 2
+    fun $ptr 2
+    var t
+      type-con $ctype 0
+    var r
+      int
+    var subject
+      type-con $field 0
+axiom 0
+    forall 4 1 3
+      var S
+        type-con $state 0
+      var r
+        int
+      var t
+        type-con $ctype 0
+      var subject
+        type-con $field 0
+      pat 2
+        fun $is_owner_approved 2
+        var t
+          type-con $ctype 0
+        var subject
+          type-con $field 0
+        fun $select.mem 2
+        fun $memory 1
+        var S
+          type-con $state 0
+        fun $dot 2
+        fun $ptr 2
+        var t
+          type-con $ctype 0
+        var r
+          int
+        var subject
+          type-con $field 0
+      attribute qid 1
+        string-attr VccPrelu.1103:15
+      attribute uniqueId 1
+        string-attr 121
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 4
+    fun $full_stop 1
+    var S
+      type-con $state 0
+    fun $closed 2
+    var S
+      type-con $state 0
+    fun $ptr 2
+    var t
+      type-con $ctype 0
+    var r
+      int
+    fun $is_owner_approved 2
+    var t
+      type-con $ctype 0
+    var subject
+      type-con $field 0
+    =
+    fun $owner 2
+    var S
+      type-con $state 0
+    fun $ptr 2
+    var t
+      type-con $ctype 0
+    var r
+      int
+    fun $me 0
+    =
+    fun $select.mem 2
+    fun $memory 1
+    var S
+      type-con $state 0
+    fun $dot 2
+    fun $ptr 2
+    var t
+      type-con $ctype 0
+    var r
+      int
+    var subject
+      type-con $field 0
+    fun $fetch_from_vv 2
+    fun $read_vol_version 2
+    var S
+      type-con $state 0
+    fun $ptr 2
+    var t
+      type-con $ctype 0
+    var r
+      int
+    fun $dot 2
+    fun $ptr 2
+    var t
+      type-con $ctype 0
+    var r
+      int
+    var subject
+      type-con $field 0
+axiom 0
+    forall 5 1 3
+      var S1
+        type-con $state 0
+      var S2
+        type-con $state 0
+      var r
+        int
+      var t
+        type-con $ctype 0
+      var subject
+        type-con $field 0
+      pat 3
+        fun $is_owner_approved 2
+        var t
+          type-con $ctype 0
+        var subject
+          type-con $field 0
+        fun $post_unwrap 2
+        var S1
+          type-con $state 0
+        var S2
+          type-con $state 0
+        fun $select.mem 2
+        fun $memory 1
+        var S1
+          type-con $state 0
+        fun $dot 2
+        fun $ptr 2
+        var t
+          type-con $ctype 0
+        var r
+          int
+        var subject
+          type-con $field 0
+      attribute qid 1
+        string-attr VccPrelu.1111:15
+      attribute uniqueId 1
+        string-attr 122
+      attribute bvZ3Native 1
+        string-attr False
+    fun $instantiate_int 1
+    fun $select.mem 2
+    fun $memory 1
+    var S2
+      type-con $state 0
+    fun $dot 2
+    fun $ptr 2
+    var t
+      type-con $ctype 0
+    var r
+      int
+    var subject
+      type-con $field 0
+axiom 0
+    forall 3 1 3
+      var S
+        type-con $state 0
+      var p
+        type-con $ptr 0
+      var q
+        type-con $ptr 0
+      pat 2
+        fun $set_in 2
+        var p
+          type-con $ptr 0
+        fun $owns 2
+        var S
+          type-con $state 0
+        var q
+          type-con $ptr 0
+        fun $is_non_primitive 1
+        fun $typ 1
+        var p
+          type-con $ptr 0
+      attribute qid 1
+        string-attr VccPrelu.1133:15
+      attribute uniqueId 1
+        string-attr 124
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 3
+    fun $good_state 1
+    var S
+      type-con $state 0
+    fun $closed 2
+    var S
+      type-con $state 0
+    var q
+      type-con $ptr 0
+    fun $is_non_primitive 1
+    fun $typ 1
+    var p
+      type-con $ptr 0
+    =
+    fun $set_in 2
+    var p
+      type-con $ptr 0
+    fun $owns 2
+    var S
+      type-con $state 0
+    var q
+      type-con $ptr 0
+    =
+    fun $owner 2
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    var q
+      type-con $ptr 0
+axiom 0
+    forall 4 1 3
+      var #s1
+        type-con $state 0
+      var #s2
+        type-con $state 0
+      var #p
+        type-con $ptr 0
+      var #t
+        type-con $ctype 0
+      pat 2
+        fun $is_arraytype 1
+        var #t
+          type-con $ctype 0
+        fun $inv2 4
+        var #s1
+          type-con $state 0
+        var #s2
+          type-con $state 0
+        var #p
+          type-con $ptr 0
+        var #t
+          type-con $ctype 0
+      attribute qid 1
+        string-attr VccPrelu.1140:15
+      attribute uniqueId 1
+        string-attr 125
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 2
+    fun $is_arraytype 1
+    var #t
+      type-con $ctype 0
+    =
+    fun $typ 1
+    var #p
+      type-con $ptr 0
+    var #t
+      type-con $ctype 0
+    and 2
+    =
+    fun $inv2 4
+    var #s1
+      type-con $state 0
+    var #s2
+      type-con $state 0
+    var #p
+      type-con $ptr 0
+    var #t
+      type-con $ctype 0
+    fun $typed 2
+    var #s2
+      type-con $state 0
+    var #p
+      type-con $ptr 0
+    fun $sequential 4
+    var #s1
+      type-con $state 0
+    var #s2
+      type-con $state 0
+    var #p
+      type-con $ptr 0
+    var #t
+      type-con $ctype 0
+axiom 0
+    forall 3 1 3
+      var S
+        type-con $state 0
+      var #r
+        int
+      var #t
+        type-con $ctype 0
+      pat 2
+        fun $owns 2
+        var S
+          type-con $state 0
+        fun $ptr 2
+        var #t
+          type-con $ctype 0
+        var #r
+          int
+        fun $is_arraytype 1
+        var #t
+          type-con $ctype 0
+      attribute qid 1
+        string-attr VccPrelu.1145:15
+      attribute uniqueId 1
+        string-attr 126
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    fun $good_state 1
+    var S
+      type-con $state 0
+    implies
+    fun $is_arraytype 1
+    var #t
+      type-con $ctype 0
+    =
+    fun $owns 2
+    var S
+      type-con $state 0
+    fun $ptr 2
+    var #t
+      type-con $ctype 0
+    var #r
+      int
+    fun $set_empty 0
+axiom 0
+    forall 3 1 3
+      var S
+        type-con $state 0
+      var #p
+        type-con $ptr 0
+      var #t
+        type-con $ctype 0
+      pat 1
+        fun $inv2 4
+        var S
+          type-con $state 0
+        var S
+          type-con $state 0
+        var #p
+          type-con $ptr 0
+        var #t
+          type-con $ctype 0
+      attribute qid 1
+        string-attr VccPrelu.1149:15
+      attribute uniqueId 1
+        string-attr 127
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 2
+    fun $invok_state 1
+    var S
+      type-con $state 0
+    fun $closed 2
+    var S
+      type-con $state 0
+    var #p
+      type-con $ptr 0
+    fun $inv2 4
+    var S
+      type-con $state 0
+    var S
+      type-con $state 0
+    var #p
+      type-con $ptr 0
+    var #t
+      type-con $ctype 0
+axiom 0
+    forall 1 1 3
+      var S
+        type-con $state 0
+      pat 1
+        fun $good_state 1
+        var S
+          type-con $state 0
+      attribute qid 1
+        string-attr VccPrelu.1152:15
+      attribute uniqueId 1
+        string-attr 128
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    fun $good_state 1
+    var S
+      type-con $state 0
+    forall 2 1 3
+      var #p
+        type-con $ptr 0
+      var #q
+        type-con $ptr 0
+      pat 1
+        fun $set_in 2
+        var #p
+          type-con $ptr 0
+        fun $owns 2
+        var S
+          type-con $state 0
+        var #q
+          type-con $ptr 0
+      attribute qid 1
+        string-attr VccPrelu.846:13
+      attribute uniqueId 1
+        string-attr 86
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 3
+    fun $good_state 1
+    var S
+      type-con $state 0
+    fun $set_in 2
+    var #p
+      type-con $ptr 0
+    fun $owns 2
+    var S
+      type-con $state 0
+    var #q
+      type-con $ptr 0
+    fun $closed 2
+    var S
+      type-con $state 0
+    var #q
+      type-con $ptr 0
+    and 2
+    fun $closed 2
+    var S
+      type-con $state 0
+    var #p
+      type-con $ptr 0
+    not
+    =
+    fun $ref 1
+    var #p
+      type-con $ptr 0
+    int-num 0
+axiom 0
+    forall 3 1 3
+      var S
+        type-con $state 0
+      var p
+        type-con $ptr 0
+      var v
+        int
+      pat 1
+        fun $update_int 3
+        var S
+          type-con $state 0
+        var p
+          type-con $ptr 0
+        var v
+          int
+      attribute qid 1
+        string-attr VccPrelu.1260:15
+      attribute uniqueId 1
+        string-attr 138
+      attribute bvZ3Native 1
+        string-attr False
+    and 6
+    =
+    fun $typemap 1
+    fun $update_int 3
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    var v
+      int
+    fun $typemap 1
+    var S
+      type-con $state 0
+    =
+    fun $statusmap 1
+    fun $update_int 3
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    var v
+      int
+    fun $statusmap 1
+    var S
+      type-con $state 0
+    =
+    fun $memory 1
+    fun $update_int 3
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    var v
+      int
+    fun $store.mem 3
+    fun $memory 1
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    var v
+      int
+    <
+    fun $current_timestamp 1
+    var S
+      type-con $state 0
+    fun $current_timestamp 1
+    fun $update_int 3
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    var v
+      int
+    forall 1 1 4
+      var p
+        type-con $ptr 0
+      pat 1
+        fun $timestamp 2
+        fun $update_int 3
+        var S
+          type-con $state 0
+        var p
+          type-con $ptr 0
+        var v
+          int
+        var p
+          type-con $ptr 0
+      attribute qid 1
+        string-attr VccPrelu.1280:13
+      attribute uniqueId 1
+        string-attr 140
+      attribute bvZ3Native 1
+        string-attr False
+      attribute weight 1
+        expr-attr
+          int-num 0
+    <=
+    fun $timestamp 2
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    fun $timestamp 2
+    fun $update_int 3
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    var v
+      int
+    var p
+      type-con $ptr 0
+    fun $call_transition 2
+    var S
+      type-con $state 0
+    fun $update_int 3
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    var v
+      int
+axiom 0
+    forall 3 1 3
+      var S
+        type-con $state 0
+      var l
+        type-con $ptr 0
+      var p
+        type-con $ptr 0
+      pat 1
+        fun $take_over 3
+        var S
+          type-con $state 0
+        var l
+          type-con $ptr 0
+        var p
+          type-con $ptr 0
+      attribute qid 1
+        string-attr VccPrelu.1309:15
+      attribute uniqueId 1
+        string-attr 141
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    not
+    =
+    fun $kind_of 1
+    fun $typ 1
+    var l
+      type-con $ptr 0
+    fun $kind_primitive 0
+    and 5
+    =
+    fun $statusmap 1
+    fun $take_over 3
+    var S
+      type-con $state 0
+    var l
+      type-con $ptr 0
+    var p
+      type-con $ptr 0
+    fun $store.sm 3
+    fun $statusmap 1
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    fun $taken_over 3
+    var S
+      type-con $state 0
+    var l
+      type-con $ptr 0
+    var p
+      type-con $ptr 0
+    fun $closed 2
+    fun $take_over 3
+    var S
+      type-con $state 0
+    var l
+      type-con $ptr 0
+    var p
+      type-con $ptr 0
+    var p
+      type-con $ptr 0
+    =
+    fun $owner 2
+    fun $take_over 3
+    var S
+      type-con $state 0
+    var l
+      type-con $ptr 0
+    var p
+      type-con $ptr 0
+    var p
+      type-con $ptr 0
+    var l
+      type-con $ptr 0
+    =
+    fun $ref_cnt 2
+    fun $take_over 3
+    var S
+      type-con $state 0
+    var l
+      type-con $ptr 0
+    var p
+      type-con $ptr 0
+    var p
+      type-con $ptr 0
+    fun $ref_cnt 2
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    true
+axiom 0
+    forall 4 1 3
+      var S0
+        type-con $state 0
+      var S
+        type-con $state 0
+      var l
+        type-con $ptr 0
+      var p
+        type-con $ptr 0
+      pat 1
+        fun $release 4
+        var S0
+          type-con $state 0
+        var S
+          type-con $state 0
+        var l
+          type-con $ptr 0
+        var p
+          type-con $ptr 0
+      attribute qid 1
+        string-attr VccPrelu.1325:15
+      attribute uniqueId 1
+        string-attr 142
+      attribute bvZ3Native 1
+        string-attr False
+    and 6
+    =
+    fun $statusmap 1
+    fun $release 4
+    var S0
+      type-con $state 0
+    var S
+      type-con $state 0
+    var l
+      type-con $ptr 0
+    var p
+      type-con $ptr 0
+    fun $store.sm 3
+    fun $statusmap 1
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    fun $released 3
+    var S
+      type-con $state 0
+    var l
+      type-con $ptr 0
+    var p
+      type-con $ptr 0
+    fun $closed 2
+    fun $release 4
+    var S0
+      type-con $state 0
+    var S
+      type-con $state 0
+    var l
+      type-con $ptr 0
+    var p
+      type-con $ptr 0
+    var p
+      type-con $ptr 0
+    =
+    fun $owner 2
+    fun $release 4
+    var S0
+      type-con $state 0
+    var S
+      type-con $state 0
+    var l
+      type-con $ptr 0
+    var p
+      type-con $ptr 0
+    var p
+      type-con $ptr 0
+    fun $me 0
+    =
+    fun $ref_cnt 2
+    fun $release 4
+    var S0
+      type-con $state 0
+    var S
+      type-con $state 0
+    var l
+      type-con $ptr 0
+    var p
+      type-con $ptr 0
+    var p
+      type-con $ptr 0
+    fun $ref_cnt 2
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    =
+    fun $timestamp 2
+    fun $release 4
+    var S0
+      type-con $state 0
+    var S
+      type-con $state 0
+    var l
+      type-con $ptr 0
+    var p
+      type-con $ptr 0
+    var p
+      type-con $ptr 0
+    fun $current_timestamp 1
+    var S0
+      type-con $state 0
+    true
+axiom 0
+    =
+    fun $get_memory_allocator 0
+    fun $ptr 2
+    fun $memory_allocator_type 0
+    fun $memory_allocator_ref 0
+axiom 0
+    =
+    fun $ptr_level 1
+    fun $memory_allocator_type 0
+    int-num 0
+axiom 0
+    forall 1 1 3
+      var S
+        type-con $state 0
+      pat 1
+        fun $program_entry_point 1
+        var S
+          type-con $state 0
+      attribute qid 1
+        string-attr VccPrelu.1661:15
+      attribute uniqueId 1
+        string-attr 175
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    fun $program_entry_point 1
+    var S
+      type-con $state 0
+    fun $program_entry_point_ch 1
+    var S
+      type-con $state 0
+axiom 0
+    forall 3 1 3
+      var S
+        type-con $state 0
+      var p
+        type-con $ptr 0
+      var q
+        type-con $ptr 0
+      pat 1
+        fun $set_in 2
+        var p
+          type-con $ptr 0
+        fun $volatile_span 2
+        var S
+          type-con $state 0
+        var q
+          type-con $ptr 0
+      attribute qid 1
+        string-attr VccPrelu.1745:15
+      attribute uniqueId 1
+        string-attr 186
+      attribute bvZ3Native 1
+        string-attr False
+    =
+    fun $set_in 2
+    var p
+      type-con $ptr 0
+    fun $volatile_span 2
+    var S
+      type-con $state 0
+    var q
+      type-con $ptr 0
+    or 2
+    =
+    var p
+      type-con $ptr 0
+    var q
+      type-con $ptr 0
+    and 2
+    fun $ts_is_volatile 1
+    fun $select.tm 2
+    fun $typemap 1
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    fun $set_in 2
+    var p
+      type-con $ptr 0
+    fun $span 1
+    var q
+      type-con $ptr 0
+axiom 0
+    forall 2 1 3
+      var a
+        type-con $ptr 0
+      var i
+        int
+      pat 1
+        fun $left_split 2
+        var a
+          type-con $ptr 0
+        var i
+          int
+      attribute qid 1
+        string-attr VccPrelu.1752:22
+      attribute uniqueId 1
+        string-attr 187
+      attribute bvZ3Native 1
+        string-attr False
+    =
+    fun $left_split 2
+    var a
+      type-con $ptr 0
+    var i
+      int
+    fun $ptr 2
+    fun $array 2
+    fun $element_type 1
+    fun $typ 1
+    var a
+      type-con $ptr 0
+    var i
+      int
+    fun $ref 1
+    var a
+      type-con $ptr 0
+axiom 0
+    forall 2 1 3
+      var a
+        type-con $ptr 0
+      var i
+        int
+      pat 1
+        fun $right_split 2
+        var a
+          type-con $ptr 0
+        var i
+          int
+      attribute qid 1
+        string-attr VccPrelu.1754:23
+      attribute uniqueId 1
+        string-attr 188
+      attribute bvZ3Native 1
+        string-attr False
+    =
+    fun $right_split 2
+    var a
+      type-con $ptr 0
+    var i
+      int
+    fun $ptr 2
+    fun $array 2
+    fun $element_type 1
+    fun $typ 1
+    var a
+      type-con $ptr 0
+    -
+    fun $array_length 1
+    fun $typ 1
+    var a
+      type-con $ptr 0
+    var i
+      int
+    fun $ref 1
+    fun $idx 3
+    fun $ptr 2
+    fun $element_type 1
+    fun $typ 1
+    var a
+      type-con $ptr 0
+    fun $ref 1
+    var a
+      type-con $ptr 0
+    var i
+      int
+    fun $element_type 1
+    fun $typ 1
+    var a
+      type-con $ptr 0
+axiom 0
+    forall 2 1 3
+      var a1
+        type-con $ptr 0
+      var a2
+        type-con $ptr 0
+      pat 1
+        fun $joined_array 2
+        var a1
+          type-con $ptr 0
+        var a2
+          type-con $ptr 0
+      attribute qid 1
+        string-attr VccPrelu.1757:24
+      attribute uniqueId 1
+        string-attr 189
+      attribute bvZ3Native 1
+        string-attr False
+    =
+    fun $joined_array 2
+    var a1
+      type-con $ptr 0
+    var a2
+      type-con $ptr 0
+    fun $ptr 2
+    fun $array 2
+    fun $element_type 1
+    fun $typ 1
+    var a1
+      type-con $ptr 0
+    +
+    fun $array_length 1
+    fun $typ 1
+    var a1
+      type-con $ptr 0
+    fun $array_length 1
+    fun $typ 1
+    var a2
+      type-con $ptr 0
+    fun $ref 1
+    var a1
+      type-con $ptr 0
+axiom 0
+    forall 1 1 4
+      var #o
+        type-con $ptr 0
+      pat 1
+        fun $set_in 2
+        var #o
+          type-con $ptr 0
+        fun $set_empty 0
+      attribute qid 1
+        string-attr VccPrelu.1854:15
+      attribute uniqueId 1
+        string-attr 198
+      attribute bvZ3Native 1
+        string-attr False
+      attribute weight 1
+        expr-attr
+          int-num 0
+    not
+    fun $set_in 2
+    var #o
+      type-con $ptr 0
+    fun $set_empty 0
+axiom 0
+    forall 2 1 4
+      var #r
+        type-con $ptr 0
+      var #o
+        type-con $ptr 0
+      pat 1
+        fun $set_in 2
+        var #o
+          type-con $ptr 0
+        fun $set_singleton 1
+        var #r
+          type-con $ptr 0
+      attribute qid 1
+        string-attr VccPrelu.1857:15
+      attribute uniqueId 1
+        string-attr 199
+      attribute bvZ3Native 1
+        string-attr False
+      attribute weight 1
+        expr-attr
+          int-num 0
+    =
+    fun $set_in 2
+    var #o
+      type-con $ptr 0
+    fun $set_singleton 1
+    var #r
+      type-con $ptr 0
+    =
+    var #r
+      type-con $ptr 0
+    var #o
+      type-con $ptr 0
+axiom 0
+    forall 2 1 4
+      var #r
+        type-con $ptr 0
+      var #o
+        type-con $ptr 0
+      pat 1
+        fun $set_in 2
+        var #o
+          type-con $ptr 0
+        fun $non_null_set_singleton 1
+        var #r
+          type-con $ptr 0
+      attribute qid 1
+        string-attr VccPrelu.1860:15
+      attribute uniqueId 1
+        string-attr 200
+      attribute bvZ3Native 1
+        string-attr False
+      attribute weight 1
+        expr-attr
+          int-num 0
+    =
+    fun $set_in 2
+    var #o
+      type-con $ptr 0
+    fun $non_null_set_singleton 1
+    var #r
+      type-con $ptr 0
+    and 2
+    =
+    var #r
+      type-con $ptr 0
+    var #o
+      type-con $ptr 0
+    not
+    =
+    fun $ref 1
+    var #r
+      type-con $ptr 0
+    fun $ref 1
+    fun $null 0
+axiom 0
+    forall 3 1 4
+      var #a
+        type-con $ptrset 0
+      var #b
+        type-con $ptrset 0
+      var #o
+        type-con $ptr 0
+      pat 1
+        fun $set_in 2
+        var #o
+          type-con $ptr 0
+        fun $set_union 2
+        var #a
+          type-con $ptrset 0
+        var #b
+          type-con $ptrset 0
+      attribute qid 1
+        string-attr VccPrelu.1863:15
+      attribute uniqueId 1
+        string-attr 201
+      attribute bvZ3Native 1
+        string-attr False
+      attribute weight 1
+        expr-attr
+          int-num 0
+    =
+    fun $set_in 2
+    var #o
+      type-con $ptr 0
+    fun $set_union 2
+    var #a
+      type-con $ptrset 0
+    var #b
+      type-con $ptrset 0
+    or 2
+    fun $set_in 2
+    var #o
+      type-con $ptr 0
+    var #a
+      type-con $ptrset 0
+    fun $set_in 2
+    var #o
+      type-con $ptr 0
+    var #b
+      type-con $ptrset 0
+axiom 0
+    forall 3 1 4
+      var #a
+        type-con $ptrset 0
+      var #b
+        type-con $ptrset 0
+      var #o
+        type-con $ptr 0
+      pat 1
+        fun $set_in 2
+        var #o
+          type-con $ptr 0
+        fun $set_difference 2
+        var #a
+          type-con $ptrset 0
+        var #b
+          type-con $ptrset 0
+      attribute qid 1
+        string-attr VccPrelu.1867:15
+      attribute uniqueId 1
+        string-attr 202
+      attribute bvZ3Native 1
+        string-attr False
+      attribute weight 1
+        expr-attr
+          int-num 0
+    =
+    fun $set_in 2
+    var #o
+      type-con $ptr 0
+    fun $set_difference 2
+    var #a
+      type-con $ptrset 0
+    var #b
+      type-con $ptrset 0
+    and 2
+    fun $set_in 2
+    var #o
+      type-con $ptr 0
+    var #a
+      type-con $ptrset 0
+    not
+    fun $set_in 2
+    var #o
+      type-con $ptr 0
+    var #b
+      type-con $ptrset 0
+axiom 0
+    forall 3 1 4
+      var #a
+        type-con $ptrset 0
+      var #b
+        type-con $ptrset 0
+      var #o
+        type-con $ptr 0
+      pat 1
+        fun $set_in 2
+        var #o
+          type-con $ptr 0
+        fun $set_intersection 2
+        var #a
+          type-con $ptrset 0
+        var #b
+          type-con $ptrset 0
+      attribute qid 1
+        string-attr VccPrelu.1871:15
+      attribute uniqueId 1
+        string-attr 203
+      attribute bvZ3Native 1
+        string-attr False
+      attribute weight 1
+        expr-attr
+          int-num 0
+    =
+    fun $set_in 2
+    var #o
+      type-con $ptr 0
+    fun $set_intersection 2
+    var #a
+      type-con $ptrset 0
+    var #b
+      type-con $ptrset 0
+    and 2
+    fun $set_in 2
+    var #o
+      type-con $ptr 0
+    var #a
+      type-con $ptrset 0
+    fun $set_in 2
+    var #o
+      type-con $ptr 0
+    var #b
+      type-con $ptrset 0
+axiom 0
+    forall 2 1 4
+      var #a
+        type-con $ptrset 0
+      var #b
+        type-con $ptrset 0
+      pat 1
+        fun $set_subset 2
+        var #a
+          type-con $ptrset 0
+        var #b
+          type-con $ptrset 0
+      attribute qid 1
+        string-attr VccPrelu.1875:14
+      attribute uniqueId 1
+        string-attr 205
+      attribute bvZ3Native 1
+        string-attr False
+      attribute weight 1
+        expr-attr
+          int-num 0
+    =
+    fun $set_subset 2
+    var #a
+      type-con $ptrset 0
+    var #b
+      type-con $ptrset 0
+    forall 1 2 4
+      var #o
+        type-con $ptr 0
+      pat 1
+        fun $set_in 2
+        var #o
+          type-con $ptr 0
+        var #a
+          type-con $ptrset 0
+      pat 1
+        fun $set_in 2
+        var #o
+          type-con $ptr 0
+        var #b
+          type-con $ptrset 0
+      attribute qid 1
+        string-attr VccPrelu.1876:35
+      attribute uniqueId 1
+        string-attr 204
+      attribute bvZ3Native 1
+        string-attr False
+      attribute weight 1
+        expr-attr
+          int-num 0
+    implies
+    fun $set_in 2
+    var #o
+      type-con $ptr 0
+    var #a
+      type-con $ptrset 0
+    fun $set_in 2
+    var #o
+      type-con $ptr 0
+    var #b
+      type-con $ptrset 0
+axiom 0
+    forall 2 1 4
+      var #a
+        type-con $ptrset 0
+      var #b
+        type-con $ptrset 0
+      pat 1
+        fun $set_eq 2
+        var #a
+          type-con $ptrset 0
+        var #b
+          type-con $ptrset 0
+      attribute qid 1
+        string-attr VccPrelu.1880:15
+      attribute uniqueId 1
+        string-attr 207
+      attribute bvZ3Native 1
+        string-attr False
+      attribute weight 1
+        expr-attr
+          int-num 0
+    implies
+    forall 1 1 4
+      var #o
+        type-con $ptr 0
+      pat 1
+        fun $dont_instantiate 1
+        var #o
+          type-con $ptr 0
+      attribute qid 1
+        string-attr VccPrelu.1881:11
+      attribute uniqueId 1
+        string-attr 206
+      attribute bvZ3Native 1
+        string-attr False
+      attribute weight 1
+        expr-attr
+          int-num 0
+    =
+    fun $set_in 2
+    var #o
+      type-con $ptr 0
+    var #a
+      type-con $ptrset 0
+    fun $set_in 2
+    var #o
+      type-con $ptr 0
+    var #b
+      type-con $ptrset 0
+    fun $set_eq 2
+    var #a
+      type-con $ptrset 0
+    var #b
+      type-con $ptrset 0
+axiom 0
+    forall 2 1 4
+      var #a
+        type-con $ptrset 0
+      var #b
+        type-con $ptrset 0
+      pat 1
+        fun $set_eq 2
+        var #a
+          type-con $ptrset 0
+        var #b
+          type-con $ptrset 0
+      attribute qid 1
+        string-attr VccPrelu.1882:15
+      attribute uniqueId 1
+        string-attr 208
+      attribute bvZ3Native 1
+        string-attr False
+      attribute weight 1
+        expr-attr
+          int-num 0
+    implies
+    fun $set_eq 2
+    var #a
+      type-con $ptrset 0
+    var #b
+      type-con $ptrset 0
+    =
+    var #a
+      type-con $ptrset 0
+    var #b
+      type-con $ptrset 0
+axiom 0
+    =
+    fun $set_cardinality 1
+    fun $set_empty 0
+    int-num 0
+axiom 0
+    forall 1 0 4
+      var p
+        type-con $ptr 0
+      attribute qid 1
+        string-attr VccPrelu.1888:15
+      attribute uniqueId 1
+        string-attr 209
+      attribute bvZ3Native 1
+        string-attr False
+      attribute weight 1
+        expr-attr
+          int-num 0
+    =
+    fun $set_cardinality 1
+    fun $set_singleton 1
+    var p
+      type-con $ptr 0
+    int-num 1
+axiom 0
+    forall 1 1 4
+      var #o
+        type-con $ptr 0
+      pat 1
+        fun $set_in 2
+        var #o
+          type-con $ptr 0
+        fun $set_universe 0
+      attribute qid 1
+        string-attr VccPrelu.1891:15
+      attribute uniqueId 1
+        string-attr 210
+      attribute bvZ3Native 1
+        string-attr False
+      attribute weight 1
+        expr-attr
+          int-num 0
+    fun $set_in 2
+    var #o
+      type-con $ptr 0
+    fun $set_universe 0
+axiom 0
+    forall 3 1 4
+      var p
+        type-con $ptr 0
+      var s1
+        type-con $ptrset 0
+      var s2
+        type-con $ptrset 0
+      pat 2
+        fun $set_disjoint 2
+        var s1
+          type-con $ptrset 0
+        var s2
+          type-con $ptrset 0
+        fun $set_in 2
+        var p
+          type-con $ptr 0
+        var s1
+          type-con $ptrset 0
+      attribute qid 1
+        string-attr VccPrelu.1896:15
+      attribute uniqueId 1
+        string-attr 211
+      attribute bvZ3Native 1
+        string-attr False
+      attribute weight 1
+        expr-attr
+          int-num 0
+    implies
+    and 2
+    fun $set_disjoint 2
+    var s1
+      type-con $ptrset 0
+    var s2
+      type-con $ptrset 0
+    fun $set_in 2
+    var p
+      type-con $ptr 0
+    var s1
+      type-con $ptrset 0
+    =
+    fun $id_set_disjoint 3
+    var p
+      type-con $ptr 0
+    var s1
+      type-con $ptrset 0
+    var s2
+      type-con $ptrset 0
+    int-num 1
+axiom 0
+    forall 3 1 4
+      var p
+        type-con $ptr 0
+      var s1
+        type-con $ptrset 0
+      var s2
+        type-con $ptrset 0
+      pat 2
+        fun $set_disjoint 2
+        var s1
+          type-con $ptrset 0
+        var s2
+          type-con $ptrset 0
+        fun $set_in 2
+        var p
+          type-con $ptr 0
+        var s2
+          type-con $ptrset 0
+      attribute qid 1
+        string-attr VccPrelu.1899:15
+      attribute uniqueId 1
+        string-attr 212
+      attribute bvZ3Native 1
+        string-attr False
+      attribute weight 1
+        expr-attr
+          int-num 0
+    implies
+    and 2
+    fun $set_disjoint 2
+    var s1
+      type-con $ptrset 0
+    var s2
+      type-con $ptrset 0
+    fun $set_in 2
+    var p
+      type-con $ptr 0
+    var s2
+      type-con $ptrset 0
+    =
+    fun $id_set_disjoint 3
+    var p
+      type-con $ptr 0
+    var s1
+      type-con $ptrset 0
+    var s2
+      type-con $ptrset 0
+    int-num 2
+axiom 0
+    forall 2 1 4
+      var s1
+        type-con $ptrset 0
+      var s2
+        type-con $ptrset 0
+      pat 1
+        fun $set_disjoint 2
+        var s1
+          type-con $ptrset 0
+        var s2
+          type-con $ptrset 0
+      attribute qid 1
+        string-attr VccPrelu.1903:15
+      attribute uniqueId 1
+        string-attr 214
+      attribute bvZ3Native 1
+        string-attr False
+      attribute weight 1
+        expr-attr
+          int-num 0
+    implies
+    forall 1 1 3
+      var p
+        type-con $ptr 0
+      pat 1
+        fun $dont_instantiate 1
+        var p
+          type-con $ptr 0
+      attribute qid 1
+        string-attr VccPrelu.1904:11
+      attribute uniqueId 1
+        string-attr 213
+      attribute bvZ3Native 1
+        string-attr False
+    and 2
+    implies
+    fun $set_in 2
+    var p
+      type-con $ptr 0
+    var s1
+      type-con $ptrset 0
+    not
+    fun $set_in 2
+    var p
+      type-con $ptr 0
+    var s2
+      type-con $ptrset 0
+    implies
+    fun $set_in 2
+    var p
+      type-con $ptr 0
+    var s2
+      type-con $ptrset 0
+    not
+    fun $set_in 2
+    var p
+      type-con $ptr 0
+    var s1
+      type-con $ptrset 0
+    fun $set_disjoint 2
+    var s1
+      type-con $ptrset 0
+    var s2
+      type-con $ptrset 0
+axiom 0
+    forall 3 1 4
+      var p
+        type-con $ptr 0
+      var S1
+        type-con $state 0
+      var p1
+        type-con $ptr 0
+      pat 1
+        fun $set_in 2
+        var p
+          type-con $ptr 0
+        fun $owns 2
+        var S1
+          type-con $state 0
+        var p1
+          type-con $ptr 0
+      attribute qid 1
+        string-attr VccPrelu.1914:15
+      attribute uniqueId 1
+        string-attr 215
+      attribute bvZ3Native 1
+        string-attr False
+      attribute weight 1
+        expr-attr
+          int-num 0
+    implies
+    fun $set_in 2
+    var p
+      type-con $ptr 0
+    fun $owns 2
+    var S1
+      type-con $state 0
+    var p1
+      type-con $ptr 0
+    fun $in_some_owns 1
+    var p
+      type-con $ptr 0
+axiom 0
+    forall 3 1 4
+      var p
+        type-con $ptr 0
+      var S1
+        type-con $state 0
+      var p1
+        type-con $ptr 0
+      pat 2
+        fun $set_in2 2
+        var p
+          type-con $ptr 0
+        fun $owns 2
+        var S1
+          type-con $state 0
+        var p1
+          type-con $ptr 0
+        fun $in_some_owns 1
+        var p
+          type-con $ptr 0
+      attribute qid 1
+        string-attr VccPrelu.1918:15
+      attribute uniqueId 1
+        string-attr 216
+      attribute bvZ3Native 1
+        string-attr False
+      attribute weight 1
+        expr-attr
+          int-num 0
+    =
+    fun $set_in 2
+    var p
+      type-con $ptr 0
+    fun $owns 2
+    var S1
+      type-con $state 0
+    var p1
+      type-con $ptr 0
+    fun $set_in2 2
+    var p
+      type-con $ptr 0
+    fun $owns 2
+    var S1
+      type-con $state 0
+    var p1
+      type-con $ptr 0
+axiom 0
+    forall 2 1 4
+      var p
+        type-con $ptr 0
+      var s
+        type-con $ptrset 0
+      pat 1
+        fun $set_in 2
+        var p
+          type-con $ptr 0
+        var s
+          type-con $ptrset 0
+      attribute qid 1
+        string-attr VccPrelu.1922:15
+      attribute uniqueId 1
+        string-attr 217
+      attribute bvZ3Native 1
+        string-attr False
+      attribute weight 1
+        expr-attr
+          int-num 0
+    =
+    fun $set_in 2
+    var p
+      type-con $ptr 0
+    var s
+      type-con $ptrset 0
+    fun $set_in2 2
+    var p
+      type-con $ptr 0
+    var s
+      type-con $ptrset 0
+axiom 0
+    forall 2 1 4
+      var p
+        type-con $ptr 0
+      var s
+        type-con $ptrset 0
+      pat 1
+        fun $set_in 2
+        var p
+          type-con $ptr 0
+        var s
+          type-con $ptrset 0
+      attribute qid 1
+        string-attr VccPrelu.1924:15
+      attribute uniqueId 1
+        string-attr 218
+      attribute bvZ3Native 1
+        string-attr False
+      attribute weight 1
+        expr-attr
+          int-num 0
+    =
+    fun $set_in 2
+    var p
+      type-con $ptr 0
+    var s
+      type-con $ptrset 0
+    fun $set_in3 2
+    var p
+      type-con $ptr 0
+    var s
+      type-con $ptrset 0
+axiom 0
+    forall 2 1 4
+      var p
+        type-con $ptr 0
+      var s
+        type-con $ptrset 0
+      pat 1
+        fun $set_in0 2
+        var p
+          type-con $ptr 0
+        var s
+          type-con $ptrset 0
+      attribute qid 1
+        string-attr VccPrelu.1928:15
+      attribute uniqueId 1
+        string-attr 219
+      attribute bvZ3Native 1
+        string-attr False
+      attribute weight 1
+        expr-attr
+          int-num 0
+    =
+    fun $set_in 2
+    var p
+      type-con $ptr 0
+    var s
+      type-con $ptrset 0
+    fun $set_in0 2
+    var p
+      type-con $ptr 0
+    var s
+      type-con $ptrset 0
+axiom 0
+    forall 2 1 3
+      var T
+        type-con $ctype 0
+      var s
+        int
+      pat 1
+        fun $array 2
+        var T
+          type-con $ctype 0
+        var s
+          int
+      attribute qid 1
+        string-attr VccPrelu.1989:15
+      attribute uniqueId 1
+        string-attr 224
+      attribute bvZ3Native 1
+        string-attr False
+    =
+    fun $element_type 1
+    fun $array 2
+    var T
+      type-con $ctype 0
+    var s
+      int
+    var T
+      type-con $ctype 0
+axiom 0
+    forall 2 1 3
+      var T
+        type-con $ctype 0
+      var s
+        int
+      pat 1
+        fun $array 2
+        var T
+          type-con $ctype 0
+        var s
+          int
+      attribute qid 1
+        string-attr VccPrelu.1990:15
+      attribute uniqueId 1
+        string-attr 225
+      attribute bvZ3Native 1
+        string-attr False
+    =
+    fun $array_length 1
+    fun $array 2
+    var T
+      type-con $ctype 0
+    var s
+      int
+    var s
+      int
+axiom 0
+    forall 2 1 3
+      var T
+        type-con $ctype 0
+      var s
+        int
+      pat 1
+        fun $array 2
+        var T
+          type-con $ctype 0
+        var s
+          int
+      attribute qid 1
+        string-attr VccPrelu.1991:15
+      attribute uniqueId 1
+        string-attr 226
+      attribute bvZ3Native 1
+        string-attr False
+    =
+    fun $ptr_level 1
+    fun $array 2
+    var T
+      type-con $ctype 0
+    var s
+      int
+    int-num 0
+axiom 0
+    forall 2 1 3
+      var T
+        type-con $ctype 0
+      var s
+        int
+      pat 1
+        fun $array 2
+        var T
+          type-con $ctype 0
+        var s
+          int
+      attribute qid 1
+        string-attr VccPrelu.1992:15
+      attribute uniqueId 1
+        string-attr 227
+      attribute bvZ3Native 1
+        string-attr False
+    fun $is_arraytype 1
+    fun $array 2
+    var T
+      type-con $ctype 0
+    var s
+      int
+axiom 0
+    forall 2 1 3
+      var T
+        type-con $ctype 0
+      var s
+        int
+      pat 1
+        fun $array 2
+        var T
+          type-con $ctype 0
+        var s
+          int
+      attribute qid 1
+        string-attr VccPrelu.1993:15
+      attribute uniqueId 1
+        string-attr 228
+      attribute bvZ3Native 1
+        string-attr False
+    not
+    fun $is_claimable 1
+    fun $array 2
+    var T
+      type-con $ctype 0
+    var s
+      int
+axiom 0
+    forall 2 1 4
+      var p
+        type-con $ptr 0
+      var T
+        type-con $ctype 0
+      pat 1
+        fun $inlined_array 2
+        var p
+          type-con $ptr 0
+        var T
+          type-con $ctype 0
+      attribute qid 1
+        string-attr VccPrelu.1998:37
+      attribute uniqueId 1
+        string-attr 229
+      attribute bvZ3Native 1
+        string-attr False
+      attribute weight 1
+        expr-attr
+          int-num 0
+    =
+    fun $inlined_array 2
+    var p
+      type-con $ptr 0
+    var T
+      type-con $ctype 0
+    var p
+      type-con $ptr 0
+axiom 0
+    forall 3 1 3
+      var #p
+        type-con $ptr 0
+      var #i
+        int
+      var #t
+        type-con $ctype 0
+      pat 1
+        fun $idx 3
+        var #p
+          type-con $ptr 0
+        var #i
+          int
+        var #t
+          type-con $ctype 0
+      attribute qid 1
+        string-attr VccPrelu.2002:15
+      attribute uniqueId 1
+        string-attr 230
+      attribute bvZ3Native 1
+        string-attr False
+    and 2
+    fun $extent_hint 2
+    fun $idx 3
+    var #p
+      type-con $ptr 0
+    var #i
+      int
+    var #t
+      type-con $ctype 0
+    var #p
+      type-con $ptr 0
+    =
+    fun $idx 3
+    var #p
+      type-con $ptr 0
+    var #i
+      int
+    var #t
+      type-con $ctype 0
+    fun $ptr 2
+    var #t
+      type-con $ctype 0
+    +
+    fun $ref 1
+    var #p
+      type-con $ptr 0
+    *
+    var #i
+      int
+    fun $sizeof 1
+    var #t
+      type-con $ctype 0
+axiom 0
+    forall 4 1 3
+      var p
+        type-con $ptr 0
+      var i
+        int
+      var j
+        int
+      var T
+        type-con $ctype 0
+      pat 1
+        fun $idx 3
+        fun $idx 3
+        var p
+          type-con $ptr 0
+        var i
+          int
+        var T
+          type-con $ctype 0
+        var j
+          int
+        var T
+          type-con $ctype 0
+      attribute qid 1
+        string-attr VccPrelu.2016:15
+      attribute uniqueId 1
+        string-attr 231
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 2
+    not
+    =
+    var i
+      int
+    int-num 0
+    not
+    =
+    var j
+      int
+    int-num 0
+    =
+    fun $idx 3
+    fun $idx 3
+    var p
+      type-con $ptr 0
+    var i
+      int
+    var T
+      type-con $ctype 0
+    var j
+      int
+    var T
+      type-con $ctype 0
+    fun $idx 3
+    var p
+      type-con $ptr 0
+    +
+    var i
+      int
+    var j
+      int
+    var T
+      type-con $ctype 0
+axiom 0
+    forall 5 1 4
+      var S
+        type-con $state 0
+      var p
+        type-con $ptr 0
+      var T
+        type-con $ctype 0
+      var sz
+        int
+      var vol
+        bool
+      pat 1
+        fun $is_array_vol_or_nonvol 5
+        var S
+          type-con $state 0
+        var p
+          type-con $ptr 0
+        var T
+          type-con $ctype 0
+        var sz
+          int
+        var vol
+          bool
+      attribute qid 1
+        string-attr VccPrelu.2020:46
+      attribute uniqueId 1
+        string-attr 233
+      attribute bvZ3Native 1
+        string-attr False
+      attribute weight 1
+        expr-attr
+          int-num 0
+    =
+    fun $is_array_vol_or_nonvol 5
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    var T
+      type-con $ctype 0
+    var sz
+      int
+    var vol
+      bool
+    and 2
+    fun $is 2
+    var p
+      type-con $ptr 0
+    var T
+      type-con $ctype 0
+    forall 1 3 3
+      var i
+        int
+      pat 1
+        fun $select.sm 2
+        fun $statusmap 1
+        var S
+          type-con $state 0
+        fun $idx 3
+        var p
+          type-con $ptr 0
+        var i
+          int
+        var T
+          type-con $ctype 0
+      pat 1
+        fun $select.tm 2
+        fun $typemap 1
+        var S
+          type-con $state 0
+        fun $idx 3
+        var p
+          type-con $ptr 0
+        var i
+          int
+        var T
+          type-con $ctype 0
+      pat 1
+        fun $select.mem 2
+        fun $memory 1
+        var S
+          type-con $state 0
+        fun $idx 3
+        var p
+          type-con $ptr 0
+        var i
+          int
+        var T
+          type-con $ctype 0
+      attribute qid 1
+        string-attr VccPrelu.2022:13
+      attribute uniqueId 1
+        string-attr 232
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 2
+    <=
+    int-num 0
+    var i
+      int
+    <
+    var i
+      int
+    var sz
+      int
+    and 3
+    =
+    fun $ts_is_volatile 1
+    fun $select.tm 2
+    fun $typemap 1
+    var S
+      type-con $state 0
+    fun $idx 3
+    var p
+      type-con $ptr 0
+    var i
+      int
+    var T
+      type-con $ctype 0
+    var vol
+      bool
+    fun $ts_is_array_elt 1
+    fun $select.tm 2
+    fun $typemap 1
+    var S
+      type-con $state 0
+    fun $idx 3
+    var p
+      type-con $ptr 0
+    var i
+      int
+    var T
+      type-con $ctype 0
+    fun $typed 2
+    var S
+      type-con $state 0
+    fun $idx 3
+    var p
+      type-con $ptr 0
+    var i
+      int
+    var T
+      type-con $ctype 0
+axiom 0
+    forall 4 1 4
+      var S
+        type-con $state 0
+      var p
+        type-con $ptr 0
+      var T
+        type-con $ctype 0
+      var sz
+        int
+      pat 1
+        fun $is_array 4
+        var S
+          type-con $state 0
+        var p
+          type-con $ptr 0
+        var T
+          type-con $ctype 0
+        var sz
+          int
+      attribute qid 1
+        string-attr VccPrelu.2026:32
+      attribute uniqueId 1
+        string-attr 235
+      attribute bvZ3Native 1
+        string-attr False
+      attribute weight 1
+        expr-attr
+          int-num 0
+    =
+    fun $is_array 4
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    var T
+      type-con $ctype 0
+    var sz
+      int
+    and 2
+    fun $is 2
+    var p
+      type-con $ptr 0
+    var T
+      type-con $ctype 0
+    forall 1 3 3
+      var i
+        int
+      pat 1
+        fun $select.sm 2
+        fun $statusmap 1
+        var S
+          type-con $state 0
+        fun $idx 3
+        var p
+          type-con $ptr 0
+        var i
+          int
+        var T
+          type-con $ctype 0
+      pat 1
+        fun $select.tm 2
+        fun $typemap 1
+        var S
+          type-con $state 0
+        fun $idx 3
+        var p
+          type-con $ptr 0
+        var i
+          int
+        var T
+          type-con $ctype 0
+      pat 1
+        fun $select.mem 2
+        fun $memory 1
+        var S
+          type-con $state 0
+        fun $idx 3
+        var p
+          type-con $ptr 0
+        var i
+          int
+        var T
+          type-con $ctype 0
+      attribute qid 1
+        string-attr VccPrelu.2028:13
+      attribute uniqueId 1
+        string-attr 234
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 2
+    <=
+    int-num 0
+    var i
+      int
+    <
+    var i
+      int
+    var sz
+      int
+    and 2
+    fun $ts_is_array_elt 1
+    fun $select.tm 2
+    fun $typemap 1
+    var S
+      type-con $state 0
+    fun $idx 3
+    var p
+      type-con $ptr 0
+    var i
+      int
+    var T
+      type-con $ctype 0
+    fun $typed 2
+    var S
+      type-con $state 0
+    fun $idx 3
+    var p
+      type-con $ptr 0
+    var i
+      int
+    var T
+      type-con $ctype 0
+axiom 0
+    forall 4 1 3
+      var p
+        type-con $ptr 0
+      var #r
+        int
+      var T
+        type-con $ctype 0
+      var sz
+        int
+      pat 1
+        fun $set_in 2
+        var p
+          type-con $ptr 0
+        fun $full_extent 1
+        fun $ptr 2
+        fun $array 2
+        var T
+          type-con $ctype 0
+        var sz
+          int
+        var #r
+          int
+      attribute qid 1
+        string-attr VccPrelu.2094:15
+      attribute uniqueId 1
+        string-attr 243
+      attribute bvZ3Native 1
+        string-attr False
+    =
+    fun $set_in 2
+    var p
+      type-con $ptr 0
+    fun $full_extent 1
+    fun $ptr 2
+    fun $array 2
+    var T
+      type-con $ctype 0
+    var sz
+      int
+    var #r
+      int
+    or 2
+    =
+    var p
+      type-con $ptr 0
+    fun $ptr 2
+    fun $array 2
+    var T
+      type-con $ctype 0
+    var sz
+      int
+    var #r
+      int
+    and 3
+    <=
+    int-num 0
+    fun $index_within 2
+    var p
+      type-con $ptr 0
+    fun $ptr 2
+    var T
+      type-con $ctype 0
+    var #r
+      int
+    <=
+    fun $index_within 2
+    var p
+      type-con $ptr 0
+    fun $ptr 2
+    var T
+      type-con $ctype 0
+    var #r
+      int
+    -
+    var sz
+      int
+    int-num 1
+    fun $set_in 2
+    var p
+      type-con $ptr 0
+    fun $full_extent 1
+    fun $idx 3
+    fun $ptr 2
+    var T
+      type-con $ctype 0
+    var #r
+      int
+    fun $index_within 2
+    var p
+      type-con $ptr 0
+    fun $ptr 2
+    var T
+      type-con $ctype 0
+    var #r
+      int
+    var T
+      type-con $ctype 0
+axiom 0
+    forall 5 1 3
+      var S
+        type-con $state 0
+      var p
+        type-con $ptr 0
+      var #r
+        int
+      var T
+        type-con $ctype 0
+      var sz
+        int
+      pat 1
+        fun $set_in 2
+        var p
+          type-con $ptr 0
+        fun $extent 2
+        var S
+          type-con $state 0
+        fun $ptr 2
+        fun $array 2
+        var T
+          type-con $ctype 0
+        var sz
+          int
+        var #r
+          int
+      attribute qid 1
+        string-attr VccPrelu.2099:15
+      attribute uniqueId 1
+        string-attr 244
+      attribute bvZ3Native 1
+        string-attr False
+    =
+    fun $set_in 2
+    var p
+      type-con $ptr 0
+    fun $extent 2
+    var S
+      type-con $state 0
+    fun $ptr 2
+    fun $array 2
+    var T
+      type-con $ctype 0
+    var sz
+      int
+    var #r
+      int
+    or 2
+    =
+    var p
+      type-con $ptr 0
+    fun $ptr 2
+    fun $array 2
+    var T
+      type-con $ctype 0
+    var sz
+      int
+    var #r
+      int
+    and 3
+    <=
+    int-num 0
+    fun $index_within 2
+    var p
+      type-con $ptr 0
+    fun $ptr 2
+    var T
+      type-con $ctype 0
+    var #r
+      int
+    <=
+    fun $index_within 2
+    var p
+      type-con $ptr 0
+    fun $ptr 2
+    var T
+      type-con $ctype 0
+    var #r
+      int
+    -
+    var sz
+      int
+    int-num 1
+    fun $set_in 2
+    var p
+      type-con $ptr 0
+    fun $extent 2
+    var S
+      type-con $state 0
+    fun $idx 3
+    fun $ptr 2
+    var T
+      type-con $ctype 0
+    var #r
+      int
+    fun $index_within 2
+    var p
+      type-con $ptr 0
+    fun $ptr 2
+    var T
+      type-con $ctype 0
+    var #r
+      int
+    var T
+      type-con $ctype 0
+axiom 0
+    forall 5 2 3
+      var S
+        type-con $state 0
+      var #r
+        int
+      var T
+        type-con $ctype 0
+      var sz
+        int
+      var i
+        int
+      pat 2
+        fun $select.sm 2
+        fun $statusmap 1
+        var S
+          type-con $state 0
+        fun $idx 3
+        fun $ptr 2
+        var T
+          type-con $ctype 0
+        var #r
+          int
+        var i
+          int
+        var T
+          type-con $ctype 0
+        fun $ptr 2
+        fun $array 2
+        var T
+          type-con $ctype 0
+        var sz
+          int
+        var #r
+          int
+      pat 2
+        fun $select.tm 2
+        fun $typemap 1
+        var S
+          type-con $state 0
+        fun $idx 3
+        fun $ptr 2
+        var T
+          type-con $ctype 0
+        var #r
+          int
+        var i
+          int
+        var T
+          type-con $ctype 0
+        fun $ptr 2
+        fun $array 2
+        var T
+          type-con $ctype 0
+        var sz
+          int
+        var #r
+          int
+      attribute qid 1
+        string-attr VccPrelu.2107:15
+      attribute uniqueId 1
+        string-attr 245
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    fun $typed 2
+    var S
+      type-con $state 0
+    fun $ptr 2
+    fun $array 2
+    var T
+      type-con $ctype 0
+    var sz
+      int
+    var #r
+      int
+    implies
+    and 2
+    <=
+    int-num 0
+    var i
+      int
+    <
+    var i
+      int
+    var sz
+      int
+    and 4
+    =
+    fun $ts_emb 1
+    fun $select.tm 2
+    fun $typemap 1
+    var S
+      type-con $state 0
+    fun $idx 3
+    fun $ptr 2
+    fun $array 2
+    var T
+      type-con $ctype 0
+    var sz
+      int
+    var #r
+      int
+    var i
+      int
+    var T
+      type-con $ctype 0
+    fun $ptr 2
+    fun $array 2
+    var T
+      type-con $ctype 0
+    var sz
+      int
+    var #r
+      int
+    not
+    fun $ts_is_volatile 1
+    fun $select.tm 2
+    fun $typemap 1
+    var S
+      type-con $state 0
+    fun $idx 3
+    fun $ptr 2
+    fun $array 2
+    var T
+      type-con $ctype 0
+    var sz
+      int
+    var #r
+      int
+    var i
+      int
+    var T
+      type-con $ctype 0
+    fun $ts_is_array_elt 1
+    fun $select.tm 2
+    fun $typemap 1
+    var S
+      type-con $state 0
+    fun $idx 3
+    fun $ptr 2
+    fun $array 2
+    var T
+      type-con $ctype 0
+    var sz
+      int
+    var #r
+      int
+    var i
+      int
+    var T
+      type-con $ctype 0
+    fun $typed 2
+    var S
+      type-con $state 0
+    fun $idx 3
+    fun $ptr 2
+    fun $array 2
+    var T
+      type-con $ctype 0
+    var sz
+      int
+    var #r
+      int
+    var i
+      int
+    var T
+      type-con $ctype 0
+axiom 0
+    forall 4 1 3
+      var p
+        type-con $ptr 0
+      var T
+        type-con $ctype 0
+      var sz
+        int
+      var elem
+        type-con $ptr 0
+      pat 1
+        fun $set_in 2
+        var elem
+          type-con $ptr 0
+        fun $array_members 3
+        var p
+          type-con $ptr 0
+        var T
+          type-con $ctype 0
+        var sz
+          int
+      attribute qid 1
+        string-attr VccPrelu.2116:15
+      attribute uniqueId 1
+        string-attr 246
+      attribute bvZ3Native 1
+        string-attr False
+    =
+    fun $set_in 2
+    var elem
+      type-con $ptr 0
+    fun $array_members 3
+    var p
+      type-con $ptr 0
+    var T
+      type-con $ctype 0
+    var sz
+      int
+    and 3
+    <=
+    int-num 0
+    fun $index_within 2
+    var elem
+      type-con $ptr 0
+    var p
+      type-con $ptr 0
+    <=
+    fun $index_within 2
+    var elem
+      type-con $ptr 0
+    var p
+      type-con $ptr 0
+    -
+    var sz
+      int
+    int-num 1
+    =
+    var elem
+      type-con $ptr 0
+    fun $idx 3
+    var p
+      type-con $ptr 0
+    fun $index_within 2
+    var elem
+      type-con $ptr 0
+    var p
+      type-con $ptr 0
+    var T
+      type-con $ctype 0
+axiom 0
+    forall 4 1 3
+      var p
+        type-con $ptr 0
+      var #r
+        int
+      var T
+        type-con $ctype 0
+      var sz
+        int
+      pat 1
+        fun $set_in 2
+        var p
+          type-con $ptr 0
+        fun $array_range 3
+        fun $ptr 2
+        var T
+          type-con $ctype 0
+        var #r
+          int
+        var T
+          type-con $ctype 0
+        var sz
+          int
+      attribute qid 1
+        string-attr VccPrelu.2122:15
+      attribute uniqueId 1
+        string-attr 247
+      attribute bvZ3Native 1
+        string-attr False
+    =
+    fun $set_in 2
+    var p
+      type-con $ptr 0
+    fun $array_range 3
+    fun $ptr 2
+    var T
+      type-con $ctype 0
+    var #r
+      int
+    var T
+      type-con $ctype 0
+    var sz
+      int
+    and 3
+    <=
+    int-num 0
+    fun $index_within 2
+    var p
+      type-con $ptr 0
+    fun $ptr 2
+    var T
+      type-con $ctype 0
+    var #r
+      int
+    <=
+    fun $index_within 2
+    var p
+      type-con $ptr 0
+    fun $ptr 2
+    var T
+      type-con $ctype 0
+    var #r
+      int
+    -
+    var sz
+      int
+    int-num 1
+    fun $set_in 2
+    var p
+      type-con $ptr 0
+    fun $full_extent 1
+    fun $idx 3
+    fun $ptr 2
+    var T
+      type-con $ctype 0
+    var #r
+      int
+    fun $index_within 2
+    var p
+      type-con $ptr 0
+    fun $ptr 2
+    var T
+      type-con $ctype 0
+    var #r
+      int
+    var T
+      type-con $ctype 0
+axiom 0
+    forall 5 1 3
+      var p
+        type-con $ptr 0
+      var T
+        type-con $ctype 0
+      var sz
+        int
+      var idx
+        int
+      var S
+        type-con $ptrset 0
+      pat 2
+        fun $idx 3
+        var p
+          type-con $ptr 0
+        var idx
+          int
+        var T
+          type-con $ctype 0
+        fun $set_disjoint 2
+        fun $array_range 3
+        var p
+          type-con $ptr 0
+        var T
+          type-con $ctype 0
+        var sz
+          int
+        var S
+          type-con $ptrset 0
+      attribute qid 1
+        string-attr VccPrelu.2126:15
+      attribute uniqueId 1
+        string-attr 248
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 2
+    <=
+    int-num 0
+    var idx
+      int
+    <
+    var idx
+      int
+    var sz
+      int
+    =
+    fun $id_set_disjoint 3
+    fun $idx 3
+    var p
+      type-con $ptr 0
+    var idx
+      int
+    var T
+      type-con $ctype 0
+    fun $array_range 3
+    var p
+      type-con $ptr 0
+    var T
+      type-con $ctype 0
+    var sz
+      int
+    var S
+      type-con $ptrset 0
+    int-num 1
+axiom 0
+    forall 5 1 3
+      var p
+        type-con $ptr 0
+      var T
+        type-con $ctype 0
+      var sz
+        int
+      var idx
+        int
+      var S
+        type-con $ptrset 0
+      pat 2
+        fun $idx 3
+        var p
+          type-con $ptr 0
+        var idx
+          int
+        var T
+          type-con $ctype 0
+        fun $set_disjoint 2
+        var S
+          type-con $ptrset 0
+        fun $array_range 3
+        var p
+          type-con $ptr 0
+        var T
+          type-con $ctype 0
+        var sz
+          int
+      attribute qid 1
+        string-attr VccPrelu.2130:15
+      attribute uniqueId 1
+        string-attr 249
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 2
+    <=
+    int-num 0
+    var idx
+      int
+    <
+    var idx
+      int
+    var sz
+      int
+    =
+    fun $id_set_disjoint 3
+    fun $idx 3
+    var p
+      type-con $ptr 0
+    var idx
+      int
+    var T
+      type-con $ctype 0
+    var S
+      type-con $ptrset 0
+    fun $array_range 3
+    var p
+      type-con $ptr 0
+    var T
+      type-con $ctype 0
+    var sz
+      int
+    int-num 2
+axiom 0
+    forall 4 1 3
+      var p
+        type-con $ptr 0
+      var #r
+        int
+      var T
+        type-con $ctype 0
+      var sz
+        int
+      pat 1
+        fun $set_in 2
+        var p
+          type-con $ptr 0
+        fun $non_null_array_range 3
+        fun $ptr 2
+        var T
+          type-con $ctype 0
+        var #r
+          int
+        var T
+          type-con $ctype 0
+        var sz
+          int
+      attribute qid 1
+        string-attr VccPrelu.2135:15
+      attribute uniqueId 1
+        string-attr 250
+      attribute bvZ3Native 1
+        string-attr False
+    =
+    fun $set_in 2
+    var p
+      type-con $ptr 0
+    fun $non_null_array_range 3
+    fun $ptr 2
+    var T
+      type-con $ctype 0
+    var #r
+      int
+    var T
+      type-con $ctype 0
+    var sz
+      int
+    and 4
+    not
+    =
+    var #r
+      int
+    int-num 0
+    <=
+    int-num 0
+    fun $index_within 2
+    var p
+      type-con $ptr 0
+    fun $ptr 2
+    var T
+      type-con $ctype 0
+    var #r
+      int
+    <=
+    fun $index_within 2
+    var p
+      type-con $ptr 0
+    fun $ptr 2
+    var T
+      type-con $ctype 0
+    var #r
+      int
+    -
+    var sz
+      int
+    int-num 1
+    fun $set_in 2
+    var p
+      type-con $ptr 0
+    fun $full_extent 1
+    fun $idx 3
+    fun $ptr 2
+    var T
+      type-con $ctype 0
+    var #r
+      int
+    fun $index_within 2
+    var p
+      type-con $ptr 0
+    fun $ptr 2
+    var T
+      type-con $ctype 0
+    var #r
+      int
+    var T
+      type-con $ctype 0
+axiom 0
+    forall 3 1 3
+      var q
+        type-con $ptr 0
+      var S
+        type-con $state 0
+      var p
+        type-con $ptr 0
+      pat 1
+        fun $set_in 2
+        var q
+          type-con $ptr 0
+        fun $non_null_extent 2
+        var S
+          type-con $state 0
+        var p
+          type-con $ptr 0
+      attribute qid 1
+        string-attr VccPrelu.2142:15
+      attribute uniqueId 1
+        string-attr 251
+      attribute bvZ3Native 1
+        string-attr False
+    =
+    fun $set_in 2
+    var q
+      type-con $ptr 0
+    fun $non_null_extent 2
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    and 2
+    not
+    =
+    fun $ref 1
+    var p
+      type-con $ptr 0
+    fun $ref 1
+    fun $null 0
+    fun $set_in 2
+    var q
+      type-con $ptr 0
+    fun $extent 2
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+axiom 0
+    forall 2 1 3
+      var p
+        type-con $ptr 0
+      var k
+        int
+      pat 1
+        fun $idx 3
+        var p
+          type-con $ptr 0
+        var k
+          int
+        fun $typ 1
+        var p
+          type-con $ptr 0
+      attribute qid 1
+        string-attr VccPrelu.2154:15
+      attribute uniqueId 1
+        string-attr 253
+      attribute bvZ3Native 1
+        string-attr False
+    =
+    fun $index_within 2
+    fun $idx 3
+    var p
+      type-con $ptr 0
+    var k
+      int
+    fun $typ 1
+    var p
+      type-con $ptr 0
+    var p
+      type-con $ptr 0
+    var k
+      int
+axiom 0
+    forall 3 1 3
+      var p
+        type-con $ptr 0
+      var k
+        int
+      var f
+        type-con $field 0
+      pat 1
+        fun $index_within 2
+        fun $dot 2
+        fun $idx 3
+        var p
+          type-con $ptr 0
+        var k
+          int
+        fun $typ 1
+        var p
+          type-con $ptr 0
+        var f
+          type-con $field 0
+        var p
+          type-con $ptr 0
+      attribute qid 1
+        string-attr VccPrelu.2156:15
+      attribute uniqueId 1
+        string-attr 254
+      attribute bvZ3Native 1
+        string-attr False
+    =
+    fun $index_within 2
+    fun $dot 2
+    fun $idx 3
+    var p
+      type-con $ptr 0
+    var k
+      int
+    fun $typ 1
+    var p
+      type-con $ptr 0
+    var f
+      type-con $field 0
+    var p
+      type-con $ptr 0
+    var k
+      int
+axiom 0
+    forall 5 1 3
+      var s1
+        type-con $state 0
+      var s2
+        type-con $state 0
+      var p
+        type-con $ptr 0
+      var t
+        type-con $ctype 0
+      var sz
+        int
+      pat 2
+        fun $state_spans_the_same 4
+        var s1
+          type-con $state 0
+        var s2
+          type-con $state 0
+        var p
+          type-con $ptr 0
+        fun $array 2
+        var t
+          type-con $ctype 0
+        var sz
+          int
+        fun $is_primitive 1
+        var t
+          type-con $ctype 0
+      attribute qid 1
+        string-attr VccPrelu.2168:15
+      attribute uniqueId 1
+        string-attr 256
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    fun $is_primitive 1
+    var t
+      type-con $ctype 0
+    implies
+    fun $state_spans_the_same 4
+    var s1
+      type-con $state 0
+    var s2
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    fun $array 2
+    var t
+      type-con $ctype 0
+    var sz
+      int
+    forall 1 1 3
+      var i
+        int
+      pat 1
+        fun $select.mem 2
+        fun $memory 1
+        var s2
+          type-con $state 0
+        fun $idx 3
+        fun $ptr 2
+        var t
+          type-con $ctype 0
+        fun $ref 1
+        var p
+          type-con $ptr 0
+        var i
+          int
+        var t
+          type-con $ctype 0
+      attribute qid 1
+        string-attr VccPrelu.2171:15
+      attribute uniqueId 1
+        string-attr 255
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 2
+    <=
+    int-num 0
+    var i
+      int
+    <
+    var i
+      int
+    var sz
+      int
+    =
+    fun $select.mem 2
+    fun $memory 1
+    var s1
+      type-con $state 0
+    fun $idx 3
+    fun $ptr 2
+    var t
+      type-con $ctype 0
+    fun $ref 1
+    var p
+      type-con $ptr 0
+    var i
+      int
+    var t
+      type-con $ctype 0
+    fun $select.mem 2
+    fun $memory 1
+    var s2
+      type-con $state 0
+    fun $idx 3
+    fun $ptr 2
+    var t
+      type-con $ctype 0
+    fun $ref 1
+    var p
+      type-con $ptr 0
+    var i
+      int
+    var t
+      type-con $ctype 0
+axiom 0
+    forall 1 1 4
+      var x
+        bool
+      pat 1
+        fun $bool_id 1
+        var x
+          bool
+      attribute qid 1
+        string-attr VccPrelu.2211:31
+      attribute uniqueId 1
+        string-attr 257
+      attribute bvZ3Native 1
+        string-attr False
+      attribute weight 1
+        expr-attr
+          int-num 0
+    =
+    fun $bool_id 1
+    var x
+      bool
+    var x
+      bool
+axiom 0
+    =
+    fun $min.i1 0
+    -
+    int-num 0
+    int-num 128
+axiom 0
+    =
+    fun $max.i1 0
+    int-num 127
+axiom 0
+    =
+    fun $min.i2 0
+    -
+    int-num 0
+    int-num 32768
+axiom 0
+    =
+    fun $max.i2 0
+    int-num 32767
+axiom 0
+    =
+    fun $min.i4 0
+    -
+    int-num 0
+    *
+    int-num 65536
+    int-num 32768
+axiom 0
+    =
+    fun $max.i4 0
+    -
+    *
+    int-num 65536
+    int-num 32768
+    int-num 1
+axiom 0
+    =
+    fun $min.i8 0
+    -
+    int-num 0
+    *
+    *
+    *
+    int-num 65536
+    int-num 65536
+    int-num 65536
+    int-num 32768
+axiom 0
+    =
+    fun $max.i8 0
+    -
+    *
+    *
+    *
+    int-num 65536
+    int-num 65536
+    int-num 65536
+    int-num 32768
+    int-num 1
+axiom 0
+    =
+    fun $max.u1 0
+    int-num 255
+axiom 0
+    =
+    fun $max.u2 0
+    int-num 65535
+axiom 0
+    =
+    fun $max.u4 0
+    -
+    *
+    int-num 65536
+    int-num 65536
+    int-num 1
+axiom 0
+    =
+    fun $max.u8 0
+    -
+    *
+    *
+    *
+    int-num 65536
+    int-num 65536
+    int-num 65536
+    int-num 65536
+    int-num 1
+axiom 0
+    forall 2 1 4
+      var S
+        type-con $state 0
+      var p
+        type-con $ptr 0
+      pat 1
+        fun $read_i1 2
+        var S
+          type-con $state 0
+        var p
+          type-con $ptr 0
+      attribute qid 1
+        string-attr VccPrelu.2253:31
+      attribute uniqueId 1
+        string-attr 258
+      attribute bvZ3Native 1
+        string-attr False
+      attribute weight 1
+        expr-attr
+          int-num 0
+    =
+    fun $read_i1 2
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    fun $select.mem 2
+    fun $memory 1
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+axiom 0
+    forall 2 1 4
+      var S
+        type-con $state 0
+      var p
+        type-con $ptr 0
+      pat 1
+        fun $read_i2 2
+        var S
+          type-con $state 0
+        var p
+          type-con $ptr 0
+      attribute qid 1
+        string-attr VccPrelu.2254:31
+      attribute uniqueId 1
+        string-attr 259
+      attribute bvZ3Native 1
+        string-attr False
+      attribute weight 1
+        expr-attr
+          int-num 0
+    =
+    fun $read_i2 2
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    fun $select.mem 2
+    fun $memory 1
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+axiom 0
+    forall 2 1 4
+      var S
+        type-con $state 0
+      var p
+        type-con $ptr 0
+      pat 1
+        fun $read_i4 2
+        var S
+          type-con $state 0
+        var p
+          type-con $ptr 0
+      attribute qid 1
+        string-attr VccPrelu.2255:31
+      attribute uniqueId 1
+        string-attr 260
+      attribute bvZ3Native 1
+        string-attr False
+      attribute weight 1
+        expr-attr
+          int-num 0
+    =
+    fun $read_i4 2
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    fun $select.mem 2
+    fun $memory 1
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+axiom 0
+    forall 2 1 4
+      var S
+        type-con $state 0
+      var p
+        type-con $ptr 0
+      pat 1
+        fun $read_i8 2
+        var S
+          type-con $state 0
+        var p
+          type-con $ptr 0
+      attribute qid 1
+        string-attr VccPrelu.2256:31
+      attribute uniqueId 1
+        string-attr 261
+      attribute bvZ3Native 1
+        string-attr False
+      attribute weight 1
+        expr-attr
+          int-num 0
+    =
+    fun $read_i8 2
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    fun $select.mem 2
+    fun $memory 1
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+axiom 0
+    forall 2 1 4
+      var S
+        type-con $state 0
+      var p
+        type-con $ptr 0
+      pat 1
+        fun $read_u1 2
+        var S
+          type-con $state 0
+        var p
+          type-con $ptr 0
+      attribute qid 1
+        string-attr VccPrelu.2257:31
+      attribute uniqueId 1
+        string-attr 262
+      attribute bvZ3Native 1
+        string-attr False
+      attribute weight 1
+        expr-attr
+          int-num 0
+    =
+    fun $read_u1 2
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    fun $select.mem 2
+    fun $memory 1
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+axiom 0
+    forall 2 1 4
+      var S
+        type-con $state 0
+      var p
+        type-con $ptr 0
+      pat 1
+        fun $read_u2 2
+        var S
+          type-con $state 0
+        var p
+          type-con $ptr 0
+      attribute qid 1
+        string-attr VccPrelu.2258:31
+      attribute uniqueId 1
+        string-attr 263
+      attribute bvZ3Native 1
+        string-attr False
+      attribute weight 1
+        expr-attr
+          int-num 0
+    =
+    fun $read_u2 2
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    fun $select.mem 2
+    fun $memory 1
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+axiom 0
+    forall 2 1 4
+      var S
+        type-con $state 0
+      var p
+        type-con $ptr 0
+      pat 1
+        fun $read_u4 2
+        var S
+          type-con $state 0
+        var p
+          type-con $ptr 0
+      attribute qid 1
+        string-attr VccPrelu.2259:31
+      attribute uniqueId 1
+        string-attr 264
+      attribute bvZ3Native 1
+        string-attr False
+      attribute weight 1
+        expr-attr
+          int-num 0
+    =
+    fun $read_u4 2
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    fun $select.mem 2
+    fun $memory 1
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+axiom 0
+    forall 2 1 4
+      var S
+        type-con $state 0
+      var p
+        type-con $ptr 0
+      pat 1
+        fun $read_u8 2
+        var S
+          type-con $state 0
+        var p
+          type-con $ptr 0
+      attribute qid 1
+        string-attr VccPrelu.2260:31
+      attribute uniqueId 1
+        string-attr 265
+      attribute bvZ3Native 1
+        string-attr False
+      attribute weight 1
+        expr-attr
+          int-num 0
+    =
+    fun $read_u8 2
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    fun $select.mem 2
+    fun $memory 1
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+axiom 0
+    =
+    fun $ptr_to_u8 1
+    fun $null 0
+    int-num 0
+axiom 0
+    =
+    fun $ptr_to_i8 1
+    fun $null 0
+    int-num 0
+axiom 0
+    =
+    fun $ptr_to_u4 1
+    fun $null 0
+    int-num 0
+axiom 0
+    =
+    fun $ptr_to_i4 1
+    fun $null 0
+    int-num 0
+axiom 0
+    forall 1 1 3
+      var p
+        type-con $ptr 0
+      pat 1
+        fun $ptr_to_u8 1
+        var p
+          type-con $ptr 0
+      attribute qid 1
+        string-attr VccPrelu.2277:15
+      attribute uniqueId 1
+        string-attr 266
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 2
+    <=
+    int-num 0
+    fun $ref 1
+    var p
+      type-con $ptr 0
+    <=
+    fun $ref 1
+    var p
+      type-con $ptr 0
+    fun $max.u8 0
+    =
+    fun $ptr_to_u8 1
+    var p
+      type-con $ptr 0
+    fun $ref 1
+    var p
+      type-con $ptr 0
+axiom 0
+    forall 1 1 3
+      var p
+        type-con $ptr 0
+      pat 1
+        fun $ptr_to_i8 1
+        var p
+          type-con $ptr 0
+      attribute qid 1
+        string-attr VccPrelu.2278:15
+      attribute uniqueId 1
+        string-attr 267
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 2
+    <=
+    fun $min.i8 0
+    fun $ref 1
+    var p
+      type-con $ptr 0
+    <=
+    fun $ref 1
+    var p
+      type-con $ptr 0
+    fun $max.i8 0
+    =
+    fun $ptr_to_i8 1
+    var p
+      type-con $ptr 0
+    fun $ref 1
+    var p
+      type-con $ptr 0
+axiom 0
+    forall 1 1 3
+      var p
+        type-con $ptr 0
+      pat 1
+        fun $ptr_to_u4 1
+        var p
+          type-con $ptr 0
+      attribute qid 1
+        string-attr VccPrelu.2279:15
+      attribute uniqueId 1
+        string-attr 268
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 2
+    <=
+    int-num 0
+    fun $ref 1
+    var p
+      type-con $ptr 0
+    <=
+    fun $ref 1
+    var p
+      type-con $ptr 0
+    fun $max.u4 0
+    =
+    fun $ptr_to_u4 1
+    var p
+      type-con $ptr 0
+    fun $ref 1
+    var p
+      type-con $ptr 0
+axiom 0
+    forall 1 1 3
+      var p
+        type-con $ptr 0
+      pat 1
+        fun $ptr_to_i4 1
+        var p
+          type-con $ptr 0
+      attribute qid 1
+        string-attr VccPrelu.2280:15
+      attribute uniqueId 1
+        string-attr 269
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 2
+    <=
+    fun $min.i4 0
+    fun $ref 1
+    var p
+      type-con $ptr 0
+    <=
+    fun $ref 1
+    var p
+      type-con $ptr 0
+    fun $max.i4 0
+    =
+    fun $ptr_to_i4 1
+    var p
+      type-con $ptr 0
+    fun $ref 1
+    var p
+      type-con $ptr 0
+axiom 0
+    forall 2 1 4
+      var p1
+        type-con $ptr 0
+      var p2
+        type-con $ptr 0
+      pat 1
+        fun $byte_ptr_subtraction 2
+        var p1
+          type-con $ptr 0
+        var p2
+          type-con $ptr 0
+      attribute qid 1
+        string-attr VccPrelu.2282:44
+      attribute uniqueId 1
+        string-attr 270
+      attribute bvZ3Native 1
+        string-attr False
+      attribute weight 1
+        expr-attr
+          int-num 0
+    =
+    fun $byte_ptr_subtraction 2
+    var p1
+      type-con $ptr 0
+    var p2
+      type-con $ptr 0
+    -
+    fun $ref 1
+    var p1
+      type-con $ptr 0
+    fun $ref 1
+    var p2
+      type-con $ptr 0
+axiom 0
+    forall 2 1 3
+      var S
+        type-con $state 0
+      var p
+        type-con $ptr 0
+      pat 1
+        fun $read_i1 2
+        var S
+          type-con $state 0
+        var p
+          type-con $ptr 0
+      attribute qid 1
+        string-attr VccPrelu.2285:15
+      attribute uniqueId 1
+        string-attr 271
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    fun $good_state 1
+    var S
+      type-con $state 0
+    and 2
+    <=
+    fun $min.i1 0
+    fun $read_i1 2
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    <=
+    fun $read_i1 2
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    fun $max.i1 0
+axiom 0
+    forall 2 1 3
+      var S
+        type-con $state 0
+      var p
+        type-con $ptr 0
+      pat 1
+        fun $read_i2 2
+        var S
+          type-con $state 0
+        var p
+          type-con $ptr 0
+      attribute qid 1
+        string-attr VccPrelu.2286:15
+      attribute uniqueId 1
+        string-attr 272
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    fun $good_state 1
+    var S
+      type-con $state 0
+    and 2
+    <=
+    fun $min.i2 0
+    fun $read_i2 2
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    <=
+    fun $read_i2 2
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    fun $max.i2 0
+axiom 0
+    forall 2 1 3
+      var S
+        type-con $state 0
+      var p
+        type-con $ptr 0
+      pat 1
+        fun $read_i4 2
+        var S
+          type-con $state 0
+        var p
+          type-con $ptr 0
+      attribute qid 1
+        string-attr VccPrelu.2287:15
+      attribute uniqueId 1
+        string-attr 273
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    fun $good_state 1
+    var S
+      type-con $state 0
+    and 2
+    <=
+    fun $min.i4 0
+    fun $read_i4 2
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    <=
+    fun $read_i4 2
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    fun $max.i4 0
+axiom 0
+    forall 2 1 3
+      var S
+        type-con $state 0
+      var p
+        type-con $ptr 0
+      pat 1
+        fun $read_i8 2
+        var S
+          type-con $state 0
+        var p
+          type-con $ptr 0
+      attribute qid 1
+        string-attr VccPrelu.2288:15
+      attribute uniqueId 1
+        string-attr 274
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    fun $good_state 1
+    var S
+      type-con $state 0
+    and 2
+    <=
+    fun $min.i8 0
+    fun $read_i8 2
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    <=
+    fun $read_i8 2
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    fun $max.i8 0
+axiom 0
+    forall 2 1 3
+      var S
+        type-con $state 0
+      var p
+        type-con $ptr 0
+      pat 1
+        fun $read_u1 2
+        var S
+          type-con $state 0
+        var p
+          type-con $ptr 0
+      attribute qid 1
+        string-attr VccPrelu.2289:15
+      attribute uniqueId 1
+        string-attr 275
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    fun $good_state 1
+    var S
+      type-con $state 0
+    and 2
+    <=
+    int-num 0
+    fun $read_u1 2
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    <=
+    fun $read_u1 2
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    fun $max.u1 0
+axiom 0
+    forall 2 1 3
+      var S
+        type-con $state 0
+      var p
+        type-con $ptr 0
+      pat 1
+        fun $read_u2 2
+        var S
+          type-con $state 0
+        var p
+          type-con $ptr 0
+      attribute qid 1
+        string-attr VccPrelu.2290:15
+      attribute uniqueId 1
+        string-attr 276
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    fun $good_state 1
+    var S
+      type-con $state 0
+    and 2
+    <=
+    int-num 0
+    fun $read_u2 2
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    <=
+    fun $read_u2 2
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    fun $max.u2 0
+axiom 0
+    forall 2 1 3
+      var S
+        type-con $state 0
+      var p
+        type-con $ptr 0
+      pat 1
+        fun $read_u4 2
+        var S
+          type-con $state 0
+        var p
+          type-con $ptr 0
+      attribute qid 1
+        string-attr VccPrelu.2291:15
+      attribute uniqueId 1
+        string-attr 277
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    fun $good_state 1
+    var S
+      type-con $state 0
+    and 2
+    <=
+    int-num 0
+    fun $read_u4 2
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    <=
+    fun $read_u4 2
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    fun $max.u4 0
+axiom 0
+    forall 2 1 3
+      var S
+        type-con $state 0
+      var p
+        type-con $ptr 0
+      pat 1
+        fun $read_u8 2
+        var S
+          type-con $state 0
+        var p
+          type-con $ptr 0
+      attribute qid 1
+        string-attr VccPrelu.2292:15
+      attribute uniqueId 1
+        string-attr 278
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    fun $good_state 1
+    var S
+      type-con $state 0
+    and 2
+    <=
+    int-num 0
+    fun $read_u8 2
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    <=
+    fun $read_u8 2
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    fun $max.u8 0
+axiom 0
+    and 64
+    =
+    fun $_pow2 1
+    int-num 0
+    int-num 1
+    =
+    fun $_pow2 1
+    int-num 1
+    int-num 2
+    =
+    fun $_pow2 1
+    int-num 2
+    int-num 4
+    =
+    fun $_pow2 1
+    int-num 3
+    int-num 8
+    =
+    fun $_pow2 1
+    int-num 4
+    int-num 16
+    =
+    fun $_pow2 1
+    int-num 5
+    int-num 32
+    =
+    fun $_pow2 1
+    int-num 6
+    int-num 64
+    =
+    fun $_pow2 1
+    int-num 7
+    int-num 128
+    =
+    fun $_pow2 1
+    int-num 8
+    int-num 256
+    =
+    fun $_pow2 1
+    int-num 9
+    int-num 512
+    =
+    fun $_pow2 1
+    int-num 10
+    int-num 1024
+    =
+    fun $_pow2 1
+    int-num 11
+    int-num 2048
+    =
+    fun $_pow2 1
+    int-num 12
+    int-num 4096
+    =
+    fun $_pow2 1
+    int-num 13
+    int-num 8192
+    =
+    fun $_pow2 1
+    int-num 14
+    int-num 16384
+    =
+    fun $_pow2 1
+    int-num 15
+    int-num 32768
+    =
+    fun $_pow2 1
+    int-num 16
+    int-num 65536
+    =
+    fun $_pow2 1
+    int-num 17
+    int-num 131072
+    =
+    fun $_pow2 1
+    int-num 18
+    int-num 262144
+    =
+    fun $_pow2 1
+    int-num 19
+    int-num 524288
+    =
+    fun $_pow2 1
+    int-num 20
+    int-num 1048576
+    =
+    fun $_pow2 1
+    int-num 21
+    int-num 2097152
+    =
+    fun $_pow2 1
+    int-num 22
+    int-num 4194304
+    =
+    fun $_pow2 1
+    int-num 23
+    int-num 8388608
+    =
+    fun $_pow2 1
+    int-num 24
+    int-num 16777216
+    =
+    fun $_pow2 1
+    int-num 25
+    int-num 33554432
+    =
+    fun $_pow2 1
+    int-num 26
+    int-num 67108864
+    =
+    fun $_pow2 1
+    int-num 27
+    int-num 134217728
+    =
+    fun $_pow2 1
+    int-num 28
+    int-num 268435456
+    =
+    fun $_pow2 1
+    int-num 29
+    int-num 536870912
+    =
+    fun $_pow2 1
+    int-num 30
+    int-num 1073741824
+    =
+    fun $_pow2 1
+    int-num 31
+    int-num 2147483648
+    =
+    fun $_pow2 1
+    int-num 32
+    int-num 4294967296
+    =
+    fun $_pow2 1
+    int-num 33
+    int-num 8589934592
+    =
+    fun $_pow2 1
+    int-num 34
+    int-num 17179869184
+    =
+    fun $_pow2 1
+    int-num 35
+    int-num 34359738368
+    =
+    fun $_pow2 1
+    int-num 36
+    int-num 68719476736
+    =
+    fun $_pow2 1
+    int-num 37
+    int-num 137438953472
+    =
+    fun $_pow2 1
+    int-num 38
+    int-num 274877906944
+    =
+    fun $_pow2 1
+    int-num 39
+    int-num 549755813888
+    =
+    fun $_pow2 1
+    int-num 40
+    int-num 1099511627776
+    =
+    fun $_pow2 1
+    int-num 41
+    int-num 2199023255552
+    =
+    fun $_pow2 1
+    int-num 42
+    int-num 4398046511104
+    =
+    fun $_pow2 1
+    int-num 43
+    int-num 8796093022208
+    =
+    fun $_pow2 1
+    int-num 44
+    int-num 17592186044416
+    =
+    fun $_pow2 1
+    int-num 45
+    int-num 35184372088832
+    =
+    fun $_pow2 1
+    int-num 46
+    int-num 70368744177664
+    =
+    fun $_pow2 1
+    int-num 47
+    int-num 140737488355328
+    =
+    fun $_pow2 1
+    int-num 48
+    int-num 281474976710656
+    =
+    fun $_pow2 1
+    int-num 49
+    int-num 562949953421312
+    =
+    fun $_pow2 1
+    int-num 50
+    int-num 1125899906842624
+    =
+    fun $_pow2 1
+    int-num 51
+    int-num 2251799813685248
+    =
+    fun $_pow2 1
+    int-num 52
+    int-num 4503599627370496
+    =
+    fun $_pow2 1
+    int-num 53
+    int-num 9007199254740992
+    =
+    fun $_pow2 1
+    int-num 54
+    int-num 18014398509481984
+    =
+    fun $_pow2 1
+    int-num 55
+    int-num 36028797018963968
+    =
+    fun $_pow2 1
+    int-num 56
+    int-num 72057594037927936
+    =
+    fun $_pow2 1
+    int-num 57
+    int-num 144115188075855872
+    =
+    fun $_pow2 1
+    int-num 58
+    int-num 288230376151711744
+    =
+    fun $_pow2 1
+    int-num 59
+    int-num 576460752303423488
+    =
+    fun $_pow2 1
+    int-num 60
+    int-num 1152921504606846976
+    =
+    fun $_pow2 1
+    int-num 61
+    int-num 2305843009213693952
+    =
+    fun $_pow2 1
+    int-num 62
+    int-num 4611686018427387904
+    =
+    fun $_pow2 1
+    int-num 63
+    int-num 9223372036854775808
+axiom 0
+    forall 3 1 4
+      var t
+        type-con $ctype 0
+      var x
+        int
+      var y
+        int
+      pat 1
+        fun $unchk_add 3
+        var t
+          type-con $ctype 0
+        var x
+          int
+        var y
+          int
+      attribute qid 1
+        string-attr VccPrelu.2318:33
+      attribute uniqueId 1
+        string-attr 279
+      attribute bvZ3Native 1
+        string-attr False
+      attribute weight 1
+        expr-attr
+          int-num 0
+    =
+    fun $unchk_add 3
+    var t
+      type-con $ctype 0
+    var x
+      int
+    var y
+      int
+    fun $unchecked 2
+    var t
+      type-con $ctype 0
+    +
+    var x
+      int
+    var y
+      int
+axiom 0
+    forall 3 1 4
+      var t
+        type-con $ctype 0
+      var x
+        int
+      var y
+        int
+      pat 1
+        fun $unchk_sub 3
+        var t
+          type-con $ctype 0
+        var x
+          int
+        var y
+          int
+      attribute qid 1
+        string-attr VccPrelu.2319:33
+      attribute uniqueId 1
+        string-attr 280
+      attribute bvZ3Native 1
+        string-attr False
+      attribute weight 1
+        expr-attr
+          int-num 0
+    =
+    fun $unchk_sub 3
+    var t
+      type-con $ctype 0
+    var x
+      int
+    var y
+      int
+    fun $unchecked 2
+    var t
+      type-con $ctype 0
+    -
+    var x
+      int
+    var y
+      int
+axiom 0
+    forall 3 1 4
+      var t
+        type-con $ctype 0
+      var x
+        int
+      var y
+        int
+      pat 1
+        fun $unchk_mul 3
+        var t
+          type-con $ctype 0
+        var x
+          int
+        var y
+          int
+      attribute qid 1
+        string-attr VccPrelu.2320:33
+      attribute uniqueId 1
+        string-attr 281
+      attribute bvZ3Native 1
+        string-attr False
+      attribute weight 1
+        expr-attr
+          int-num 0
+    =
+    fun $unchk_mul 3
+    var t
+      type-con $ctype 0
+    var x
+      int
+    var y
+      int
+    fun $unchecked 2
+    var t
+      type-con $ctype 0
+    *
+    var x
+      int
+    var y
+      int
+axiom 0
+    forall 3 1 4
+      var t
+        type-con $ctype 0
+      var x
+        int
+      var y
+        int
+      pat 1
+        fun $unchk_div 3
+        var t
+          type-con $ctype 0
+        var x
+          int
+        var y
+          int
+      attribute qid 1
+        string-attr VccPrelu.2321:33
+      attribute uniqueId 1
+        string-attr 282
+      attribute bvZ3Native 1
+        string-attr False
+      attribute weight 1
+        expr-attr
+          int-num 0
+    =
+    fun $unchk_div 3
+    var t
+      type-con $ctype 0
+    var x
+      int
+    var y
+      int
+    fun $unchecked 2
+    var t
+      type-con $ctype 0
+    /
+    var x
+      int
+    var y
+      int
+axiom 0
+    forall 3 1 4
+      var t
+        type-con $ctype 0
+      var x
+        int
+      var y
+        int
+      pat 1
+        fun $unchk_mod 3
+        var t
+          type-con $ctype 0
+        var x
+          int
+        var y
+          int
+      attribute qid 1
+        string-attr VccPrelu.2322:33
+      attribute uniqueId 1
+        string-attr 283
+      attribute bvZ3Native 1
+        string-attr False
+      attribute weight 1
+        expr-attr
+          int-num 0
+    =
+    fun $unchk_mod 3
+    var t
+      type-con $ctype 0
+    var x
+      int
+    var y
+      int
+    fun $unchecked 2
+    var t
+      type-con $ctype 0
+    %
+    var x
+      int
+    var y
+      int
+axiom 0
+    forall 3 1 4
+      var t
+        type-con $ctype 0
+      var x
+        int
+      var y
+        int
+      pat 1
+        fun $_shl 3
+        var t
+          type-con $ctype 0
+        var x
+          int
+        var y
+          int
+      attribute qid 1
+        string-attr VccPrelu.2324:28
+      attribute uniqueId 1
+        string-attr 284
+      attribute bvZ3Native 1
+        string-attr False
+      attribute weight 1
+        expr-attr
+          int-num 0
+    =
+    fun $_shl 3
+    var t
+      type-con $ctype 0
+    var x
+      int
+    var y
+      int
+    fun $unchecked 2
+    var t
+      type-con $ctype 0
+    *
+    var x
+      int
+    fun $_pow2 1
+    var y
+      int
+axiom 0
+    forall 2 1 4
+      var x
+        int
+      var y
+        int
+      pat 1
+        fun $_shr 2
+        var x
+          int
+        var y
+          int
+      attribute qid 1
+        string-attr VccPrelu.2326:28
+      attribute uniqueId 1
+        string-attr 285
+      attribute bvZ3Native 1
+        string-attr False
+      attribute weight 1
+        expr-attr
+          int-num 0
+    =
+    fun $_shr 2
+    var x
+      int
+    var y
+      int
+    /
+    var x
+      int
+    fun $_pow2 1
+    var y
+      int
+axiom 0
+    forall 5 1 3
+      var x
+        int
+      var from
+        int
+      var to
+        int
+      var xs
+        int
+      var val
+        int
+      pat 1
+        fun $bv_update 5
+        var x
+          int
+        var xs
+          int
+        var from
+          int
+        var to
+          int
+        var val
+          int
+      attribute qid 1
+        string-attr VccPrelu.2333:15
+      attribute uniqueId 1
+        string-attr 286
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 3
+    <=
+    int-num 0
+    var from
+      int
+    <
+    var from
+      int
+    var to
+      int
+    <=
+    var to
+      int
+    var xs
+      int
+    implies
+    and 2
+    <=
+    int-num 0
+    var val
+      int
+    <
+    var val
+      int
+    fun $_pow2 1
+    -
+    var to
+      int
+    var from
+      int
+    and 2
+    <=
+    int-num 0
+    fun $bv_update 5
+    var x
+      int
+    var xs
+      int
+    var from
+      int
+    var to
+      int
+    var val
+      int
+    <
+    fun $bv_update 5
+    var x
+      int
+    var xs
+      int
+    var from
+      int
+    var to
+      int
+    var val
+      int
+    fun $_pow2 1
+    var xs
+      int
+axiom 0
+    forall 3 1 3
+      var from
+        int
+      var to
+        int
+      var xs
+        int
+      pat 1
+        fun $bv_update 5
+        int-num 0
+        var xs
+          int
+        var from
+          int
+        var to
+          int
+        int-num 0
+      attribute qid 1
+        string-attr VccPrelu.2339:15
+      attribute uniqueId 1
+        string-attr 287
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 3
+    <=
+    int-num 0
+    var from
+      int
+    <
+    var from
+      int
+    var to
+      int
+    <=
+    var to
+      int
+    var xs
+      int
+    =
+    fun $bv_update 5
+    int-num 0
+    var xs
+      int
+    var from
+      int
+    var to
+      int
+    int-num 0
+    int-num 0
+axiom 0
+    forall 5 1 3
+      var from
+        int
+      var to
+        int
+      var val
+        int
+      var x
+        int
+      var xs
+        int
+      pat 1
+        fun $bv_extract_signed 4
+        fun $bv_update 5
+        var x
+          int
+        var xs
+          int
+        var from
+          int
+        var to
+          int
+        var val
+          int
+        var xs
+          int
+        var from
+          int
+        var to
+          int
+      attribute qid 1
+        string-attr VccPrelu.2343:15
+      attribute uniqueId 1
+        string-attr 288
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 3
+    <=
+    int-num 0
+    var from
+      int
+    <
+    var from
+      int
+    var to
+      int
+    <=
+    var to
+      int
+    var xs
+      int
+    implies
+    and 2
+    <=
+    -
+    int-num 0
+    fun $_pow2 1
+    -
+    -
+    var to
+      int
+    var from
+      int
+    int-num 1
+    var val
+      int
+    <
+    var val
+      int
+    fun $_pow2 1
+    -
+    -
+    var to
+      int
+    var from
+      int
+    int-num 1
+    =
+    fun $bv_extract_signed 4
+    fun $bv_update 5
+    var x
+      int
+    var xs
+      int
+    var from
+      int
+    var to
+      int
+    var val
+      int
+    var xs
+      int
+    var from
+      int
+    var to
+      int
+    var val
+      int
+axiom 0
+    forall 5 1 3
+      var from
+        int
+      var to
+        int
+      var val
+        int
+      var x
+        int
+      var xs
+        int
+      pat 1
+        fun $bv_extract_unsigned 4
+        fun $bv_update 5
+        var x
+          int
+        var xs
+          int
+        var from
+          int
+        var to
+          int
+        var val
+          int
+        var xs
+          int
+        var from
+          int
+        var to
+          int
+      attribute qid 1
+        string-attr VccPrelu.2349:15
+      attribute uniqueId 1
+        string-attr 289
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 3
+    <=
+    int-num 0
+    var from
+      int
+    <
+    var from
+      int
+    var to
+      int
+    <=
+    var to
+      int
+    var xs
+      int
+    implies
+    and 2
+    <=
+    int-num 0
+    var val
+      int
+    <
+    var val
+      int
+    fun $_pow2 1
+    -
+    var to
+      int
+    var from
+      int
+    =
+    fun $bv_extract_unsigned 4
+    fun $bv_update 5
+    var x
+      int
+    var xs
+      int
+    var from
+      int
+    var to
+      int
+    var val
+      int
+    var xs
+      int
+    var from
+      int
+    var to
+      int
+    var val
+      int
+axiom 0
+    forall 4 1 3
+      var from
+        int
+      var to
+        int
+      var x
+        int
+      var xs
+        int
+      pat 1
+        fun $bv_extract_signed 4
+        var x
+          int
+        var xs
+          int
+        var from
+          int
+        var to
+          int
+      attribute qid 1
+        string-attr VccPrelu.2355:15
+      attribute uniqueId 1
+        string-attr 290
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 3
+    <=
+    int-num 0
+    var from
+      int
+    <
+    var from
+      int
+    var to
+      int
+    <=
+    var to
+      int
+    var xs
+      int
+    and 2
+    <=
+    -
+    int-num 0
+    fun $_pow2 1
+    -
+    -
+    var to
+      int
+    var from
+      int
+    int-num 1
+    fun $bv_extract_signed 4
+    var x
+      int
+    var xs
+      int
+    var from
+      int
+    var to
+      int
+    <=
+    fun $bv_extract_signed 4
+    var x
+      int
+    var xs
+      int
+    var from
+      int
+    var to
+      int
+    -
+    fun $_pow2 1
+    -
+    -
+    var to
+      int
+    var from
+      int
+    int-num 1
+    int-num 1
+axiom 0
+    forall 4 1 3
+      var from
+        int
+      var to
+        int
+      var x
+        int
+      var xs
+        int
+      pat 1
+        fun $bv_extract_unsigned 4
+        var x
+          int
+        var xs
+          int
+        var from
+          int
+        var to
+          int
+      attribute qid 1
+        string-attr VccPrelu.2360:15
+      attribute uniqueId 1
+        string-attr 291
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 3
+    <=
+    int-num 0
+    var from
+      int
+    <
+    var from
+      int
+    var to
+      int
+    <=
+    var to
+      int
+    var xs
+      int
+    and 2
+    <=
+    int-num 0
+    fun $bv_extract_unsigned 4
+    var x
+      int
+    var xs
+      int
+    var from
+      int
+    var to
+      int
+    <=
+    fun $bv_extract_unsigned 4
+    var x
+      int
+    var xs
+      int
+    var from
+      int
+    var to
+      int
+    -
+    fun $_pow2 1
+    -
+    var to
+      int
+    var from
+      int
+    int-num 1
+axiom 0
+    forall 7 1 3
+      var from
+        int
+      var to
+        int
+      var val
+        int
+      var x
+        int
+      var xs
+        int
+      var from2
+        int
+      var to2
+        int
+      pat 1
+        fun $bv_extract_signed 4
+        fun $bv_update 5
+        var x
+          int
+        var xs
+          int
+        var from
+          int
+        var to
+          int
+        var val
+          int
+        var xs
+          int
+        var from2
+          int
+        var to2
+          int
+      attribute qid 1
+        string-attr VccPrelu.2365:15
+      attribute uniqueId 1
+        string-attr 292
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 3
+    <=
+    int-num 0
+    var from
+      int
+    <
+    var from
+      int
+    var to
+      int
+    <=
+    var to
+      int
+    var xs
+      int
+    implies
+    and 3
+    <=
+    int-num 0
+    var from2
+      int
+    <
+    var from2
+      int
+    var to2
+      int
+    <=
+    var to2
+      int
+    var xs
+      int
+    implies
+    or 2
+    <=
+    var to2
+      int
+    var from
+      int
+    <=
+    var to
+      int
+    var from2
+      int
+    =
+    fun $bv_extract_signed 4
+    fun $bv_update 5
+    var x
+      int
+    var xs
+      int
+    var from
+      int
+    var to
+      int
+    var val
+      int
+    var xs
+      int
+    var from2
+      int
+    var to2
+      int
+    fun $bv_extract_signed 4
+    var x
+      int
+    var xs
+      int
+    var from2
+      int
+    var to2
+      int
+axiom 0
+    forall 7 1 3
+      var from
+        int
+      var to
+        int
+      var val
+        int
+      var x
+        int
+      var xs
+        int
+      var from2
+        int
+      var to2
+        int
+      pat 1
+        fun $bv_extract_unsigned 4
+        fun $bv_update 5
+        var x
+          int
+        var xs
+          int
+        var from
+          int
+        var to
+          int
+        var val
+          int
+        var xs
+          int
+        var from2
+          int
+        var to2
+          int
+      attribute qid 1
+        string-attr VccPrelu.2372:15
+      attribute uniqueId 1
+        string-attr 293
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 3
+    <=
+    int-num 0
+    var from
+      int
+    <
+    var from
+      int
+    var to
+      int
+    <=
+    var to
+      int
+    var xs
+      int
+    implies
+    and 3
+    <=
+    int-num 0
+    var from2
+      int
+    <
+    var from2
+      int
+    var to2
+      int
+    <=
+    var to2
+      int
+    var xs
+      int
+    implies
+    or 2
+    <=
+    var to2
+      int
+    var from
+      int
+    <=
+    var to
+      int
+    var from2
+      int
+    =
+    fun $bv_extract_unsigned 4
+    fun $bv_update 5
+    var x
+      int
+    var xs
+      int
+    var from
+      int
+    var to
+      int
+    var val
+      int
+    var xs
+      int
+    var from2
+      int
+    var to2
+      int
+    fun $bv_extract_unsigned 4
+    var x
+      int
+    var xs
+      int
+    var from2
+      int
+    var to2
+      int
+axiom 0
+    forall 3 1 3
+      var from
+        int
+      var to
+        int
+      var xs
+        int
+      pat 1
+        fun $bv_extract_signed 4
+        int-num 0
+        var xs
+          int
+        var from
+          int
+        var to
+          int
+      attribute qid 1
+        string-attr VccPrelu.2379:15
+      attribute uniqueId 1
+        string-attr 294
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 3
+    <=
+    int-num 0
+    var from
+      int
+    <
+    var from
+      int
+    var to
+      int
+    <=
+    var to
+      int
+    var xs
+      int
+    =
+    fun $bv_extract_signed 4
+    int-num 0
+    var xs
+      int
+    var from
+      int
+    var to
+      int
+    int-num 0
+axiom 0
+    forall 3 1 3
+      var from
+        int
+      var to
+        int
+      var xs
+        int
+      pat 1
+        fun $bv_extract_unsigned 4
+        int-num 0
+        var xs
+          int
+        var from
+          int
+        var to
+          int
+      attribute qid 1
+        string-attr VccPrelu.2384:15
+      attribute uniqueId 1
+        string-attr 295
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 3
+    <=
+    int-num 0
+    var from
+      int
+    <
+    var from
+      int
+    var to
+      int
+    <=
+    var to
+      int
+    var xs
+      int
+    =
+    fun $bv_extract_unsigned 4
+    int-num 0
+    var xs
+      int
+    var from
+      int
+    var to
+      int
+    int-num 0
+axiom 0
+    forall 4 1 3
+      var from
+        int
+      var to
+        int
+      var val
+        int
+      var xs
+        int
+      pat 1
+        fun $bv_extract_unsigned 4
+        var val
+          int
+        var xs
+          int
+        var from
+          int
+        var to
+          int
+      attribute qid 1
+        string-attr VccPrelu.2389:15
+      attribute uniqueId 1
+        string-attr 296
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 4
+    <=
+    int-num 0
+    var from
+      int
+    <
+    var from
+      int
+    var to
+      int
+    <=
+    var to
+      int
+    var xs
+      int
+    <=
+    int-num 0
+    var val
+      int
+    =
+    fun $bv_extract_unsigned 4
+    var val
+      int
+    var xs
+      int
+    var from
+      int
+    var to
+      int
+    %
+    /
+    var val
+      int
+    fun $_pow2 1
+    var from
+      int
+    fun $_pow2 1
+    -
+    var to
+      int
+    var from
+      int
+axiom 0
+    forall 4 1 3
+      var from
+        int
+      var to
+        int
+      var val
+        int
+      var xs
+        int
+      pat 1
+        fun $bv_extract_signed 4
+        var val
+          int
+        var xs
+          int
+        var from
+          int
+        var to
+          int
+      attribute qid 1
+        string-attr VccPrelu.2394:15
+      attribute uniqueId 1
+        string-attr 297
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 5
+    <=
+    int-num 0
+    var from
+      int
+    <
+    var from
+      int
+    var to
+      int
+    <=
+    var to
+      int
+    var xs
+      int
+    <=
+    int-num 0
+    var val
+      int
+    <
+    %
+    /
+    var val
+      int
+    fun $_pow2 1
+    var from
+      int
+    fun $_pow2 1
+    -
+    var to
+      int
+    var from
+      int
+    fun $_pow2 1
+    -
+    -
+    var to
+      int
+    var from
+      int
+    int-num 1
+    =
+    fun $bv_extract_signed 4
+    var val
+      int
+    var xs
+      int
+    var from
+      int
+    var to
+      int
+    %
+    /
+    var val
+      int
+    fun $_pow2 1
+    var from
+      int
+    fun $_pow2 1
+    -
+    var to
+      int
+    var from
+      int
+axiom 0
+    forall 4 1 3
+      var from
+        int
+      var to
+        int
+      var val
+        int
+      var xs
+        int
+      pat 1
+        fun $bv_extract_signed 4
+        var val
+          int
+        var xs
+          int
+        var from
+          int
+        var to
+          int
+      attribute qid 1
+        string-attr VccPrelu.2399:15
+      attribute uniqueId 1
+        string-attr 298
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 5
+    <=
+    int-num 0
+    var from
+      int
+    <
+    var from
+      int
+    var to
+      int
+    <=
+    var to
+      int
+    var xs
+      int
+    <=
+    int-num 0
+    var val
+      int
+    >=
+    %
+    /
+    var val
+      int
+    fun $_pow2 1
+    var from
+      int
+    fun $_pow2 1
+    -
+    var to
+      int
+    var from
+      int
+    fun $_pow2 1
+    -
+    -
+    var to
+      int
+    var from
+      int
+    int-num 1
+    =
+    fun $bv_extract_signed 4
+    var val
+      int
+    var xs
+      int
+    var from
+      int
+    var to
+      int
+    -
+    fun $_pow2 1
+    -
+    -
+    var to
+      int
+    var from
+      int
+    int-num 1
+    %
+    /
+    var val
+      int
+    fun $_pow2 1
+    var from
+      int
+    fun $_pow2 1
+    -
+    var to
+      int
+    var from
+      int
+axiom 0
+    forall 1 1 3
+      var val
+        int
+      pat 1
+        fun $in_range_t 2
+        fun ^^i1 0
+        var val
+          int
+      attribute qid 1
+        string-attr VccPrelu.2427:15
+      attribute uniqueId 1
+        string-attr 299
+      attribute bvZ3Native 1
+        string-attr False
+    =
+    fun $in_range_t 2
+    fun ^^i1 0
+    var val
+      int
+    and 2
+    <=
+    fun $min.i1 0
+    var val
+      int
+    <=
+    var val
+      int
+    fun $max.i1 0
+axiom 0
+    forall 1 1 3
+      var val
+        int
+      pat 1
+        fun $in_range_t 2
+        fun ^^i2 0
+        var val
+          int
+      attribute qid 1
+        string-attr VccPrelu.2428:15
+      attribute uniqueId 1
+        string-attr 300
+      attribute bvZ3Native 1
+        string-attr False
+    =
+    fun $in_range_t 2
+    fun ^^i2 0
+    var val
+      int
+    and 2
+    <=
+    fun $min.i2 0
+    var val
+      int
+    <=
+    var val
+      int
+    fun $max.i2 0
+axiom 0
+    forall 1 1 3
+      var val
+        int
+      pat 1
+        fun $in_range_t 2
+        fun ^^i4 0
+        var val
+          int
+      attribute qid 1
+        string-attr VccPrelu.2429:15
+      attribute uniqueId 1
+        string-attr 301
+      attribute bvZ3Native 1
+        string-attr False
+    =
+    fun $in_range_t 2
+    fun ^^i4 0
+    var val
+      int
+    and 2
+    <=
+    fun $min.i4 0
+    var val
+      int
+    <=
+    var val
+      int
+    fun $max.i4 0
+axiom 0
+    forall 1 1 3
+      var val
+        int
+      pat 1
+        fun $in_range_t 2
+        fun ^^i8 0
+        var val
+          int
+      attribute qid 1
+        string-attr VccPrelu.2430:15
+      attribute uniqueId 1
+        string-attr 302
+      attribute bvZ3Native 1
+        string-attr False
+    =
+    fun $in_range_t 2
+    fun ^^i8 0
+    var val
+      int
+    and 2
+    <=
+    fun $min.i8 0
+    var val
+      int
+    <=
+    var val
+      int
+    fun $max.i8 0
+axiom 0
+    forall 1 1 3
+      var val
+        int
+      pat 1
+        fun $in_range_t 2
+        fun ^^u1 0
+        var val
+          int
+      attribute qid 1
+        string-attr VccPrelu.2431:15
+      attribute uniqueId 1
+        string-attr 303
+      attribute bvZ3Native 1
+        string-attr False
+    =
+    fun $in_range_t 2
+    fun ^^u1 0
+    var val
+      int
+    and 2
+    <=
+    int-num 0
+    var val
+      int
+    <=
+    var val
+      int
+    fun $max.u1 0
+axiom 0
+    forall 1 1 3
+      var val
+        int
+      pat 1
+        fun $in_range_t 2
+        fun ^^u2 0
+        var val
+          int
+      attribute qid 1
+        string-attr VccPrelu.2432:15
+      attribute uniqueId 1
+        string-attr 304
+      attribute bvZ3Native 1
+        string-attr False
+    =
+    fun $in_range_t 2
+    fun ^^u2 0
+    var val
+      int
+    and 2
+    <=
+    int-num 0
+    var val
+      int
+    <=
+    var val
+      int
+    fun $max.u2 0
+axiom 0
+    forall 1 1 3
+      var val
+        int
+      pat 1
+        fun $in_range_t 2
+        fun ^^u4 0
+        var val
+          int
+      attribute qid 1
+        string-attr VccPrelu.2433:15
+      attribute uniqueId 1
+        string-attr 305
+      attribute bvZ3Native 1
+        string-attr False
+    =
+    fun $in_range_t 2
+    fun ^^u4 0
+    var val
+      int
+    and 2
+    <=
+    int-num 0
+    var val
+      int
+    <=
+    var val
+      int
+    fun $max.u4 0
+axiom 0
+    forall 1 1 3
+      var val
+        int
+      pat 1
+        fun $in_range_t 2
+        fun ^^u8 0
+        var val
+          int
+      attribute qid 1
+        string-attr VccPrelu.2434:15
+      attribute uniqueId 1
+        string-attr 306
+      attribute bvZ3Native 1
+        string-attr False
+    =
+    fun $in_range_t 2
+    fun ^^u8 0
+    var val
+      int
+    and 2
+    <=
+    int-num 0
+    var val
+      int
+    <=
+    var val
+      int
+    fun $max.u8 0
+axiom 0
+    forall 2 1 3
+      var t
+        type-con $ctype 0
+      var val
+        int
+      pat 1
+        fun $unchecked 2
+        var t
+          type-con $ctype 0
+        var val
+          int
+      attribute qid 1
+        string-attr VccPrelu.2436:15
+      attribute uniqueId 1
+        string-attr 307
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    fun $in_range_t 2
+    var t
+      type-con $ctype 0
+    var val
+      int
+    =
+    fun $unchecked 2
+    var t
+      type-con $ctype 0
+    var val
+      int
+    var val
+      int
+axiom 0
+    forall 2 1 3
+      var t
+        type-con $ctype 0
+      var val
+        int
+      pat 1
+        fun $unchecked 2
+        var t
+          type-con $ctype 0
+        var val
+          int
+      attribute qid 1
+        string-attr VccPrelu.2437:15
+      attribute uniqueId 1
+        string-attr 308
+      attribute bvZ3Native 1
+        string-attr False
+    fun $in_range_t 2
+    var t
+      type-con $ctype 0
+    fun $unchecked 2
+    var t
+      type-con $ctype 0
+    var val
+      int
+axiom 0
+    forall 1 1 3
+      var val
+        int
+      pat 1
+        fun $unchecked 2
+        fun ^^u1 0
+        fun $unchecked 2
+        fun ^^i1 0
+        var val
+          int
+      attribute qid 1
+        string-attr VccPrelu.2439:15
+      attribute uniqueId 1
+        string-attr 309
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 2
+    <=
+    int-num 0
+    var val
+      int
+    <=
+    var val
+      int
+    fun $max.u1 0
+    =
+    fun $unchecked 2
+    fun ^^u1 0
+    fun $unchecked 2
+    fun ^^i1 0
+    var val
+      int
+    var val
+      int
+axiom 0
+    forall 1 1 3
+      var val
+        int
+      pat 1
+        fun $unchecked 2
+        fun ^^u2 0
+        fun $unchecked 2
+        fun ^^i2 0
+        var val
+          int
+      attribute qid 1
+        string-attr VccPrelu.2440:15
+      attribute uniqueId 1
+        string-attr 310
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 2
+    <=
+    int-num 0
+    var val
+      int
+    <=
+    var val
+      int
+    fun $max.u2 0
+    =
+    fun $unchecked 2
+    fun ^^u2 0
+    fun $unchecked 2
+    fun ^^i2 0
+    var val
+      int
+    var val
+      int
+axiom 0
+    forall 1 1 3
+      var val
+        int
+      pat 1
+        fun $unchecked 2
+        fun ^^u4 0
+        fun $unchecked 2
+        fun ^^i4 0
+        var val
+          int
+      attribute qid 1
+        string-attr VccPrelu.2441:15
+      attribute uniqueId 1
+        string-attr 311
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 2
+    <=
+    int-num 0
+    var val
+      int
+    <=
+    var val
+      int
+    fun $max.u4 0
+    =
+    fun $unchecked 2
+    fun ^^u4 0
+    fun $unchecked 2
+    fun ^^i4 0
+    var val
+      int
+    var val
+      int
+axiom 0
+    forall 1 1 3
+      var val
+        int
+      pat 1
+        fun $unchecked 2
+        fun ^^u8 0
+        fun $unchecked 2
+        fun ^^i8 0
+        var val
+          int
+      attribute qid 1
+        string-attr VccPrelu.2442:15
+      attribute uniqueId 1
+        string-attr 312
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 2
+    <=
+    int-num 0
+    var val
+      int
+    <=
+    var val
+      int
+    fun $max.u8 0
+    =
+    fun $unchecked 2
+    fun ^^u8 0
+    fun $unchecked 2
+    fun ^^i8 0
+    var val
+      int
+    var val
+      int
+axiom 0
+    forall 1 1 3
+      var val
+        int
+      pat 1
+        fun $unchecked 2
+        fun ^^i1 0
+        fun $unchecked 2
+        fun ^^u1 0
+        var val
+          int
+      attribute qid 1
+        string-attr VccPrelu.2443:15
+      attribute uniqueId 1
+        string-attr 313
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 2
+    <=
+    fun $min.i1 0
+    var val
+      int
+    <=
+    var val
+      int
+    fun $max.i1 0
+    =
+    fun $unchecked 2
+    fun ^^i1 0
+    fun $unchecked 2
+    fun ^^u1 0
+    var val
+      int
+    var val
+      int
+axiom 0
+    forall 1 1 3
+      var val
+        int
+      pat 1
+        fun $unchecked 2
+        fun ^^i2 0
+        fun $unchecked 2
+        fun ^^u2 0
+        var val
+          int
+      attribute qid 1
+        string-attr VccPrelu.2444:15
+      attribute uniqueId 1
+        string-attr 314
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 2
+    <=
+    fun $min.i2 0
+    var val
+      int
+    <=
+    var val
+      int
+    fun $max.i2 0
+    =
+    fun $unchecked 2
+    fun ^^i2 0
+    fun $unchecked 2
+    fun ^^u2 0
+    var val
+      int
+    var val
+      int
+axiom 0
+    forall 1 1 3
+      var val
+        int
+      pat 1
+        fun $unchecked 2
+        fun ^^i4 0
+        fun $unchecked 2
+        fun ^^u4 0
+        var val
+          int
+      attribute qid 1
+        string-attr VccPrelu.2445:15
+      attribute uniqueId 1
+        string-attr 315
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 2
+    <=
+    fun $min.i4 0
+    var val
+      int
+    <=
+    var val
+      int
+    fun $max.i4 0
+    =
+    fun $unchecked 2
+    fun ^^i4 0
+    fun $unchecked 2
+    fun ^^u4 0
+    var val
+      int
+    var val
+      int
+axiom 0
+    forall 1 1 3
+      var val
+        int
+      pat 1
+        fun $unchecked 2
+        fun ^^i8 0
+        fun $unchecked 2
+        fun ^^u8 0
+        var val
+          int
+      attribute qid 1
+        string-attr VccPrelu.2446:15
+      attribute uniqueId 1
+        string-attr 316
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 2
+    <=
+    fun $min.i8 0
+    var val
+      int
+    <=
+    var val
+      int
+    fun $max.i8 0
+    =
+    fun $unchecked 2
+    fun ^^i8 0
+    fun $unchecked 2
+    fun ^^u8 0
+    var val
+      int
+    var val
+      int
+axiom 0
+    forall 4 1 3
+      var t
+        type-con $ctype 0
+      var x
+        int
+      var y
+        int
+      var z
+        int
+      pat 2
+        %
+        var x
+          int
+        fun $_pow2 1
+        var y
+          int
+        fun $_and 3
+        var t
+          type-con $ctype 0
+        var x
+          int
+        var z
+          int
+      attribute qid 1
+        string-attr VccPrelu.2452:15
+      attribute uniqueId 1
+        string-attr 317
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 3
+    fun $in_range_t 2
+    var t
+      type-con $ctype 0
+    var x
+      int
+    fun $in_range_t 2
+    var t
+      type-con $ctype 0
+    -
+    fun $_pow2 1
+    var y
+      int
+    int-num 1
+    >=
+    var x
+      int
+    int-num 0
+    =
+    %
+    var x
+      int
+    fun $_pow2 1
+    var y
+      int
+    fun $_and 3
+    var t
+      type-con $ctype 0
+    var x
+      int
+    -
+    fun $_pow2 1
+    var y
+      int
+    int-num 1
+axiom 0
+    forall 2 1 3
+      var i
+        int
+      var j
+        int
+      pat 1
+        /
+        var i
+          int
+        var j
+          int
+      attribute qid 1
+        string-attr VccPrelu.2458:15
+      attribute uniqueId 1
+        string-attr 318
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 2
+    <=
+    int-num 0
+    var i
+      int
+    <
+    int-num 0
+    var j
+      int
+    <=
+    /
+    var i
+      int
+    var j
+      int
+    var i
+      int
+axiom 0
+    forall 2 1 3
+      var i
+        int
+      var j
+        int
+      pat 1
+        /
+        var i
+          int
+        var j
+          int
+      attribute qid 1
+        string-attr VccPrelu.2460:15
+      attribute uniqueId 1
+        string-attr 319
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 2
+    >
+    var i
+      int
+    int-num 0
+    >
+    var j
+      int
+    int-num 0
+    and 2
+    <
+    -
+    var i
+      int
+    var j
+      int
+    *
+    /
+    var i
+      int
+    var j
+      int
+    var j
+      int
+    <=
+    *
+    /
+    var i
+      int
+    var j
+      int
+    var j
+      int
+    var i
+      int
+axiom 0
+    forall 1 1 3
+      var i
+        int
+      pat 1
+        /
+        var i
+          int
+        var i
+          int
+      attribute qid 1
+        string-attr VccPrelu.2461:15
+      attribute uniqueId 1
+        string-attr 320
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    not
+    =
+    var i
+      int
+    int-num 0
+    =
+    /
+    var i
+      int
+    var i
+      int
+    int-num 1
+axiom 0
+    forall 2 2 3
+      var x
+        int
+      var y
+        int
+      pat 1
+        %
+        var x
+          int
+        var y
+          int
+      pat 1
+        /
+        var x
+          int
+        var y
+          int
+      attribute qid 1
+        string-attr VccPrelu.2464:15
+      attribute uniqueId 1
+        string-attr 321
+      attribute bvZ3Native 1
+        string-attr False
+    =
+    %
+    var x
+      int
+    var y
+      int
+    -
+    var x
+      int
+    *
+    /
+    var x
+      int
+    var y
+      int
+    var y
+      int
+axiom 0
+    forall 2 1 3
+      var x
+        int
+      var y
+        int
+      pat 1
+        %
+        var x
+          int
+        var y
+          int
+      attribute qid 1
+        string-attr VccPrelu.2465:15
+      attribute uniqueId 1
+        string-attr 322
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 2
+    <=
+    int-num 0
+    var x
+      int
+    <
+    int-num 0
+    var y
+      int
+    and 2
+    <=
+    int-num 0
+    %
+    var x
+      int
+    var y
+      int
+    <
+    %
+    var x
+      int
+    var y
+      int
+    var y
+      int
+axiom 0
+    forall 2 1 3
+      var x
+        int
+      var y
+        int
+      pat 1
+        %
+        var x
+          int
+        var y
+          int
+      attribute qid 1
+        string-attr VccPrelu.2466:15
+      attribute uniqueId 1
+        string-attr 323
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 2
+    <=
+    int-num 0
+    var x
+      int
+    <
+    var y
+      int
+    int-num 0
+    and 2
+    <=
+    int-num 0
+    %
+    var x
+      int
+    var y
+      int
+    <
+    %
+    var x
+      int
+    var y
+      int
+    -
+    int-num 0
+    var y
+      int
+axiom 0
+    forall 2 1 3
+      var x
+        int
+      var y
+        int
+      pat 1
+        %
+        var x
+          int
+        var y
+          int
+      attribute qid 1
+        string-attr VccPrelu.2467:15
+      attribute uniqueId 1
+        string-attr 324
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 2
+    <=
+    var x
+      int
+    int-num 0
+    <
+    int-num 0
+    var y
+      int
+    and 2
+    <
+    -
+    int-num 0
+    var y
+      int
+    %
+    var x
+      int
+    var y
+      int
+    <=
+    %
+    var x
+      int
+    var y
+      int
+    int-num 0
+axiom 0
+    forall 2 1 3
+      var x
+        int
+      var y
+        int
+      pat 1
+        %
+        var x
+          int
+        var y
+          int
+      attribute qid 1
+        string-attr VccPrelu.2468:15
+      attribute uniqueId 1
+        string-attr 325
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 2
+    <=
+    var x
+      int
+    int-num 0
+    <
+    var y
+      int
+    int-num 0
+    and 2
+    <
+    var y
+      int
+    %
+    var x
+      int
+    var y
+      int
+    <=
+    %
+    var x
+      int
+    var y
+      int
+    int-num 0
+axiom 0
+    forall 3 1 3
+      var t
+        type-con $ctype 0
+      var x
+        int
+      var y
+        int
+      pat 1
+        fun $_and 3
+        var t
+          type-con $ctype 0
+        var x
+          int
+        var y
+          int
+      attribute qid 1
+        string-attr VccPrelu.2477:15
+      attribute uniqueId 1
+        string-attr 326
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 2
+    <=
+    int-num 0
+    var x
+      int
+    fun $in_range_t 2
+    var t
+      type-con $ctype 0
+    var x
+      int
+    and 2
+    <=
+    int-num 0
+    fun $_and 3
+    var t
+      type-con $ctype 0
+    var x
+      int
+    var y
+      int
+    <=
+    fun $_and 3
+    var t
+      type-con $ctype 0
+    var x
+      int
+    var y
+      int
+    var x
+      int
+axiom 0
+    forall 3 1 3
+      var t
+        type-con $ctype 0
+      var x
+        int
+      var y
+        int
+      pat 1
+        fun $_and 3
+        var t
+          type-con $ctype 0
+        var x
+          int
+        var y
+          int
+      attribute qid 1
+        string-attr VccPrelu.2478:15
+      attribute uniqueId 1
+        string-attr 327
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 4
+    <=
+    int-num 0
+    var x
+      int
+    <=
+    int-num 0
+    var y
+      int
+    fun $in_range_t 2
+    var t
+      type-con $ctype 0
+    var x
+      int
+    fun $in_range_t 2
+    var t
+      type-con $ctype 0
+    var y
+      int
+    and 2
+    <=
+    fun $_and 3
+    var t
+      type-con $ctype 0
+    var x
+      int
+    var y
+      int
+    var x
+      int
+    <=
+    fun $_and 3
+    var t
+      type-con $ctype 0
+    var x
+      int
+    var y
+      int
+    var y
+      int
+axiom 0
+    forall 3 1 3
+      var t
+        type-con $ctype 0
+      var x
+        int
+      var y
+        int
+      pat 1
+        fun $_or 3
+        var t
+          type-con $ctype 0
+        var x
+          int
+        var y
+          int
+      attribute qid 1
+        string-attr VccPrelu.2479:15
+      attribute uniqueId 1
+        string-attr 328
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 4
+    <=
+    int-num 0
+    var x
+      int
+    <=
+    int-num 0
+    var y
+      int
+    fun $in_range_t 2
+    var t
+      type-con $ctype 0
+    var x
+      int
+    fun $in_range_t 2
+    var t
+      type-con $ctype 0
+    var y
+      int
+    and 2
+    <=
+    int-num 0
+    fun $_or 3
+    var t
+      type-con $ctype 0
+    var x
+      int
+    var y
+      int
+    <=
+    fun $_or 3
+    var t
+      type-con $ctype 0
+    var x
+      int
+    var y
+      int
+    +
+    var x
+      int
+    var y
+      int
+axiom 0
+    forall 3 1 3
+      var t
+        type-con $ctype 0
+      var x
+        int
+      var y
+        int
+      pat 1
+        fun $_or 3
+        var t
+          type-con $ctype 0
+        var x
+          int
+        var y
+          int
+      attribute qid 1
+        string-attr VccPrelu.2480:15
+      attribute uniqueId 1
+        string-attr 329
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 4
+    <=
+    int-num 0
+    var x
+      int
+    <=
+    int-num 0
+    var y
+      int
+    fun $in_range_t 2
+    var t
+      type-con $ctype 0
+    var x
+      int
+    fun $in_range_t 2
+    var t
+      type-con $ctype 0
+    var y
+      int
+    and 2
+    <=
+    var x
+      int
+    fun $_or 3
+    var t
+      type-con $ctype 0
+    var x
+      int
+    var y
+      int
+    <=
+    var y
+      int
+    fun $_or 3
+    var t
+      type-con $ctype 0
+    var x
+      int
+    var y
+      int
+axiom 0
+    forall 4 1 3
+      var t
+        type-con $ctype 0
+      var x
+        int
+      var y
+        int
+      var z
+        int
+      pat 2
+        fun $_or 3
+        var t
+          type-con $ctype 0
+        var x
+          int
+        var y
+          int
+        fun $_pow2 1
+        var z
+          int
+      attribute qid 1
+        string-attr VccPrelu.2481:15
+      attribute uniqueId 1
+        string-attr 330
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 8
+    <=
+    int-num 0
+    var x
+      int
+    <=
+    int-num 0
+    var y
+      int
+    <=
+    int-num 0
+    var z
+      int
+    <
+    var z
+      int
+    int-num 64
+    <
+    var x
+      int
+    fun $_pow2 1
+    var z
+      int
+    <
+    var y
+      int
+    fun $_pow2 1
+    var z
+      int
+    fun $in_range_t 2
+    var t
+      type-con $ctype 0
+    var x
+      int
+    fun $in_range_t 2
+    var t
+      type-con $ctype 0
+    var y
+      int
+    <
+    fun $_or 3
+    var t
+      type-con $ctype 0
+    var x
+      int
+    var y
+      int
+    fun $_pow2 1
+    var z
+      int
+axiom 0
+    forall 3 1 3
+      var t
+        type-con $ctype 0
+      var x
+        int
+      var y
+        int
+      pat 1
+        fun $_or 3
+        var t
+          type-con $ctype 0
+        var x
+          int
+        var y
+          int
+      attribute qid 1
+        string-attr VccPrelu.2484:15
+      attribute uniqueId 1
+        string-attr 331
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 4
+    <=
+    int-num 0
+    var x
+      int
+    <=
+    var x
+      int
+    fun $max.u1 0
+    <=
+    int-num 0
+    var y
+      int
+    <=
+    var y
+      int
+    fun $max.u1 0
+    and 2
+    <=
+    int-num 0
+    fun $_or 3
+    var t
+      type-con $ctype 0
+    var x
+      int
+    var y
+      int
+    <=
+    fun $_or 3
+    var t
+      type-con $ctype 0
+    var x
+      int
+    var y
+      int
+    fun $max.u1 0
+axiom 0
+    forall 3 1 3
+      var t
+        type-con $ctype 0
+      var x
+        int
+      var y
+        int
+      pat 1
+        fun $_or 3
+        var t
+          type-con $ctype 0
+        var x
+          int
+        var y
+          int
+      attribute qid 1
+        string-attr VccPrelu.2485:15
+      attribute uniqueId 1
+        string-attr 332
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 4
+    <=
+    int-num 0
+    var x
+      int
+    <=
+    var x
+      int
+    fun $max.u2 0
+    <=
+    int-num 0
+    var y
+      int
+    <=
+    var y
+      int
+    fun $max.u2 0
+    and 2
+    <=
+    int-num 0
+    fun $_or 3
+    var t
+      type-con $ctype 0
+    var x
+      int
+    var y
+      int
+    <=
+    fun $_or 3
+    var t
+      type-con $ctype 0
+    var x
+      int
+    var y
+      int
+    fun $max.u2 0
+axiom 0
+    forall 3 1 3
+      var t
+        type-con $ctype 0
+      var x
+        int
+      var y
+        int
+      pat 1
+        fun $_or 3
+        var t
+          type-con $ctype 0
+        var x
+          int
+        var y
+          int
+      attribute qid 1
+        string-attr VccPrelu.2486:15
+      attribute uniqueId 1
+        string-attr 333
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 4
+    <=
+    int-num 0
+    var x
+      int
+    <=
+    var x
+      int
+    fun $max.u4 0
+    <=
+    int-num 0
+    var y
+      int
+    <=
+    var y
+      int
+    fun $max.u4 0
+    and 2
+    <=
+    int-num 0
+    fun $_or 3
+    var t
+      type-con $ctype 0
+    var x
+      int
+    var y
+      int
+    <=
+    fun $_or 3
+    var t
+      type-con $ctype 0
+    var x
+      int
+    var y
+      int
+    fun $max.u4 0
+axiom 0
+    forall 3 1 3
+      var t
+        type-con $ctype 0
+      var x
+        int
+      var y
+        int
+      pat 1
+        fun $_or 3
+        var t
+          type-con $ctype 0
+        var x
+          int
+        var y
+          int
+      attribute qid 1
+        string-attr VccPrelu.2487:15
+      attribute uniqueId 1
+        string-attr 334
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 4
+    <=
+    int-num 0
+    var x
+      int
+    <=
+    var x
+      int
+    fun $max.u8 0
+    <=
+    int-num 0
+    var y
+      int
+    <=
+    var y
+      int
+    fun $max.u8 0
+    and 2
+    <=
+    int-num 0
+    fun $_or 3
+    var t
+      type-con $ctype 0
+    var x
+      int
+    var y
+      int
+    <=
+    fun $_or 3
+    var t
+      type-con $ctype 0
+    var x
+      int
+    var y
+      int
+    fun $max.u8 0
+axiom 0
+    forall 3 1 3
+      var t
+        type-con $ctype 0
+      var x
+        int
+      var y
+        int
+      pat 1
+        fun $_and 3
+        var t
+          type-con $ctype 0
+        var x
+          int
+        var y
+          int
+      attribute qid 1
+        string-attr VccPrelu.2489:15
+      attribute uniqueId 1
+        string-attr 335
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 4
+    <=
+    int-num 0
+    var x
+      int
+    <=
+    var x
+      int
+    fun $max.u1 0
+    <=
+    int-num 0
+    var y
+      int
+    <=
+    var y
+      int
+    fun $max.u1 0
+    and 2
+    <=
+    int-num 0
+    fun $_and 3
+    var t
+      type-con $ctype 0
+    var x
+      int
+    var y
+      int
+    <=
+    fun $_and 3
+    var t
+      type-con $ctype 0
+    var x
+      int
+    var y
+      int
+    fun $max.u1 0
+axiom 0
+    forall 3 1 3
+      var t
+        type-con $ctype 0
+      var x
+        int
+      var y
+        int
+      pat 1
+        fun $_and 3
+        var t
+          type-con $ctype 0
+        var x
+          int
+        var y
+          int
+      attribute qid 1
+        string-attr VccPrelu.2490:15
+      attribute uniqueId 1
+        string-attr 336
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 4
+    <=
+    int-num 0
+    var x
+      int
+    <=
+    var x
+      int
+    fun $max.u2 0
+    <=
+    int-num 0
+    var y
+      int
+    <=
+    var y
+      int
+    fun $max.u2 0
+    and 2
+    <=
+    int-num 0
+    fun $_and 3
+    var t
+      type-con $ctype 0
+    var x
+      int
+    var y
+      int
+    <=
+    fun $_and 3
+    var t
+      type-con $ctype 0
+    var x
+      int
+    var y
+      int
+    fun $max.u2 0
+axiom 0
+    forall 3 1 3
+      var t
+        type-con $ctype 0
+      var x
+        int
+      var y
+        int
+      pat 1
+        fun $_and 3
+        var t
+          type-con $ctype 0
+        var x
+          int
+        var y
+          int
+      attribute qid 1
+        string-attr VccPrelu.2491:15
+      attribute uniqueId 1
+        string-attr 337
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 4
+    <=
+    int-num 0
+    var x
+      int
+    <=
+    var x
+      int
+    fun $max.u4 0
+    <=
+    int-num 0
+    var y
+      int
+    <=
+    var y
+      int
+    fun $max.u4 0
+    and 2
+    <=
+    int-num 0
+    fun $_and 3
+    var t
+      type-con $ctype 0
+    var x
+      int
+    var y
+      int
+    <=
+    fun $_and 3
+    var t
+      type-con $ctype 0
+    var x
+      int
+    var y
+      int
+    fun $max.u4 0
+axiom 0
+    forall 3 1 3
+      var t
+        type-con $ctype 0
+      var x
+        int
+      var y
+        int
+      pat 1
+        fun $_and 3
+        var t
+          type-con $ctype 0
+        var x
+          int
+        var y
+          int
+      attribute qid 1
+        string-attr VccPrelu.2492:15
+      attribute uniqueId 1
+        string-attr 338
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 4
+    <=
+    int-num 0
+    var x
+      int
+    <=
+    var x
+      int
+    fun $max.u8 0
+    <=
+    int-num 0
+    var y
+      int
+    <=
+    var y
+      int
+    fun $max.u8 0
+    and 2
+    <=
+    int-num 0
+    fun $_and 3
+    var t
+      type-con $ctype 0
+    var x
+      int
+    var y
+      int
+    <=
+    fun $_and 3
+    var t
+      type-con $ctype 0
+    var x
+      int
+    var y
+      int
+    fun $max.u8 0
+axiom 0
+    forall 3 1 3
+      var t
+        type-con $ctype 0
+      var x
+        int
+      var y
+        int
+      pat 1
+        fun $_xor 3
+        var t
+          type-con $ctype 0
+        var x
+          int
+        var y
+          int
+      attribute qid 1
+        string-attr VccPrelu.2494:15
+      attribute uniqueId 1
+        string-attr 339
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 4
+    <=
+    int-num 0
+    var x
+      int
+    <=
+    var x
+      int
+    fun $max.u1 0
+    <=
+    int-num 0
+    var y
+      int
+    <=
+    var y
+      int
+    fun $max.u1 0
+    and 2
+    <=
+    int-num 0
+    fun $_xor 3
+    var t
+      type-con $ctype 0
+    var x
+      int
+    var y
+      int
+    <=
+    fun $_xor 3
+    var t
+      type-con $ctype 0
+    var x
+      int
+    var y
+      int
+    fun $max.u1 0
+axiom 0
+    forall 3 1 3
+      var t
+        type-con $ctype 0
+      var x
+        int
+      var y
+        int
+      pat 1
+        fun $_xor 3
+        var t
+          type-con $ctype 0
+        var x
+          int
+        var y
+          int
+      attribute qid 1
+        string-attr VccPrelu.2495:15
+      attribute uniqueId 1
+        string-attr 340
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 4
+    <=
+    int-num 0
+    var x
+      int
+    <=
+    var x
+      int
+    fun $max.u2 0
+    <=
+    int-num 0
+    var y
+      int
+    <=
+    var y
+      int
+    fun $max.u2 0
+    and 2
+    <=
+    int-num 0
+    fun $_xor 3
+    var t
+      type-con $ctype 0
+    var x
+      int
+    var y
+      int
+    <=
+    fun $_xor 3
+    var t
+      type-con $ctype 0
+    var x
+      int
+    var y
+      int
+    fun $max.u2 0
+axiom 0
+    forall 3 1 3
+      var t
+        type-con $ctype 0
+      var x
+        int
+      var y
+        int
+      pat 1
+        fun $_xor 3
+        var t
+          type-con $ctype 0
+        var x
+          int
+        var y
+          int
+      attribute qid 1
+        string-attr VccPrelu.2496:15
+      attribute uniqueId 1
+        string-attr 341
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 4
+    <=
+    int-num 0
+    var x
+      int
+    <=
+    var x
+      int
+    fun $max.u4 0
+    <=
+    int-num 0
+    var y
+      int
+    <=
+    var y
+      int
+    fun $max.u4 0
+    and 2
+    <=
+    int-num 0
+    fun $_xor 3
+    var t
+      type-con $ctype 0
+    var x
+      int
+    var y
+      int
+    <=
+    fun $_xor 3
+    var t
+      type-con $ctype 0
+    var x
+      int
+    var y
+      int
+    fun $max.u4 0
+axiom 0
+    forall 3 1 3
+      var t
+        type-con $ctype 0
+      var x
+        int
+      var y
+        int
+      pat 1
+        fun $_xor 3
+        var t
+          type-con $ctype 0
+        var x
+          int
+        var y
+          int
+      attribute qid 1
+        string-attr VccPrelu.2497:15
+      attribute uniqueId 1
+        string-attr 342
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 4
+    <=
+    int-num 0
+    var x
+      int
+    <=
+    var x
+      int
+    fun $max.u8 0
+    <=
+    int-num 0
+    var y
+      int
+    <=
+    var y
+      int
+    fun $max.u8 0
+    and 2
+    <=
+    int-num 0
+    fun $_xor 3
+    var t
+      type-con $ctype 0
+    var x
+      int
+    var y
+      int
+    <=
+    fun $_xor 3
+    var t
+      type-con $ctype 0
+    var x
+      int
+    var y
+      int
+    fun $max.u8 0
+axiom 0
+    forall 2 1 3
+      var t
+        type-con $ctype 0
+      var x
+        int
+      pat 1
+        fun $_not 2
+        var t
+          type-con $ctype 0
+        var x
+          int
+      attribute qid 1
+        string-attr VccPrelu.2499:15
+      attribute uniqueId 1
+        string-attr 343
+      attribute bvZ3Native 1
+        string-attr False
+    fun $in_range_t 2
+    var t
+      type-con $ctype 0
+    fun $_not 2
+    var t
+      type-con $ctype 0
+    var x
+      int
+axiom 0
+    forall 2 1 3
+      var t
+        type-con $ctype 0
+      var x
+        int
+      pat 1
+        fun $_or 3
+        var t
+          type-con $ctype 0
+        var x
+          int
+        fun $_not 2
+        var t
+          type-con $ctype 0
+        var x
+          int
+      attribute qid 1
+        string-attr VccPrelu.2504:15
+      attribute uniqueId 1
+        string-attr 344
+      attribute bvZ3Native 1
+        string-attr False
+    =
+    fun $_or 3
+    var t
+      type-con $ctype 0
+    var x
+      int
+    fun $_not 2
+    var t
+      type-con $ctype 0
+    var x
+      int
+    fun $_not 2
+    var t
+      type-con $ctype 0
+    int-num 0
+axiom 0
+    forall 2 1 3
+      var t
+        type-con $ctype 0
+      var x
+        int
+      pat 1
+        fun $_and 3
+        var t
+          type-con $ctype 0
+        var x
+          int
+        fun $_not 2
+        var t
+          type-con $ctype 0
+        var x
+          int
+      attribute qid 1
+        string-attr VccPrelu.2505:15
+      attribute uniqueId 1
+        string-attr 345
+      attribute bvZ3Native 1
+        string-attr False
+    =
+    fun $_and 3
+    var t
+      type-con $ctype 0
+    var x
+      int
+    fun $_not 2
+    var t
+      type-con $ctype 0
+    var x
+      int
+    int-num 0
+axiom 0
+    forall 2 1 3
+      var t
+        type-con $ctype 0
+      var x
+        int
+      pat 1
+        fun $_or 3
+        var t
+          type-con $ctype 0
+        var x
+          int
+        int-num 0
+      attribute qid 1
+        string-attr VccPrelu.2506:15
+      attribute uniqueId 1
+        string-attr 346
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    fun $in_range_t 2
+    var t
+      type-con $ctype 0
+    var x
+      int
+    =
+    fun $_or 3
+    var t
+      type-con $ctype 0
+    var x
+      int
+    int-num 0
+    var x
+      int
+axiom 0
+    forall 2 1 3
+      var t
+        type-con $ctype 0
+      var x
+        int
+      pat 1
+        fun $_or 3
+        var t
+          type-con $ctype 0
+        var x
+          int
+        fun $_not 2
+        var t
+          type-con $ctype 0
+        int-num 0
+      attribute qid 1
+        string-attr VccPrelu.2507:15
+      attribute uniqueId 1
+        string-attr 347
+      attribute bvZ3Native 1
+        string-attr False
+    =
+    fun $_or 3
+    var t
+      type-con $ctype 0
+    var x
+      int
+    fun $_not 2
+    var t
+      type-con $ctype 0
+    int-num 0
+    fun $_not 2
+    var t
+      type-con $ctype 0
+    int-num 0
+axiom 0
+    forall 2 1 3
+      var t
+        type-con $ctype 0
+      var x
+        int
+      pat 1
+        fun $_or 3
+        var t
+          type-con $ctype 0
+        var x
+          int
+        var x
+          int
+      attribute qid 1
+        string-attr VccPrelu.2508:15
+      attribute uniqueId 1
+        string-attr 348
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    fun $in_range_t 2
+    var t
+      type-con $ctype 0
+    var x
+      int
+    =
+    fun $_or 3
+    var t
+      type-con $ctype 0
+    var x
+      int
+    var x
+      int
+    var x
+      int
+axiom 0
+    forall 2 1 3
+      var t
+        type-con $ctype 0
+      var x
+        int
+      pat 1
+        fun $_and 3
+        var t
+          type-con $ctype 0
+        var x
+          int
+        int-num 0
+      attribute qid 1
+        string-attr VccPrelu.2509:15
+      attribute uniqueId 1
+        string-attr 349
+      attribute bvZ3Native 1
+        string-attr False
+    =
+    fun $_and 3
+    var t
+      type-con $ctype 0
+    var x
+      int
+    int-num 0
+    int-num 0
+axiom 0
+    forall 2 1 3
+      var t
+        type-con $ctype 0
+      var x
+        int
+      pat 1
+        fun $_and 3
+        var t
+          type-con $ctype 0
+        var x
+          int
+        fun $_not 2
+        var t
+          type-con $ctype 0
+        int-num 0
+      attribute qid 1
+        string-attr VccPrelu.2510:15
+      attribute uniqueId 1
+        string-attr 350
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    fun $in_range_t 2
+    var t
+      type-con $ctype 0
+    var x
+      int
+    =
+    fun $_and 3
+    var t
+      type-con $ctype 0
+    var x
+      int
+    fun $_not 2
+    var t
+      type-con $ctype 0
+    int-num 0
+    var x
+      int
+axiom 0
+    forall 2 1 3
+      var t
+        type-con $ctype 0
+      var x
+        int
+      pat 1
+        fun $_and 3
+        var t
+          type-con $ctype 0
+        var x
+          int
+        var x
+          int
+      attribute qid 1
+        string-attr VccPrelu.2511:15
+      attribute uniqueId 1
+        string-attr 351
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    fun $in_range_t 2
+    var t
+      type-con $ctype 0
+    var x
+      int
+    =
+    fun $_and 3
+    var t
+      type-con $ctype 0
+    var x
+      int
+    var x
+      int
+    var x
+      int
+axiom 0
+    forall 3 1 3
+      var t
+        type-con $ctype 0
+      var x
+        int
+      var y
+        int
+      pat 1
+        fun $_and 3
+        var t
+          type-con $ctype 0
+        fun $_or 3
+        var t
+          type-con $ctype 0
+        var x
+          int
+        var y
+          int
+        var y
+          int
+      attribute qid 1
+        string-attr VccPrelu.2512:15
+      attribute uniqueId 1
+        string-attr 352
+      attribute bvZ3Native 1
+        string-attr False
+    =
+    fun $_and 3
+    var t
+      type-con $ctype 0
+    fun $_or 3
+    var t
+      type-con $ctype 0
+    var x
+      int
+    var y
+      int
+    var y
+      int
+    var y
+      int
+axiom 0
+    forall 3 1 3
+      var t
+        type-con $ctype 0
+      var x
+        int
+      var y
+        int
+      pat 1
+        fun $_and 3
+        var t
+          type-con $ctype 0
+        fun $_or 3
+        var t
+          type-con $ctype 0
+        var x
+          int
+        var y
+          int
+        var x
+          int
+      attribute qid 1
+        string-attr VccPrelu.2513:15
+      attribute uniqueId 1
+        string-attr 353
+      attribute bvZ3Native 1
+        string-attr False
+    =
+    fun $_and 3
+    var t
+      type-con $ctype 0
+    fun $_or 3
+    var t
+      type-con $ctype 0
+    var x
+      int
+    var y
+      int
+    var x
+      int
+    var x
+      int
+axiom 0
+    forall 2 1 3
+      var t
+        type-con $ctype 0
+      var x
+        int
+      pat 1
+        fun $_xor 3
+        var t
+          type-con $ctype 0
+        var x
+          int
+        int-num 0
+      attribute qid 1
+        string-attr VccPrelu.2514:15
+      attribute uniqueId 1
+        string-attr 354
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    fun $in_range_t 2
+    var t
+      type-con $ctype 0
+    var x
+      int
+    =
+    fun $_xor 3
+    var t
+      type-con $ctype 0
+    var x
+      int
+    int-num 0
+    var x
+      int
+axiom 0
+    forall 2 1 3
+      var t
+        type-con $ctype 0
+      var x
+        int
+      pat 1
+        fun $_xor 3
+        var t
+          type-con $ctype 0
+        var x
+          int
+        var x
+          int
+      attribute qid 1
+        string-attr VccPrelu.2515:15
+      attribute uniqueId 1
+        string-attr 355
+      attribute bvZ3Native 1
+        string-attr False
+    =
+    fun $_xor 3
+    var t
+      type-con $ctype 0
+    var x
+      int
+    var x
+      int
+    int-num 0
+axiom 0
+    forall 2 1 3
+      var t
+        type-con $ctype 0
+      var x
+        int
+      pat 1
+        fun $_xor 3
+        var t
+          type-con $ctype 0
+        var x
+          int
+        fun $_not 2
+        var t
+          type-con $ctype 0
+        int-num 0
+      attribute qid 1
+        string-attr VccPrelu.2516:15
+      attribute uniqueId 1
+        string-attr 356
+      attribute bvZ3Native 1
+        string-attr False
+    =
+    fun $_xor 3
+    var t
+      type-con $ctype 0
+    var x
+      int
+    fun $_not 2
+    var t
+      type-con $ctype 0
+    int-num 0
+    fun $_not 2
+    var t
+      type-con $ctype 0
+    var x
+      int
+axiom 0
+    forall 2 1 3
+      var t
+        type-con $ctype 0
+      var x
+        int
+      pat 1
+        fun $_not 2
+        var t
+          type-con $ctype 0
+        fun $_not 2
+        var t
+          type-con $ctype 0
+        var x
+          int
+      attribute qid 1
+        string-attr VccPrelu.2517:15
+      attribute uniqueId 1
+        string-attr 357
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    fun $in_range_t 2
+    var t
+      type-con $ctype 0
+    var x
+      int
+    =
+    fun $_not 2
+    var t
+      type-con $ctype 0
+    fun $_not 2
+    var t
+      type-con $ctype 0
+    var x
+      int
+    var x
+      int
+axiom 0
+    forall 3 1 3
+      var t
+        type-con $ctype 0
+      var x
+        int
+      var y
+        int
+      pat 1
+        fun $_or 3
+        var t
+          type-con $ctype 0
+        var x
+          int
+        var y
+          int
+      attribute qid 1
+        string-attr VccPrelu.2518:15
+      attribute uniqueId 1
+        string-attr 358
+      attribute bvZ3Native 1
+        string-attr False
+    =
+    fun $_or 3
+    var t
+      type-con $ctype 0
+    var x
+      int
+    var y
+      int
+    fun $_or 3
+    var t
+      type-con $ctype 0
+    var y
+      int
+    var x
+      int
+axiom 0
+    forall 3 1 3
+      var t
+        type-con $ctype 0
+      var x
+        int
+      var y
+        int
+      pat 1
+        fun $_xor 3
+        var t
+          type-con $ctype 0
+        var x
+          int
+        var y
+          int
+      attribute qid 1
+        string-attr VccPrelu.2519:15
+      attribute uniqueId 1
+        string-attr 359
+      attribute bvZ3Native 1
+        string-attr False
+    =
+    fun $_xor 3
+    var t
+      type-con $ctype 0
+    var x
+      int
+    var y
+      int
+    fun $_xor 3
+    var t
+      type-con $ctype 0
+    var y
+      int
+    var x
+      int
+axiom 0
+    forall 3 1 3
+      var t
+        type-con $ctype 0
+      var x
+        int
+      var y
+        int
+      pat 1
+        fun $_and 3
+        var t
+          type-con $ctype 0
+        var x
+          int
+        var y
+          int
+      attribute qid 1
+        string-attr VccPrelu.2520:15
+      attribute uniqueId 1
+        string-attr 360
+      attribute bvZ3Native 1
+        string-attr False
+    =
+    fun $_and 3
+    var t
+      type-con $ctype 0
+    var x
+      int
+    var y
+      int
+    fun $_and 3
+    var t
+      type-con $ctype 0
+    var y
+      int
+    var x
+      int
+axiom 0
+    forall 2 1 4
+      var x
+        int
+      var y
+        int
+      pat 1
+        fun $_mul 2
+        var x
+          int
+        var y
+          int
+      attribute qid 1
+        string-attr VccPrelu.2524:28
+      attribute uniqueId 1
+        string-attr 361
+      attribute bvZ3Native 1
+        string-attr False
+      attribute weight 1
+        expr-attr
+          int-num 0
+    =
+    fun $_mul 2
+    var x
+      int
+    var y
+      int
+    *
+    var x
+      int
+    var y
+      int
+axiom 0
+    forall 2 1 3
+      var id
+        int
+      var length
+        int
+      pat 1
+        fun $get_string_literal 2
+        var id
+          int
+        var length
+          int
+      attribute qid 1
+        string-attr VccPrelu.2531:15
+      attribute uniqueId 1
+        string-attr 362
+      attribute bvZ3Native 1
+        string-attr False
+    fun $is 2
+    fun $get_string_literal 2
+    var id
+      int
+    var length
+      int
+    fun ^^u1 0
+axiom 0
+    forall 3 2 3
+      var id
+        int
+      var length
+        int
+      var S
+        type-con $state 0
+      pat 1
+        fun $typed 2
+        var S
+          type-con $state 0
+        fun $get_string_literal 2
+        var id
+          int
+        var length
+          int
+      pat 1
+        fun $is_array 4
+        var S
+          type-con $state 0
+        fun $get_string_literal 2
+        var id
+          int
+        var length
+          int
+        fun ^^u1 0
+        var length
+          int
+      attribute qid 1
+        string-attr VccPrelu.2532:15
+      attribute uniqueId 1
+        string-attr 363
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    fun $good_state 1
+    var S
+      type-con $state 0
+    and 2
+    fun $typed 2
+    var S
+      type-con $state 0
+    fun $get_string_literal 2
+    var id
+      int
+    var length
+      int
+    forall 1 2 3
+      var i
+        int
+      pat 1
+        fun $select.sm 2
+        fun $statusmap 1
+        var S
+          type-con $state 0
+        fun $idx 3
+        fun $get_string_literal 2
+        var id
+          int
+        var length
+          int
+        var i
+          int
+        fun ^^u1 0
+      pat 1
+        fun $select.tm 2
+        fun $typemap 1
+        var S
+          type-con $state 0
+        fun $idx 3
+        fun $get_string_literal 2
+        var id
+          int
+        var length
+          int
+        var i
+          int
+        fun ^^u1 0
+      attribute qid 1
+        string-attr VccPrelu.2043:13
+      attribute uniqueId 1
+        string-attr 236
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 2
+    <=
+    int-num 0
+    var i
+      int
+    <
+    var i
+      int
+    var length
+      int
+    and 3
+    fun $ts_is_array_elt 1
+    fun $select.tm 2
+    fun $typemap 1
+    var S
+      type-con $state 0
+    fun $idx 3
+    fun $get_string_literal 2
+    var id
+      int
+    var length
+      int
+    var i
+      int
+    fun ^^u1 0
+    fun $is 2
+    fun $idx 3
+    fun $get_string_literal 2
+    var id
+      int
+    var length
+      int
+    var i
+      int
+    fun ^^u1 0
+    fun ^^u1 0
+    fun $thread_local 2
+    var S
+      type-con $state 0
+    fun $idx 3
+    fun $get_string_literal 2
+    var id
+      int
+    var length
+      int
+    var i
+      int
+    fun ^^u1 0
+axiom 0
+    forall 2 1 3
+      var no
+        int
+      var t
+        type-con $ctype 0
+      pat 1
+        fun $get_fnptr 2
+        var no
+          int
+        var t
+          type-con $ctype 0
+      attribute qid 1
+        string-attr VccPrelu.2543:21
+      attribute uniqueId 1
+        string-attr 364
+      attribute bvZ3Native 1
+        string-attr False
+    =
+    fun $get_fnptr 2
+    var no
+      int
+    var t
+      type-con $ctype 0
+    fun $ptr 2
+    var t
+      type-con $ctype 0
+    fun $get_fnptr_ref 1
+    var no
+      int
+axiom 0
+    forall 1 0 3
+      var no
+        int
+      attribute qid 1
+        string-attr VccPrelu.2550:15
+      attribute uniqueId 1
+        string-attr 365
+      attribute bvZ3Native 1
+        string-attr False
+    =
+    fun $get_fnptr_inv 1
+    fun $get_fnptr_ref 1
+    var no
+      int
+    var no
+      int
+axiom 0
+    forall 3 2 3
+      var S
+        type-con $state 0
+      var no
+        int
+      var t
+        type-con $ctype 0
+      pat 1
+        fun $select.tm 2
+        fun $typemap 1
+        var S
+          type-con $state 0
+        fun $get_fnptr 2
+        var no
+          int
+        var t
+          type-con $ctype 0
+      pat 1
+        fun $select.sm 2
+        fun $statusmap 1
+        var S
+          type-con $state 0
+        fun $get_fnptr 2
+        var no
+          int
+        var t
+          type-con $ctype 0
+      attribute qid 1
+        string-attr VccPrelu.2553:15
+      attribute uniqueId 1
+        string-attr 366
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 2
+    fun $is_fnptr_type 1
+    var t
+      type-con $ctype 0
+    fun $good_state 1
+    var S
+      type-con $state 0
+    fun $mutable 2
+    var S
+      type-con $state 0
+    fun $get_fnptr 2
+    var no
+      int
+    var t
+      type-con $ctype 0
+axiom 0
+    forall 1 1 3
+      var t
+        type-con $ctype 0
+      pat 1
+        fun $is_math_type 1
+        var t
+          type-con $ctype 0
+      attribute qid 1
+        string-attr VccPrelu.2560:15
+      attribute uniqueId 1
+        string-attr 367
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    fun $is_math_type 1
+    var t
+      type-con $ctype 0
+    fun $is_primitive 1
+    var t
+      type-con $ctype 0
+axiom 0
+    forall 1 1 3
+      var t
+        type-con $ctype 0
+      pat 1
+        fun $is_fnptr_type 1
+        var t
+          type-con $ctype 0
+      attribute qid 1
+        string-attr VccPrelu.2561:15
+      attribute uniqueId 1
+        string-attr 368
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    fun $is_fnptr_type 1
+    var t
+      type-con $ctype 0
+    fun $is_primitive 1
+    var t
+      type-con $ctype 0
+axiom 0
+    forall 2 1 3
+      var S
+        type-con $state 0
+      var c
+        type-con $ptr 0
+      pat 2
+        fun $full_stop 1
+        var S
+          type-con $state 0
+        fun $valid_claim 2
+        var S
+          type-con $state 0
+        var c
+          type-con $ptr 0
+      attribute qid 1
+        string-attr VccPrelu.2572:15
+      attribute uniqueId 1
+        string-attr 369
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 2
+    fun $full_stop 1
+    var S
+      type-con $state 0
+    fun $closed 2
+    var S
+      type-con $state 0
+    var c
+      type-con $ptr 0
+    fun $valid_claim 2
+    var S
+      type-con $state 0
+    var c
+      type-con $ptr 0
+axiom 0
+    forall 2 1 3
+      var S
+        type-con $state 0
+      var c
+        type-con $ptr 0
+      pat 1
+        fun $valid_claim 2
+        var S
+          type-con $state 0
+        var c
+          type-con $ptr 0
+      attribute qid 1
+        string-attr VccPrelu.2575:15
+      attribute uniqueId 1
+        string-attr 370
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    fun $valid_claim 2
+    var S
+      type-con $state 0
+    var c
+      type-con $ptr 0
+    and 2
+    fun $closed 2
+    var S
+      type-con $state 0
+    var c
+      type-con $ptr 0
+    fun $invok_state 1
+    var S
+      type-con $state 0
+axiom 0
+    forall 2 1 3
+      var c1
+        type-con $ptr 0
+      var c2
+        type-con $ptr 0
+      pat 1
+        fun $claims_claim 2
+        var c1
+          type-con $ptr 0
+        var c2
+          type-con $ptr 0
+      attribute qid 1
+        string-attr VccPrelu.2600:15
+      attribute uniqueId 1
+        string-attr 373
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 3
+    fun $is 2
+    var c1
+      type-con $ptr 0
+    fun ^^claim 0
+    fun $is 2
+    var c2
+      type-con $ptr 0
+    fun ^^claim 0
+    forall 1 0 3
+      var S
+        type-con $state 0
+      attribute qid 1
+        string-attr VccPrelu.2602:11
+      attribute uniqueId 1
+        string-attr 372
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    fun $valid_claim 2
+    var S
+      type-con $state 0
+    var c1
+      type-con $ptr 0
+    fun $closed 2
+    var S
+      type-con $state 0
+    var c2
+      type-con $ptr 0
+    fun $claims_claim 2
+    var c1
+      type-con $ptr 0
+    var c2
+      type-con $ptr 0
+axiom 0
+    forall 3 1 3
+      var S
+        type-con $state 0
+      var c1
+        type-con $ptr 0
+      var c2
+        type-con $ptr 0
+      pat 2
+        fun $valid_claim 2
+        var S
+          type-con $state 0
+        var c1
+          type-con $ptr 0
+        fun $claims_claim 2
+        var c1
+          type-con $ptr 0
+        var c2
+          type-con $ptr 0
+      attribute qid 1
+        string-attr VccPrelu.2606:15
+      attribute uniqueId 1
+        string-attr 374
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 2
+    fun $valid_claim 2
+    var S
+      type-con $state 0
+    var c1
+      type-con $ptr 0
+    fun $claims_claim 2
+    var c1
+      type-con $ptr 0
+    var c2
+      type-con $ptr 0
+    fun $valid_claim 2
+    var S
+      type-con $state 0
+    var c2
+      type-con $ptr 0
+axiom 0
+    forall 3 1 3
+      var S
+        type-con $state 0
+      var c
+        type-con $ptr 0
+      var o
+        type-con $ptr 0
+      pat 2
+        fun $closed 2
+        var S
+          type-con $state 0
+        var c
+          type-con $ptr 0
+        fun $claims_obj 2
+        var c
+          type-con $ptr 0
+        var o
+          type-con $ptr 0
+      attribute qid 1
+        string-attr VccPrelu.2609:15
+      attribute uniqueId 1
+        string-attr 375
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    fun $good_state 1
+    var S
+      type-con $state 0
+    implies
+    and 2
+    fun $claims_obj 2
+    var c
+      type-con $ptr 0
+    var o
+      type-con $ptr 0
+    fun $closed 2
+    var S
+      type-con $state 0
+    var c
+      type-con $ptr 0
+    and 3
+    fun $instantiate_ptrset 1
+    fun $owns 2
+    var S
+      type-con $state 0
+    var o
+      type-con $ptr 0
+    fun $closed 2
+    var S
+      type-con $state 0
+    var o
+      type-con $ptr 0
+    >
+    fun $ref_cnt 2
+    var S
+      type-con $state 0
+    var o
+      type-con $ptr 0
+    int-num 0
+axiom 0
+    forall 3 1 3
+      var S
+        type-con $state 0
+      var c
+        type-con $ptr 0
+      var o
+        type-con $ptr 0
+      pat 2
+        fun $valid_claim 2
+        var S
+          type-con $state 0
+        var c
+          type-con $ptr 0
+        fun $claims_obj 2
+        var c
+          type-con $ptr 0
+        var o
+          type-con $ptr 0
+      attribute qid 1
+        string-attr VccPrelu.2614:15
+      attribute uniqueId 1
+        string-attr 376
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 2
+    fun $valid_claim 2
+    var S
+      type-con $state 0
+    var c
+      type-con $ptr 0
+    fun $claims_obj 2
+    var c
+      type-con $ptr 0
+    var o
+      type-con $ptr 0
+    fun $inv2 4
+    var S
+      type-con $state 0
+    var S
+      type-con $state 0
+    var o
+      type-con $ptr 0
+    fun $typ 1
+    var o
+      type-con $ptr 0
+axiom 0
+    forall 3 1 3
+      var S
+        type-con $state 0
+      var c
+        type-con $ptr 0
+      var r
+        int
+      pat 2
+        fun $valid_claim 2
+        var S
+          type-con $state 0
+        var c
+          type-con $ptr 0
+        fun $claims_obj 2
+        var c
+          type-con $ptr 0
+        fun $ptr 2
+        fun ^^claim 0
+        var r
+          int
+      attribute qid 1
+        string-attr VccPrelu.2618:15
+      attribute uniqueId 1
+        string-attr 377
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 2
+    fun $valid_claim 2
+    var S
+      type-con $state 0
+    var c
+      type-con $ptr 0
+    fun $claims_obj 2
+    var c
+      type-con $ptr 0
+    fun $ptr 2
+    fun ^^claim 0
+    var r
+      int
+    fun $valid_claim 2
+    var S
+      type-con $state 0
+    fun $ptr 2
+    fun ^^claim 0
+    var r
+      int
+axiom 0
+    forall 2 1 4
+      var S
+        type-con $state 0
+      var p
+        type-con $ptr 0
+      pat 1
+        fun $not_shared 2
+        var S
+          type-con $state 0
+        var p
+          type-con $ptr 0
+      attribute qid 1
+        string-attr VccPrelu.2623:34
+      attribute uniqueId 1
+        string-attr 378
+      attribute bvZ3Native 1
+        string-attr False
+      attribute weight 1
+        expr-attr
+          int-num 0
+    =
+    fun $not_shared 2
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    and 7
+    fun $closed 2
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    =
+    fun $owner 2
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    fun $me 0
+    fun $is 2
+    var p
+      type-con $ptr 0
+    fun $typ 1
+    var p
+      type-con $ptr 0
+    fun $typed 2
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    not
+    =
+    fun $kind_of 1
+    fun $typ 1
+    var p
+      type-con $ptr 0
+    fun $kind_primitive 0
+    fun $is_non_primitive 1
+    fun $typ 1
+    var p
+      type-con $ptr 0
+    or 2
+    not
+    fun $is_claimable 1
+    fun $typ 1
+    var p
+      type-con $ptr 0
+    =
+    fun $ref_cnt 2
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    int-num 0
+axiom 0
+    forall 2 1 4
+      var s
+        type-con $state 0
+      var p
+        type-con $ptr 0
+      pat 1
+        fun $claimed_closed 2
+        var s
+          type-con $state 0
+        var p
+          type-con $ptr 0
+      attribute qid 1
+        string-attr VccPrelu.2626:38
+      attribute uniqueId 1
+        string-attr 379
+      attribute bvZ3Native 1
+        string-attr False
+      attribute weight 1
+        expr-attr
+          int-num 0
+    =
+    fun $claimed_closed 2
+    var s
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    fun $closed 2
+    var s
+      type-con $state 0
+    var p
+      type-con $ptr 0
+axiom 0
+    forall 2 1 3
+      var S
+        type-con $state 0
+      var p
+        type-con $ptr 0
+      pat 2
+        fun $invok_state 1
+        var S
+          type-con $state 0
+        fun $claimed_closed 2
+        var S
+          type-con $state 0
+        var p
+          type-con $ptr 0
+      attribute qid 1
+        string-attr VccPrelu.2629:15
+      attribute uniqueId 1
+        string-attr 380
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 2
+    fun $invok_state 1
+    var S
+      type-con $state 0
+    fun $claimed_closed 2
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    fun $inv2 4
+    var S
+      type-con $state 0
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    fun $typ 1
+    var p
+      type-con $ptr 0
+axiom 0
+    =
+    fun $no_claim 0
+    fun $ptr 2
+    fun ^^claim 0
+    int-num 0
+axiom 0
+    forall 2 1 4
+      var S
+        type-con $state 0
+      var p
+        type-con $ptr 0
+      pat 1
+        fun $ref_cnt 2
+        var S
+          type-con $state 0
+        var p
+          type-con $ptr 0
+      attribute qid 1
+        string-attr VccPrelu.2685:31
+      attribute uniqueId 1
+        string-attr 388
+      attribute bvZ3Native 1
+        string-attr False
+      attribute weight 1
+        expr-attr
+          int-num 0
+    =
+    fun $ref_cnt 2
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    fun $st_ref_cnt 1
+    fun $select.sm 2
+    fun $statusmap 1
+    var S
+      type-con $state 0
+    var p
+      type-con $ptr 0
+axiom 0
+    fun $is_claimable 1
+    fun ^^claim 0
+axiom 0
+    forall 1 0 3
+      var p
+        type-con $ptrset 0
+      attribute qid 1
+        string-attr VccPrelu.2729:15
+      attribute uniqueId 1
+        string-attr 390
+      attribute bvZ3Native 1
+        string-attr False
+    =
+    fun $int_to_ptrset 1
+    fun $ptrset_to_int 1
+    var p
+      type-con $ptrset 0
+    var p
+      type-con $ptrset 0
+axiom 0
+    forall 1 0 3
+      var p
+        type-con $version 0
+      attribute qid 1
+        string-attr VccPrelu.2733:15
+      attribute uniqueId 1
+        string-attr 391
+      attribute bvZ3Native 1
+        string-attr False
+    =
+    fun $int_to_version 1
+    fun $version_to_int 1
+    var p
+      type-con $version 0
+    var p
+      type-con $version 0
+axiom 0
+    forall 1 0 3
+      var p
+        type-con $vol_version 0
+      attribute qid 1
+        string-attr VccPrelu.2737:15
+      attribute uniqueId 1
+        string-attr 392
+      attribute bvZ3Native 1
+        string-attr False
+    =
+    fun $int_to_vol_version 1
+    fun $vol_version_to_int 1
+    var p
+      type-con $vol_version 0
+    var p
+      type-con $vol_version 0
+axiom 0
+    forall 1 0 3
+      var p
+        type-con $ptr 0
+      attribute qid 1
+        string-attr VccPrelu.2741:15
+      attribute uniqueId 1
+        string-attr 393
+      attribute bvZ3Native 1
+        string-attr False
+    =
+    fun $int_to_ptr 1
+    fun $ptr_to_int 1
+    var p
+      type-con $ptr 0
+    var p
+      type-con $ptr 0
+axiom 0
+    forall 3 1 3
+      var S1
+        type-con $state 0
+      var S2
+        type-con $state 0
+      var W
+        type-con $ptrset 0
+      pat 1
+        fun $updated_only_values 3
+        var S1
+          type-con $state 0
+        var S2
+          type-con $state 0
+        var W
+          type-con $ptrset 0
+      attribute qid 1
+        string-attr VccPrelu.2752:15
+      attribute uniqueId 1
+        string-attr 395
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    forall 1 1 3
+      var p
+        type-con $ptr 0
+      pat 1
+        fun $dont_instantiate 1
+        var p
+          type-con $ptr 0
+      attribute qid 1
+        string-attr VccPrelu.2754:11
+      attribute uniqueId 1
+        string-attr 394
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    or 2
+    fun $is_primitive 1
+    fun $typ 1
+    var p
+      type-con $ptr 0
+    fun $is_non_primitive 1
+    fun $typ 1
+    var p
+      type-con $ptr 0
+    implies
+    and 2
+    fun $typed 2
+    var S1
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    not
+    or 2
+    not
+    =
+    fun $owner 2
+    var S1
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    fun $me 0
+    and 2
+    =
+    fun $kind_of 1
+    fun $typ 1
+    var p
+      type-con $ptr 0
+    fun $kind_primitive 0
+    fun $closed 2
+    var S1
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    or 2
+    =
+    fun $select.mem 2
+    fun $memory 1
+    var S1
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    fun $select.mem 2
+    fun $memory 1
+    var S2
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    fun $set_in 2
+    var p
+      type-con $ptr 0
+    var W
+      type-con $ptrset 0
+    fun $updated_only_values 3
+    var S1
+      type-con $state 0
+    var S2
+      type-con $state 0
+    var W
+      type-con $ptrset 0
+axiom 0
+    forall 3 1 3
+      var S1
+        type-con $state 0
+      var S2
+        type-con $state 0
+      var W
+        type-con $ptrset 0
+      pat 1
+        fun $updated_only_domains 3
+        var S1
+          type-con $state 0
+        var S2
+          type-con $state 0
+        var W
+          type-con $ptrset 0
+      attribute qid 1
+        string-attr VccPrelu.2759:15
+      attribute uniqueId 1
+        string-attr 397
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    forall 1 1 3
+      var p
+        type-con $ptr 0
+      pat 1
+        fun $dont_instantiate 1
+        var p
+          type-con $ptr 0
+      attribute qid 1
+        string-attr VccPrelu.2761:11
+      attribute uniqueId 1
+        string-attr 396
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 2
+    fun $set_in 2
+    var p
+      type-con $ptr 0
+    var W
+      type-con $ptrset 0
+    not
+    =
+    fun $kind_of 1
+    fun $typ 1
+    var p
+      type-con $ptr 0
+    fun $kind_primitive 0
+    or 2
+    =
+    fun $select.mem 2
+    fun $memory 1
+    var S1
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    fun $select.mem 2
+    fun $memory 1
+    var S2
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    fun $domain_updated_at 4
+    var S1
+      type-con $state 0
+    var S2
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    var W
+      type-con $ptrset 0
+    fun $updated_only_domains 3
+    var S1
+      type-con $state 0
+    var S2
+      type-con $state 0
+    var W
+      type-con $ptrset 0
+axiom 0
+    forall 4 1 3
+      var S1
+        type-con $state 0
+      var S2
+        type-con $state 0
+      var p
+        type-con $ptr 0
+      var W
+        type-con $ptrset 0
+      pat 1
+        fun $domain_updated_at 4
+        var S1
+          type-con $state 0
+        var S2
+          type-con $state 0
+        var p
+          type-con $ptr 0
+        var W
+          type-con $ptrset 0
+      attribute qid 1
+        string-attr VccPrelu.2777:29
+      attribute uniqueId 1
+        string-attr 399
+      attribute bvZ3Native 1
+        string-attr False
+    =
+    fun $domain_updated_at 4
+    var S1
+      type-con $state 0
+    var S2
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    var W
+      type-con $ptrset 0
+    and 2
+    forall 1 1 3
+      var q
+        type-con $ptr 0
+      pat 1
+        fun $fetch_from_domain 2
+        fun $read_version 2
+        var S2
+          type-con $state 0
+        var p
+          type-con $ptr 0
+        var q
+          type-con $ptr 0
+      attribute qid 1
+        string-attr VccPrelu.2778:13
+      attribute uniqueId 1
+        string-attr 398
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 2
+    =
+    fun $kind_of 1
+    fun $typ 1
+    var q
+      type-con $ptr 0
+    fun $kind_primitive 0
+    not
+    fun $set_in 2
+    var q
+      type-con $ptr 0
+    var W
+      type-con $ptrset 0
+    =
+    fun $fetch_from_domain 2
+    fun $read_version 2
+    var S1
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    var q
+      type-con $ptr 0
+    fun $fetch_from_domain 2
+    fun $read_version 2
+    var S2
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    var q
+      type-con $ptr 0
+    =
+    fun $domain 2
+    var S1
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    fun $domain 2
+    var S2
+      type-con $state 0
+    var p
+      type-con $ptr 0
+axiom 0
+    =
+    fun #distTp1 0
+    fun $ptr_to 1
+    fun ^^u1 0
+axiom 0
+    fun $type_code_is 2
+    int-num 1
+    fun ^^u4 0
+axiom 0
+    fun $file_name_is 2
+    int-num 1
+    fun #file^Z?3A?5CC?5Cmax.c 0
+var-decl $s 0
+    type-con $state 0
+vc maximum 1
+    implies
+    label pos 0 0
+    true
+    implies
+    and 2
+    <=
+    int-num 0
+    var L#max
+      int
+    <=
+    var L#max
+      int
+    fun $max.u1 0
+    implies
+    and 2
+    <=
+    int-num 0
+    var L#p
+      int
+    <=
+    var L#p
+      int
+    fun $max.u4 0
+    implies
+    and 2
+    <=
+    int-num 0
+    var SL#witness
+      int
+    <=
+    var SL#witness
+      int
+    fun $max.u4 0
+    implies
+    <
+    var P#len
+      int
+    int-num 1099511627776
+    implies
+    <
+    int-num 0
+    var P#len
+      int
+    implies
+    and 6
+    fun $closed 2
+    var $s
+      type-con $state 0
+    fun $ptr 2
+    fun $array 2
+    fun ^^u1 0
+    var P#len
+      int
+    fun $ref 1
+    fun $ptr 2
+    fun ^^u1 0
+    var P#arr
+      int
+    =
+    fun $owner 2
+    var $s
+      type-con $state 0
+    fun $ptr 2
+    fun $array 2
+    fun ^^u1 0
+    var P#len
+      int
+    fun $ref 1
+    fun $ptr 2
+    fun ^^u1 0
+    var P#arr
+      int
+    fun $me 0
+    fun $is 2
+    fun $ptr 2
+    fun $array 2
+    fun ^^u1 0
+    var P#len
+      int
+    fun $ref 1
+    fun $ptr 2
+    fun ^^u1 0
+    var P#arr
+      int
+    fun $array 2
+    fun ^^u1 0
+    var P#len
+      int
+    fun $typed 2
+    var $s
+      type-con $state 0
+    fun $ptr 2
+    fun $array 2
+    fun ^^u1 0
+    var P#len
+      int
+    fun $ref 1
+    fun $ptr 2
+    fun ^^u1 0
+    var P#arr
+      int
+    not
+    =
+    fun $kind_of 1
+    fun $array 2
+    fun ^^u1 0
+    var P#len
+      int
+    fun $kind_primitive 0
+    fun $is_non_primitive 1
+    fun $array 2
+    fun ^^u1 0
+    var P#len
+      int
+    implies
+    true
+    implies
+    fun $function_entry 1
+    var $s
+      type-con $state 0
+    implies
+    and 2
+    fun $good_state_ext 2
+    fun #tok$1^6.1 0
+    var $s
+      type-con $state 0
+    fun $full_stop 1
+    var $s
+      type-con $state 0
+    implies
+    forall 1 1 3
+      var f
+        type-con $pure_function 0
+      pat 1
+        fun $frame_level 1
+        var f
+          type-con $pure_function 0
+      attribute qid 1
+        string-attr VccPrelu.2703:13
+      attribute uniqueId 1
+        string-attr 389
+      attribute bvZ3Native 1
+        string-attr False
+    <
+    fun $frame_level 1
+    var f
+      type-con $pure_function 0
+    fun $current_frame_level 0
+    implies
+    and 2
+    fun $local_value_is 5
+    var $s
+      type-con $state 0
+    fun #tok$1^6.1 0
+    fun #loc.arr 0
+    fun $ptr_to_int 1
+    fun $ptr 2
+    fun ^^u1 0
+    var P#arr
+      int
+    fun $ptr_to 1
+    fun ^^u1 0
+    fun $local_value_is_ptr 5
+    var $s
+      type-con $state 0
+    fun #tok$1^6.1 0
+    fun #loc.arr 0
+    fun $ptr 2
+    fun ^^u1 0
+    var P#arr
+      int
+    fun $ptr_to 1
+    fun ^^u1 0
+    implies
+    fun $local_value_is 5
+    var $s
+      type-con $state 0
+    fun #tok$1^6.1 0
+    fun #loc.len 0
+    var P#len
+      int
+    fun ^^u4 0
+    implies
+    =
+    var #wrTime$1^6.1
+      int
+    fun $current_timestamp 1
+    var $s
+      type-con $state 0
+    implies
+    forall 1 1 3
+      var #p
+        type-con $ptr 0
+      pat 1
+        fun $in_writes_at 2
+        var #wrTime$1^6.1
+          int
+        var #p
+          type-con $ptr 0
+      attribute qid 1
+        string-attr nofile.0:0
+      attribute uniqueId 1
+        string-attr 2
+      attribute bvZ3Native 1
+        string-attr False
+    =
+    fun $in_writes_at 2
+    var #wrTime$1^6.1
+      int
+    var #p
+      type-con $ptr 0
+    false
+    implies
+    and 2
+    <=
+    int-num 0
+    var P#len
+      int
+    <=
+    var P#len
+      int
+    fun $max.u4 0
+    and 2
+    label neg 7 27
+    fun $in_domain_lab 4
+    var $s
+      type-con $state 0
+    fun $ptr 2
+    fun $array 2
+    fun ^^u1 0
+    var P#len
+      int
+    fun $ref 1
+    fun $ptr 2
+    fun ^^u1 0
+    var P#arr
+      int
+    fun $ptr 2
+    fun $array 2
+    fun ^^u1 0
+    var P#len
+      int
+    fun $ref 1
+    fun $ptr 2
+    fun ^^u1 0
+    var P#arr
+      int
+    fun l#public 0
+    implies
+    fun $in_domain_lab 4
+    var $s
+      type-con $state 0
+    fun $ptr 2
+    fun $array 2
+    fun ^^u1 0
+    var P#len
+      int
+    fun $ref 1
+    fun $ptr 2
+    fun ^^u1 0
+    var P#arr
+      int
+    fun $ptr 2
+    fun $array 2
+    fun ^^u1 0
+    var P#len
+      int
+    fun $ref 1
+    fun $ptr 2
+    fun ^^u1 0
+    var P#arr
+      int
+    fun l#public 0
+    and 2
+    label neg 12 14
+    and 2
+    fun $is 2
+    fun $idx 3
+    fun $ptr 2
+    fun ^^u1 0
+    var P#arr
+      int
+    int-num 0
+    fun ^^u1 0
+    fun ^^u1 0
+    fun $typed 2
+    var $s
+      type-con $state 0
+    fun $idx 3
+    fun $ptr 2
+    fun ^^u1 0
+    var P#arr
+      int
+    int-num 0
+    fun ^^u1 0
+    implies
+    and 2
+    fun $is 2
+    fun $idx 3
+    fun $ptr 2
+    fun ^^u1 0
+    var P#arr
+      int
+    int-num 0
+    fun ^^u1 0
+    fun ^^u1 0
+    fun $typed 2
+    var $s
+      type-con $state 0
+    fun $idx 3
+    fun $ptr 2
+    fun ^^u1 0
+    var P#arr
+      int
+    int-num 0
+    fun ^^u1 0
+    and 2
+    label neg 12 14
+    and 2
+    fun $is 2
+    fun $idx 3
+    fun $ptr 2
+    fun ^^u1 0
+    var P#arr
+      int
+    int-num 0
+    fun ^^u1 0
+    fun ^^u1 0
+    fun $thread_local 2
+    var $s
+      type-con $state 0
+    fun $idx 3
+    fun $ptr 2
+    fun ^^u1 0
+    var P#arr
+      int
+    int-num 0
+    fun ^^u1 0
+    implies
+    and 2
+    fun $is 2
+    fun $idx 3
+    fun $ptr 2
+    fun ^^u1 0
+    var P#arr
+      int
+    int-num 0
+    fun ^^u1 0
+    fun ^^u1 0
+    fun $thread_local 2
+    var $s
+      type-con $state 0
+    fun $idx 3
+    fun $ptr 2
+    fun ^^u1 0
+    var P#arr
+      int
+    int-num 0
+    fun ^^u1 0
+    implies
+    =
+    var L#max@0
+      int
+    fun $read_u1 2
+    var $s
+      type-con $state 0
+    fun $idx 3
+    fun $ptr 2
+    fun ^^u1 0
+    var P#arr
+      int
+    int-num 0
+    fun ^^u1 0
+    implies
+    fun $local_value_is 5
+    var $s
+      type-con $state 0
+    fun #tok$1^12.3 0
+    fun #loc.max 0
+    var L#max@0
+      int
+    fun ^^u1 0
+    implies
+    fun $local_value_is 5
+    var $s
+      type-con $state 0
+    fun #tok$1^14.3 0
+    fun #loc.witness 0
+    int-num 0
+    fun ^^u4 0
+    implies
+    fun $local_value_is 5
+    var $s
+      type-con $state 0
+    fun #tok$1^16.8 0
+    fun #loc.p 0
+    int-num 1
+    fun ^^u4 0
+    implies
+    and 4
+    <=
+    int-num 1
+    int-num 1
+    <=
+    int-num 1
+    int-num 1
+    <=
+    int-num 0
+    int-num 0
+    <=
+    int-num 0
+    int-num 0
+    and 2
+    label neg 17 17
+    <=
+    int-num 1
+    var P#len
+      int
+    implies
+    <=
+    int-num 1
+    var P#len
+      int
+    and 2
+    label neg 18 17
+    forall 1 0 3
+      var Q#i$1^18.17#tc1
+        int
+      attribute qid 1
+        string-attr maxc.18:17
+      attribute uniqueId 1
+        string-attr 3
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 2
+    <=
+    int-num 0
+    var Q#i$1^18.17#tc1
+      int
+    <=
+    var Q#i$1^18.17#tc1
+      int
+    fun $max.u4 0
+    implies
+    <
+    var Q#i$1^18.17#tc1
+      int
+    int-num 1
+    <=
+    fun $read_u1 2
+    var $s
+      type-con $state 0
+    fun $idx 3
+    fun $ptr 2
+    fun ^^u1 0
+    var P#arr
+      int
+    var Q#i$1^18.17#tc1
+      int
+    fun ^^u1 0
+    var L#max@0
+      int
+    implies
+    forall 1 0 3
+      var Q#i$1^18.17#tc1
+        int
+      attribute qid 1
+        string-attr maxc.18:17
+      attribute uniqueId 1
+        string-attr 3
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 2
+    <=
+    int-num 0
+    var Q#i$1^18.17#tc1
+      int
+    <=
+    var Q#i$1^18.17#tc1
+      int
+    fun $max.u4 0
+    implies
+    <
+    var Q#i$1^18.17#tc1
+      int
+    int-num 1
+    <=
+    fun $read_u1 2
+    var $s
+      type-con $state 0
+    fun $idx 3
+    fun $ptr 2
+    fun ^^u1 0
+    var P#arr
+      int
+    var Q#i$1^18.17#tc1
+      int
+    fun ^^u1 0
+    var L#max@0
+      int
+    and 2
+    label neg 19 17
+    and 2
+    <
+    int-num 0
+    var P#len
+      int
+    =
+    fun $read_u1 2
+    var $s
+      type-con $state 0
+    fun $idx 3
+    fun $ptr 2
+    fun ^^u1 0
+    var P#arr
+      int
+    int-num 0
+    fun ^^u1 0
+    var L#max@0
+      int
+    implies
+    and 2
+    <
+    int-num 0
+    var P#len
+      int
+    =
+    fun $read_u1 2
+    var $s
+      type-con $state 0
+    fun $idx 3
+    fun $ptr 2
+    fun ^^u1 0
+    var P#arr
+      int
+    int-num 0
+    fun ^^u1 0
+    var L#max@0
+      int
+    implies
+    label pos 0 0
+    true
+    implies
+    and 2
+    <=
+    int-num 0
+    var L#max@1
+      int
+    <=
+    var L#max@1
+      int
+    fun $max.u1 0
+    implies
+    and 2
+    <=
+    int-num 0
+    var SL#witness@0
+      int
+    <=
+    var SL#witness@0
+      int
+    fun $max.u4 0
+    implies
+    and 2
+    <=
+    int-num 0
+    var L#p@0
+      int
+    <=
+    var L#p@0
+      int
+    fun $max.u4 0
+    implies
+    and 2
+    <=
+    int-num 1
+    var L#p@0
+      int
+    <=
+    int-num 0
+    var SL#witness@0
+      int
+    implies
+    <=
+    var L#p@0
+      int
+    var P#len
+      int
+    implies
+    forall 1 0 3
+      var Q#i$1^18.17#tc1
+        int
+      attribute qid 1
+        string-attr maxc.18:17
+      attribute uniqueId 1
+        string-attr 3
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 2
+    <=
+    int-num 0
+    var Q#i$1^18.17#tc1
+      int
+    <=
+    var Q#i$1^18.17#tc1
+      int
+    fun $max.u4 0
+    implies
+    <
+    var Q#i$1^18.17#tc1
+      int
+    var L#p@0
+      int
+    <=
+    fun $read_u1 2
+    var $s
+      type-con $state 0
+    fun $idx 3
+    fun $ptr 2
+    fun ^^u1 0
+    var P#arr
+      int
+    var Q#i$1^18.17#tc1
+      int
+    fun ^^u1 0
+    var L#max@1
+      int
+    implies
+    and 2
+    <
+    var SL#witness@0
+      int
+    var P#len
+      int
+    =
+    fun $read_u1 2
+    var $s
+      type-con $state 0
+    fun $idx 3
+    fun $ptr 2
+    fun ^^u1 0
+    var P#arr
+      int
+    var SL#witness@0
+      int
+    fun ^^u1 0
+    var L#max@1
+      int
+    implies
+    and 2
+    <=
+    int-num 1
+    var L#p@0
+      int
+    <=
+    int-num 0
+    var SL#witness@0
+      int
+    and 2
+    implies
+    label pos 0 0
+    true
+    implies
+    and 2
+    <=
+    int-num 1
+    var L#p@0
+      int
+    <=
+    int-num 0
+    var SL#witness@0
+      int
+    implies
+    not
+    true
+    implies
+    and 2
+    <=
+    int-num 1
+    var L#p@0
+      int
+    <=
+    int-num 0
+    var SL#witness@0
+      int
+    implies
+    label pos 0 0
+    true
+    implies
+    and 2
+    <=
+    int-num 1
+    var L#p@0
+      int
+    <=
+    int-num 0
+    var SL#witness@0
+      int
+    implies
+    and 2
+    fun $good_state_ext 2
+    fun #tok$1^16.3 0
+    var $s
+      type-con $state 0
+    fun $full_stop 1
+    var $s
+      type-con $state 0
+    implies
+    and 2
+    <=
+    int-num 1
+    var L#p@0
+      int
+    <=
+    int-num 0
+    var SL#witness@0
+      int
+    implies
+    label pos 0 0
+    true
+    implies
+    and 2
+    <=
+    int-num 1
+    var L#p@0
+      int
+    <=
+    int-num 0
+    var SL#witness@0
+      int
+    and 2
+    label neg 27 3
+    fun $position_marker 0
+    implies
+    fun $position_marker 0
+    implies
+    and 2
+    <=
+    int-num 1
+    var L#p@0
+      int
+    <=
+    int-num 0
+    var SL#witness@0
+      int
+    implies
+    label pos 0 0
+    true
+    implies
+    and 2
+    <=
+    int-num 1
+    var L#p@0
+      int
+    <=
+    int-num 0
+    var SL#witness@0
+      int
+    implies
+    and 2
+    <=
+    int-num 1
+    var L#p@0
+      int
+    <=
+    int-num 0
+    var SL#witness@0
+      int
+    implies
+    label pos 0 0
+    true
+    implies
+    =
+    var L#max@4
+      int
+    var L#max@1
+      int
+    implies
+    =
+    var L#p@2
+      int
+    var L#p@0
+      int
+    implies
+    =
+    var SL#witness@2
+      int
+    var SL#witness@0
+      int
+    implies
+    =
+    var $result@0
+      int
+    var L#max@1
+      int
+    implies
+    label pos 0 0
+    true
+    and 2
+    label neg 9 14
+    forall 1 0 3
+      var Q#i$1^9.14#tc1
+        int
+      attribute qid 1
+        string-attr maxc.9:14
+      attribute uniqueId 1
+        string-attr 1
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 2
+    <=
+    int-num 0
+    var Q#i$1^9.14#tc1
+      int
+    <=
+    var Q#i$1^9.14#tc1
+      int
+    fun $max.u4 0
+    implies
+    <
+    var Q#i$1^9.14#tc1
+      int
+    var P#len
+      int
+    <=
+    fun $read_u1 2
+    var $s
+      type-con $state 0
+    fun $idx 3
+    fun $ptr 2
+    fun ^^u1 0
+    var P#arr
+      int
+    var Q#i$1^9.14#tc1
+      int
+    fun ^^u1 0
+    var $result@0
+      int
+    implies
+    forall 1 0 3
+      var Q#i$1^9.14#tc1
+        int
+      attribute qid 1
+        string-attr maxc.9:14
+      attribute uniqueId 1
+        string-attr 1
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 2
+    <=
+    int-num 0
+    var Q#i$1^9.14#tc1
+      int
+    <=
+    var Q#i$1^9.14#tc1
+      int
+    fun $max.u4 0
+    implies
+    <
+    var Q#i$1^9.14#tc1
+      int
+    var P#len
+      int
+    <=
+    fun $read_u1 2
+    var $s
+      type-con $state 0
+    fun $idx 3
+    fun $ptr 2
+    fun ^^u1 0
+    var P#arr
+      int
+    var Q#i$1^9.14#tc1
+      int
+    fun ^^u1 0
+    var $result@0
+      int
+    and 2
+    label neg 10 14
+    exists 1 0 3
+      var Q#i$1^10.14#tc1
+        int
+      attribute qid 1
+        string-attr maxc.10:14
+      attribute uniqueId 1
+        string-attr 0
+      attribute bvZ3Native 1
+        string-attr False
+    and 4
+    <=
+    int-num 0
+    var Q#i$1^10.14#tc1
+      int
+    <=
+    var Q#i$1^10.14#tc1
+      int
+    fun $max.u4 0
+    <
+    var Q#i$1^10.14#tc1
+      int
+    var P#len
+      int
+    =
+    fun $read_u1 2
+    var $s
+      type-con $state 0
+    fun $idx 3
+    fun $ptr 2
+    fun ^^u1 0
+    var P#arr
+      int
+    var Q#i$1^10.14#tc1
+      int
+    fun ^^u1 0
+    var $result@0
+      int
+    implies
+    exists 1 0 3
+      var Q#i$1^10.14#tc1
+        int
+      attribute qid 1
+        string-attr maxc.10:14
+      attribute uniqueId 1
+        string-attr 0
+      attribute bvZ3Native 1
+        string-attr False
+    and 4
+    <=
+    int-num 0
+    var Q#i$1^10.14#tc1
+      int
+    <=
+    var Q#i$1^10.14#tc1
+      int
+    fun $max.u4 0
+    <
+    var Q#i$1^10.14#tc1
+      int
+    var P#len
+      int
+    =
+    fun $read_u1 2
+    var $s
+      type-con $state 0
+    fun $idx 3
+    fun $ptr 2
+    fun ^^u1 0
+    var P#arr
+      int
+    var Q#i$1^10.14#tc1
+      int
+    fun ^^u1 0
+    var $result@0
+      int
+    true
+    implies
+    label pos 0 0
+    true
+    implies
+    and 2
+    <=
+    int-num 1
+    var L#p@0
+      int
+    <=
+    int-num 0
+    var SL#witness@0
+      int
+    implies
+    true
+    implies
+    and 2
+    <=
+    int-num 1
+    var L#p@0
+      int
+    <=
+    int-num 0
+    var SL#witness@0
+      int
+    implies
+    label pos 0 0
+    true
+    implies
+    and 2
+    <=
+    int-num 1
+    var L#p@0
+      int
+    <=
+    int-num 0
+    var SL#witness@0
+      int
+    implies
+    and 7
+    forall 1 1 3
+      var p
+        type-con $ptr 0
+      pat 1
+        fun $select.sm 2
+        fun $statusmap 1
+        var $s
+          type-con $state 0
+        var p
+          type-con $ptr 0
+      attribute qid 1
+        string-attr VccPrelu.1970:13
+      attribute uniqueId 1
+        string-attr 220
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    not
+    =
+    fun $kind_of 1
+    fun $typ 1
+    fun $owner 2
+    var $s
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    fun $kind_thread 0
+    not
+    =
+    fun $kind_of 1
+    fun $typ 1
+    fun $owner 2
+    var $s
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    fun $kind_thread 0
+    forall 1 1 3
+      var p
+        type-con $ptr 0
+      pat 1
+        fun $select.mem 2
+        fun $memory 1
+        var $s
+          type-con $state 0
+        var p
+          type-con $ptr 0
+      attribute qid 1
+        string-attr VccPrelu.1972:13
+      attribute uniqueId 1
+        string-attr 221
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    fun $thread_local 2
+    var $s
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    and 2
+    =
+    fun $select.mem 2
+    fun $memory 1
+    var $s
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    fun $select.mem 2
+    fun $memory 1
+    var $s
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    fun $thread_local 2
+    var $s
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    forall 1 1 3
+      var p
+        type-con $ptr 0
+      pat 1
+        fun $select.sm 2
+        fun $statusmap 1
+        var $s
+          type-con $state 0
+        var p
+          type-con $ptr 0
+      attribute qid 1
+        string-attr VccPrelu.1974:13
+      attribute uniqueId 1
+        string-attr 222
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    fun $thread_local 2
+    var $s
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    and 2
+    =
+    fun $select.sm 2
+    fun $statusmap 1
+    var $s
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    fun $select.sm 2
+    fun $statusmap 1
+    var $s
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    fun $thread_local 2
+    var $s
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    forall 1 1 3
+      var p
+        type-con $ptr 0
+      pat 1
+        fun $select.tm 2
+        fun $typemap 1
+        var $s
+          type-con $state 0
+        var p
+          type-con $ptr 0
+      attribute qid 1
+        string-attr VccPrelu.1976:13
+      attribute uniqueId 1
+        string-attr 223
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    fun $thread_local 2
+    var $s
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    and 2
+    =
+    fun $select.tm 2
+    fun $typemap 1
+    var $s
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    fun $select.tm 2
+    fun $typemap 1
+    var $s
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    fun $thread_local 2
+    var $s
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    <=
+    fun $current_timestamp 1
+    var $s
+      type-con $state 0
+    fun $current_timestamp 1
+    var $s
+      type-con $state 0
+    forall 1 1 4
+      var p
+        type-con $ptr 0
+      pat 1
+        fun $timestamp 2
+        var $s
+          type-con $state 0
+        var p
+          type-con $ptr 0
+      attribute qid 1
+        string-attr VccPrelu.1274:13
+      attribute uniqueId 1
+        string-attr 139
+      attribute bvZ3Native 1
+        string-attr False
+      attribute weight 1
+        expr-attr
+          int-num 0
+    <=
+    fun $timestamp 2
+    var $s
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    fun $timestamp 2
+    var $s
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    fun $call_transition 2
+    var $s
+      type-con $state 0
+    var $s
+      type-con $state 0
+    implies
+    and 3
+    <=
+    fun $current_timestamp 1
+    var $s
+      type-con $state 0
+    fun $current_timestamp 1
+    var $s
+      type-con $state 0
+    forall 1 1 4
+      var p
+        type-con $ptr 0
+      pat 1
+        fun $timestamp 2
+        var $s
+          type-con $state 0
+        var p
+          type-con $ptr 0
+      attribute qid 1
+        string-attr VccPrelu.1274:13
+      attribute uniqueId 1
+        string-attr 139
+      attribute bvZ3Native 1
+        string-attr False
+      attribute weight 1
+        expr-attr
+          int-num 0
+    <=
+    fun $timestamp 2
+    var $s
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    fun $timestamp 2
+    var $s
+      type-con $state 0
+    var p
+      type-con $ptr 0
+    fun $call_transition 2
+    var $s
+      type-con $state 0
+    var $s
+      type-con $state 0
+    implies
+    and 2
+    fun $good_state_ext 2
+    fun #tok$1^16.3 0
+    var $s
+      type-con $state 0
+    fun $full_stop 1
+    var $s
+      type-con $state 0
+    implies
+    fun $local_value_is 5
+    var $s
+      type-con $state 0
+    fun #tok$1^16.3 0
+    fun #loc.p 0
+    var L#p@0
+      int
+    fun ^^u4 0
+    implies
+    fun $local_value_is 5
+    var $s
+      type-con $state 0
+    fun #tok$1^16.3 0
+    fun #loc.witness 0
+    var SL#witness@0
+      int
+    fun ^^u4 0
+    implies
+    fun $local_value_is 5
+    var $s
+      type-con $state 0
+    fun #tok$1^16.3 0
+    fun #loc.max 0
+    var L#max@1
+      int
+    fun ^^u1 0
+    implies
+    fun $local_value_is 5
+    var $s
+      type-con $state 0
+    fun #tok$1^16.3 0
+    fun #loc.len 0
+    var P#len
+      int
+    fun ^^u4 0
+    implies
+    and 2
+    fun $local_value_is 5
+    var $s
+      type-con $state 0
+    fun #tok$1^16.3 0
+    fun #loc.arr 0
+    fun $ptr_to_int 1
+    fun $ptr 2
+    fun ^^u1 0
+    var P#arr
+      int
+    fun $ptr_to 1
+    fun ^^u1 0
+    fun $local_value_is_ptr 5
+    var $s
+      type-con $state 0
+    fun #tok$1^16.3 0
+    fun #loc.arr 0
+    fun $ptr 2
+    fun ^^u1 0
+    var P#arr
+      int
+    fun $ptr_to 1
+    fun ^^u1 0
+    implies
+    and 2
+    =
+    fun $typemap 1
+    var $s
+      type-con $state 0
+    fun $typemap 1
+    var $s
+      type-con $state 0
+    =
+    fun $statusmap 1
+    var $s
+      type-con $state 0
+    fun $statusmap 1
+    var $s
+      type-con $state 0
+    implies
+    and 2
+    <=
+    int-num 1
+    var L#p@0
+      int
+    <=
+    int-num 0
+    var SL#witness@0
+      int
+    and 2
+    implies
+    label pos 0 0
+    true
+    implies
+    and 2
+    <=
+    int-num 1
+    var L#p@0
+      int
+    <=
+    int-num 0
+    var SL#witness@0
+      int
+    implies
+    <
+    var L#p@0
+      int
+    var P#len
+      int
+    implies
+    and 2
+    <=
+    int-num 1
+    var L#p@0
+      int
+    <=
+    int-num 0
+    var SL#witness@0
+      int
+    implies
+    label pos 0 0
+    true
+    implies
+    and 2
+    <=
+    int-num 1
+    var L#p@0
+      int
+    <=
+    int-num 0
+    var SL#witness@0
+      int
+    and 2
+    label neg 21 9
+    and 2
+    fun $is 2
+    fun $idx 3
+    fun $ptr 2
+    fun ^^u1 0
+    var P#arr
+      int
+    var L#p@0
+      int
+    fun ^^u1 0
+    fun ^^u1 0
+    fun $typed 2
+    var $s
+      type-con $state 0
+    fun $idx 3
+    fun $ptr 2
+    fun ^^u1 0
+    var P#arr
+      int
+    var L#p@0
+      int
+    fun ^^u1 0
+    implies
+    and 2
+    fun $is 2
+    fun $idx 3
+    fun $ptr 2
+    fun ^^u1 0
+    var P#arr
+      int
+    var L#p@0
+      int
+    fun ^^u1 0
+    fun ^^u1 0
+    fun $typed 2
+    var $s
+      type-con $state 0
+    fun $idx 3
+    fun $ptr 2
+    fun ^^u1 0
+    var P#arr
+      int
+    var L#p@0
+      int
+    fun ^^u1 0
+    and 2
+    label neg 21 9
+    and 2
+    fun $is 2
+    fun $idx 3
+    fun $ptr 2
+    fun ^^u1 0
+    var P#arr
+      int
+    var L#p@0
+      int
+    fun ^^u1 0
+    fun ^^u1 0
+    fun $thread_local 2
+    var $s
+      type-con $state 0
+    fun $idx 3
+    fun $ptr 2
+    fun ^^u1 0
+    var P#arr
+      int
+    var L#p@0
+      int
+    fun ^^u1 0
+    implies
+    and 2
+    fun $is 2
+    fun $idx 3
+    fun $ptr 2
+    fun ^^u1 0
+    var P#arr
+      int
+    var L#p@0
+      int
+    fun ^^u1 0
+    fun ^^u1 0
+    fun $thread_local 2
+    var $s
+      type-con $state 0
+    fun $idx 3
+    fun $ptr 2
+    fun ^^u1 0
+    var P#arr
+      int
+    var L#p@0
+      int
+    fun ^^u1 0
+    implies
+    and 2
+    <=
+    int-num 1
+    var L#p@0
+      int
+    <=
+    int-num 0
+    var SL#witness@0
+      int
+    and 2
+    implies
+    label pos 0 0
+    true
+    implies
+    and 2
+    <=
+    int-num 1
+    var L#p@0
+      int
+    <=
+    int-num 0
+    var SL#witness@0
+      int
+    implies
+    >
+    fun $read_u1 2
+    var $s
+      type-con $state 0
+    fun $idx 3
+    fun $ptr 2
+    fun ^^u1 0
+    var P#arr
+      int
+    var L#p@0
+      int
+    fun ^^u1 0
+    var L#max@1
+      int
+    implies
+    and 2
+    <=
+    int-num 1
+    var L#p@0
+      int
+    <=
+    int-num 0
+    var SL#witness@0
+      int
+    implies
+    label pos 0 0
+    true
+    implies
+    and 2
+    <=
+    int-num 1
+    var L#p@0
+      int
+    <=
+    int-num 0
+    var SL#witness@0
+      int
+    and 2
+    label neg 23 13
+    and 2
+    fun $is 2
+    fun $idx 3
+    fun $ptr 2
+    fun ^^u1 0
+    var P#arr
+      int
+    var L#p@0
+      int
+    fun ^^u1 0
+    fun ^^u1 0
+    fun $typed 2
+    var $s
+      type-con $state 0
+    fun $idx 3
+    fun $ptr 2
+    fun ^^u1 0
+    var P#arr
+      int
+    var L#p@0
+      int
+    fun ^^u1 0
+    implies
+    and 2
+    fun $is 2
+    fun $idx 3
+    fun $ptr 2
+    fun ^^u1 0
+    var P#arr
+      int
+    var L#p@0
+      int
+    fun ^^u1 0
+    fun ^^u1 0
+    fun $typed 2
+    var $s
+      type-con $state 0
+    fun $idx 3
+    fun $ptr 2
+    fun ^^u1 0
+    var P#arr
+      int
+    var L#p@0
+      int
+    fun ^^u1 0
+    and 2
+    label neg 23 13
+    and 2
+    fun $is 2
+    fun $idx 3
+    fun $ptr 2
+    fun ^^u1 0
+    var P#arr
+      int
+    var L#p@0
+      int
+    fun ^^u1 0
+    fun ^^u1 0
+    fun $thread_local 2
+    var $s
+      type-con $state 0
+    fun $idx 3
+    fun $ptr 2
+    fun ^^u1 0
+    var P#arr
+      int
+    var L#p@0
+      int
+    fun ^^u1 0
+    implies
+    and 2
+    fun $is 2
+    fun $idx 3
+    fun $ptr 2
+    fun ^^u1 0
+    var P#arr
+      int
+    var L#p@0
+      int
+    fun ^^u1 0
+    fun ^^u1 0
+    fun $thread_local 2
+    var $s
+      type-con $state 0
+    fun $idx 3
+    fun $ptr 2
+    fun ^^u1 0
+    var P#arr
+      int
+    var L#p@0
+      int
+    fun ^^u1 0
+    implies
+    =
+    var L#max@2
+      int
+    fun $read_u1 2
+    var $s
+      type-con $state 0
+    fun $idx 3
+    fun $ptr 2
+    fun ^^u1 0
+    var P#arr
+      int
+    var L#p@0
+      int
+    fun ^^u1 0
+    implies
+    fun $local_value_is 5
+    var $s
+      type-con $state 0
+    fun #tok$1^23.7 0
+    fun #loc.max 0
+    var L#max@2
+      int
+    fun ^^u1 0
+    implies
+    fun $local_value_is 5
+    var $s
+      type-con $state 0
+    fun #tok$1^24.47 0
+    fun #loc.witness 0
+    var L#p@0
+      int
+    fun ^^u4 0
+    implies
+    and 2
+    <=
+    int-num 1
+    var L#p@0
+      int
+    <=
+    int-num 1
+    var L#p@0
+      int
+    implies
+    label pos 0 0
+    true
+    implies
+    =
+    var L#max@3
+      int
+    var L#max@2
+      int
+    implies
+    =
+    var SL#witness@1
+      int
+    var L#p@0
+      int
+    implies
+    label pos 0 0
+    true
+    implies
+    and 2
+    <=
+    int-num 1
+    var L#p@0
+      int
+    <=
+    int-num 0
+    var SL#witness@1
+      int
+    and 2
+    label neg 16 24
+    and 2
+    <=
+    int-num 0
+    +
+    var L#p@0
+      int
+    int-num 1
+    <=
+    +
+    var L#p@0
+      int
+    int-num 1
+    fun $max.u4 0
+    implies
+    and 2
+    <=
+    int-num 0
+    +
+    var L#p@0
+      int
+    int-num 1
+    <=
+    +
+    var L#p@0
+      int
+    int-num 1
+    fun $max.u4 0
+    implies
+    =
+    var L#p@1
+      int
+    +
+    var L#p@0
+      int
+    int-num 1
+    implies
+    fun $local_value_is 5
+    var $s
+      type-con $state 0
+    fun #tok$1^16.24 0
+    fun #loc.p 0
+    var L#p@1
+      int
+    fun ^^u4 0
+    implies
+    and 2
+    <=
+    int-num 2
+    var L#p@1
+      int
+    <=
+    int-num 0
+    var SL#witness@1
+      int
+    implies
+    label pos 0 0
+    true
+    and 2
+    label neg 17 17
+    <=
+    var L#p@1
+      int
+    var P#len
+      int
+    implies
+    <=
+    var L#p@1
+      int
+    var P#len
+      int
+    and 2
+    label neg 18 17
+    forall 1 0 3
+      var Q#i$1^18.17#tc1
+        int
+      attribute qid 1
+        string-attr maxc.18:17
+      attribute uniqueId 1
+        string-attr 3
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 2
+    <=
+    int-num 0
+    var Q#i$1^18.17#tc1
+      int
+    <=
+    var Q#i$1^18.17#tc1
+      int
+    fun $max.u4 0
+    implies
+    <
+    var Q#i$1^18.17#tc1
+      int
+    var L#p@1
+      int
+    <=
+    fun $read_u1 2
+    var $s
+      type-con $state 0
+    fun $idx 3
+    fun $ptr 2
+    fun ^^u1 0
+    var P#arr
+      int
+    var Q#i$1^18.17#tc1
+      int
+    fun ^^u1 0
+    var L#max@3
+      int
+    implies
+    forall 1 0 3
+      var Q#i$1^18.17#tc1
+        int
+      attribute qid 1
+        string-attr maxc.18:17
+      attribute uniqueId 1
+        string-attr 3
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 2
+    <=
+    int-num 0
+    var Q#i$1^18.17#tc1
+      int
+    <=
+    var Q#i$1^18.17#tc1
+      int
+    fun $max.u4 0
+    implies
+    <
+    var Q#i$1^18.17#tc1
+      int
+    var L#p@1
+      int
+    <=
+    fun $read_u1 2
+    var $s
+      type-con $state 0
+    fun $idx 3
+    fun $ptr 2
+    fun ^^u1 0
+    var P#arr
+      int
+    var Q#i$1^18.17#tc1
+      int
+    fun ^^u1 0
+    var L#max@3
+      int
+    and 2
+    label neg 19 17
+    and 2
+    <
+    var SL#witness@1
+      int
+    var P#len
+      int
+    =
+    fun $read_u1 2
+    var $s
+      type-con $state 0
+    fun $idx 3
+    fun $ptr 2
+    fun ^^u1 0
+    var P#arr
+      int
+    var SL#witness@1
+      int
+    fun ^^u1 0
+    var L#max@3
+      int
+    implies
+    and 2
+    <
+    var SL#witness@1
+      int
+    var P#len
+      int
+    =
+    fun $read_u1 2
+    var $s
+      type-con $state 0
+    fun $idx 3
+    fun $ptr 2
+    fun ^^u1 0
+    var P#arr
+      int
+    var SL#witness@1
+      int
+    fun ^^u1 0
+    var L#max@3
+      int
+    implies
+    false
+    true
+    implies
+    label pos 0 0
+    true
+    implies
+    and 2
+    <=
+    int-num 1
+    var L#p@0
+      int
+    <=
+    int-num 0
+    var SL#witness@0
+      int
+    implies
+    <=
+    fun $read_u1 2
+    var $s
+      type-con $state 0
+    fun $idx 3
+    fun $ptr 2
+    fun ^^u1 0
+    var P#arr
+      int
+    var L#p@0
+      int
+    fun ^^u1 0
+    var L#max@1
+      int
+    implies
+    and 2
+    <=
+    int-num 1
+    var L#p@0
+      int
+    <=
+    int-num 0
+    var SL#witness@0
+      int
+    implies
+    label pos 0 0
+    true
+    implies
+    and 2
+    <=
+    int-num 1
+    var L#p@0
+      int
+    <=
+    int-num 0
+    var SL#witness@0
+      int
+    implies
+    and 2
+    <=
+    int-num 1
+    var L#p@0
+      int
+    <=
+    int-num 0
+    var SL#witness@0
+      int
+    implies
+    label pos 0 0
+    true
+    implies
+    =
+    var L#max@3
+      int
+    var L#max@1
+      int
+    implies
+    =
+    var SL#witness@1
+      int
+    var SL#witness@0
+      int
+    implies
+    label pos 0 0
+    true
+    implies
+    and 2
+    <=
+    int-num 1
+    var L#p@0
+      int
+    <=
+    int-num 0
+    var SL#witness@1
+      int
+    and 2
+    label neg 16 24
+    and 2
+    <=
+    int-num 0
+    +
+    var L#p@0
+      int
+    int-num 1
+    <=
+    +
+    var L#p@0
+      int
+    int-num 1
+    fun $max.u4 0
+    implies
+    and 2
+    <=
+    int-num 0
+    +
+    var L#p@0
+      int
+    int-num 1
+    <=
+    +
+    var L#p@0
+      int
+    int-num 1
+    fun $max.u4 0
+    implies
+    =
+    var L#p@1
+      int
+    +
+    var L#p@0
+      int
+    int-num 1
+    implies
+    fun $local_value_is 5
+    var $s
+      type-con $state 0
+    fun #tok$1^16.24 0
+    fun #loc.p 0
+    var L#p@1
+      int
+    fun ^^u4 0
+    implies
+    and 2
+    <=
+    int-num 2
+    var L#p@1
+      int
+    <=
+    int-num 0
+    var SL#witness@1
+      int
+    implies
+    label pos 0 0
+    true
+    and 2
+    label neg 17 17
+    <=
+    var L#p@1
+      int
+    var P#len
+      int
+    implies
+    <=
+    var L#p@1
+      int
+    var P#len
+      int
+    and 2
+    label neg 18 17
+    forall 1 0 3
+      var Q#i$1^18.17#tc1
+        int
+      attribute qid 1
+        string-attr maxc.18:17
+      attribute uniqueId 1
+        string-attr 3
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 2
+    <=
+    int-num 0
+    var Q#i$1^18.17#tc1
+      int
+    <=
+    var Q#i$1^18.17#tc1
+      int
+    fun $max.u4 0
+    implies
+    <
+    var Q#i$1^18.17#tc1
+      int
+    var L#p@1
+      int
+    <=
+    fun $read_u1 2
+    var $s
+      type-con $state 0
+    fun $idx 3
+    fun $ptr 2
+    fun ^^u1 0
+    var P#arr
+      int
+    var Q#i$1^18.17#tc1
+      int
+    fun ^^u1 0
+    var L#max@3
+      int
+    implies
+    forall 1 0 3
+      var Q#i$1^18.17#tc1
+        int
+      attribute qid 1
+        string-attr maxc.18:17
+      attribute uniqueId 1
+        string-attr 3
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 2
+    <=
+    int-num 0
+    var Q#i$1^18.17#tc1
+      int
+    <=
+    var Q#i$1^18.17#tc1
+      int
+    fun $max.u4 0
+    implies
+    <
+    var Q#i$1^18.17#tc1
+      int
+    var L#p@1
+      int
+    <=
+    fun $read_u1 2
+    var $s
+      type-con $state 0
+    fun $idx 3
+    fun $ptr 2
+    fun ^^u1 0
+    var P#arr
+      int
+    var Q#i$1^18.17#tc1
+      int
+    fun ^^u1 0
+    var L#max@3
+      int
+    and 2
+    label neg 19 17
+    and 2
+    <
+    var SL#witness@1
+      int
+    var P#len
+      int
+    =
+    fun $read_u1 2
+    var $s
+      type-con $state 0
+    fun $idx 3
+    fun $ptr 2
+    fun ^^u1 0
+    var P#arr
+      int
+    var SL#witness@1
+      int
+    fun ^^u1 0
+    var L#max@3
+      int
+    implies
+    and 2
+    <
+    var SL#witness@1
+      int
+    var P#len
+      int
+    =
+    fun $read_u1 2
+    var $s
+      type-con $state 0
+    fun $idx 3
+    fun $ptr 2
+    fun ^^u1 0
+    var P#arr
+      int
+    var SL#witness@1
+      int
+    fun ^^u1 0
+    var L#max@3
+      int
+    implies
+    false
+    true
+    implies
+    label pos 0 0
+    true
+    implies
+    and 2
+    <=
+    int-num 1
+    var L#p@0
+      int
+    <=
+    int-num 0
+    var SL#witness@0
+      int
+    implies
+    >=
+    var L#p@0
+      int
+    var P#len
+      int
+    implies
+    and 2
+    <=
+    int-num 1
+    var L#p@0
+      int
+    <=
+    int-num 0
+    var SL#witness@0
+      int
+    implies
+    label pos 0 0
+    true
+    implies
+    and 2
+    <=
+    int-num 1
+    var L#p@0
+      int
+    <=
+    int-num 0
+    var SL#witness@0
+      int
+    implies
+    and 2
+    <=
+    int-num 1
+    var L#p@0
+      int
+    <=
+    int-num 0
+    var SL#witness@0
+      int
+    implies
+    label pos 0 0
+    true
+    implies
+    and 2
+    <=
+    int-num 1
+    var L#p@0
+      int
+    <=
+    int-num 0
+    var SL#witness@0
+      int
+    and 2
+    label neg 27 3
+    fun $position_marker 0
+    implies
+    fun $position_marker 0
+    implies
+    and 2
+    <=
+    int-num 1
+    var L#p@0
+      int
+    <=
+    int-num 0
+    var SL#witness@0
+      int
+    implies
+    label pos 0 0
+    true
+    implies
+    and 2
+    <=
+    int-num 1
+    var L#p@0
+      int
+    <=
+    int-num 0
+    var SL#witness@0
+      int
+    implies
+    and 2
+    <=
+    int-num 1
+    var L#p@0
+      int
+    <=
+    int-num 0
+    var SL#witness@0
+      int
+    implies
+    label pos 0 0
+    true
+    implies
+    =
+    var L#max@4
+      int
+    var L#max@1
+      int
+    implies
+    =
+    var L#p@2
+      int
+    var L#p@0
+      int
+    implies
+    =
+    var SL#witness@2
+      int
+    var SL#witness@0
+      int
+    implies
+    =
+    var $result@0
+      int
+    var L#max@1
+      int
+    implies
+    label pos 0 0
+    true
+    and 2
+    label neg 9 14
+    forall 1 0 3
+      var Q#i$1^9.14#tc1
+        int
+      attribute qid 1
+        string-attr maxc.9:14
+      attribute uniqueId 1
+        string-attr 1
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 2
+    <=
+    int-num 0
+    var Q#i$1^9.14#tc1
+      int
+    <=
+    var Q#i$1^9.14#tc1
+      int
+    fun $max.u4 0
+    implies
+    <
+    var Q#i$1^9.14#tc1
+      int
+    var P#len
+      int
+    <=
+    fun $read_u1 2
+    var $s
+      type-con $state 0
+    fun $idx 3
+    fun $ptr 2
+    fun ^^u1 0
+    var P#arr
+      int
+    var Q#i$1^9.14#tc1
+      int
+    fun ^^u1 0
+    var $result@0
+      int
+    implies
+    forall 1 0 3
+      var Q#i$1^9.14#tc1
+        int
+      attribute qid 1
+        string-attr maxc.9:14
+      attribute uniqueId 1
+        string-attr 1
+      attribute bvZ3Native 1
+        string-attr False
+    implies
+    and 2
+    <=
+    int-num 0
+    var Q#i$1^9.14#tc1
+      int
+    <=
+    var Q#i$1^9.14#tc1
+      int
+    fun $max.u4 0
+    implies
+    <
+    var Q#i$1^9.14#tc1
+      int
+    var P#len
+      int
+    <=
+    fun $read_u1 2
+    var $s
+      type-con $state 0
+    fun $idx 3
+    fun $ptr 2
+    fun ^^u1 0
+    var P#arr
+      int
+    var Q#i$1^9.14#tc1
+      int
+    fun ^^u1 0
+    var $result@0
+      int
+    and 2
+    label neg 10 14
+    exists 1 0 3
+      var Q#i$1^10.14#tc1
+        int
+      attribute qid 1
+        string-attr maxc.10:14
+      attribute uniqueId 1
+        string-attr 0
+      attribute bvZ3Native 1
+        string-attr False
+    and 4
+    <=
+    int-num 0
+    var Q#i$1^10.14#tc1
+      int
+    <=
+    var Q#i$1^10.14#tc1
+      int
+    fun $max.u4 0
+    <
+    var Q#i$1^10.14#tc1
+      int
+    var P#len
+      int
+    =
+    fun $read_u1 2
+    var $s
+      type-con $state 0
+    fun $idx 3
+    fun $ptr 2
+    fun ^^u1 0
+    var P#arr
+      int
+    var Q#i$1^10.14#tc1
+      int
+    fun ^^u1 0
+    var $result@0
+      int
+    implies
+    exists 1 0 3
+      var Q#i$1^10.14#tc1
+        int
+      attribute qid 1
+        string-attr maxc.10:14
+      attribute uniqueId 1
+        string-attr 0
+      attribute bvZ3Native 1
+        string-attr False
+    and 4
+    <=
+    int-num 0
+    var Q#i$1^10.14#tc1
+      int
+    <=
+    var Q#i$1^10.14#tc1
+      int
+    fun $max.u4 0
+    <
+    var Q#i$1^10.14#tc1
+      int
+    var P#len
+      int
+    =
+    fun $read_u1 2
+    var $s
+      type-con $state 0
+    fun $idx 3
+    fun $ptr 2
+    fun ^^u1 0
+    var P#arr
+      int
+    var Q#i$1^10.14#tc1
+      int
+    fun ^^u1 0
+    var $result@0
+      int
+    true
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/SMT_Examples/VCC_Max.certs	Tue Jul 23 18:36:23 2013 +0200
@@ -0,0 +1,7207 @@
+2122eda8d3638c072eaaa16a2c285fe3e5c83f7e 7206 0
+WARNING: For problems containing quantifiers, the model finding capabilities of Z3 work better when the formula does not contain nested quantifiers. You can use PULL_NESTED_QUANTIFIERS=true to eliminate nested quantifiers.
+#2 := false
+decl f470 :: Int
+#4962 := f470
+decl f55 :: (-> S17 S11 Int)
+decl f139 :: (-> S60 S4 S11)
+decl f35 :: S4
+#356 := f35
+decl f140 :: (-> S61 Int S60)
+decl f471 :: Int
+#4964 := f471
+decl f153 :: (-> S68 S11 S61)
+decl f87 :: (-> S34 Int S11)
+decl f445 :: Int
+#4655 := f445
+decl f113 :: (-> S49 S4 S34)
+decl f114 :: S49
+#1136 := f114
+#4654 := (f113 f114 f35)
+#4656 := (f87 #4654 f445)
+decl f154 :: S68
+#1350 := f154
+#4734 := (f153 f154 #4656)
+#4988 := (f140 #4734 f471)
+#4989 := (f139 #4988 f35)
+decl f103 :: (-> S42 S10 S17)
+decl f444 :: S10
+#4649 := f444
+decl f199 :: S42
+#2434 := f199
+#4748 := (f103 f199 f444)
+#4990 := (f55 #4748 #4989)
+#4991 := (= #4990 f470)
+#20985 := (not #4991)
+#1138 := 0::Int
+#6707 := -1::Int
+#12718 := (* -1::Int f471)
+decl f443 :: Int
+#4646 := f443
+#12719 := (+ f443 #12718)
+#12720 := (<= #12719 0::Int)
+#20986 := (or #12720 #20985)
+#20987 := (not #20986)
+#1197 := (:var 0 Int)
+#4773 := (f140 #4734 #1197)
+#21878 := (pattern #4773)
+#12696 := (* -1::Int f470)
+#4774 := (f139 #4773 f35)
+#4775 := (f55 #4748 #4774)
+#12697 := (+ #4775 #12696)
+#12698 := (<= #12697 0::Int)
+decl f472 :: Int
+#4972 := f472
+#12677 := (* -1::Int f472)
+#12685 := (+ #1197 #12677)
+#12684 := (>= #12685 0::Int)
+#9359 := 4294967295::Int
+#14917 := (<= #1197 4294967295::Int)
+#18181 := (not #14917)
+#6706 := (>= #1197 0::Int)
+#7428 := (not #6706)
+#20977 := (or #7428 #18181 #12684 #12698)
+#21895 := (forall (vars (?v0 Int)) (:pat #21878) #20977)
+#21900 := (not #21895)
+#21903 := (or #21900 #20987)
+#21906 := (not #21903)
+decl ?v0!14 :: Int
+#17232 := ?v0!14
+#17239 := (f140 #4734 ?v0!14)
+#17240 := (f139 #17239 f35)
+#17241 := (f55 #4748 #17240)
+#17543 := (* -1::Int #17241)
+#17544 := (+ f470 #17543)
+#17545 := (>= #17544 0::Int)
+#17530 := (* -1::Int ?v0!14)
+#17531 := (+ f472 #17530)
+#17532 := (<= #17531 0::Int)
+#17234 := (<= ?v0!14 4294967295::Int)
+#20951 := (not #17234)
+#17233 := (>= ?v0!14 0::Int)
+#20950 := (not #17233)
+#20966 := (or #20950 #20951 #17532 #17545)
+#20971 := (not #20966)
+#21909 := (or #20971 #21906)
+#21912 := (not #21909)
+#12678 := (+ f443 #12677)
+#12676 := (>= #12678 0::Int)
+#12681 := (not #12676)
+#21915 := (or #12681 #21912)
+#21918 := (not #21915)
+#21921 := (or #12681 #21918)
+#21924 := (not #21921)
+#12651 := (>= f471 0::Int)
+#21027 := (not #12651)
+#2098 := 2::Int
+#12668 := (>= f472 2::Int)
+#21026 := (not #12668)
+decl f1 :: S1
+#3 := f1
+decl f45 :: (-> S7 S4 S1)
+decl f33 :: S4
+#350 := f33
+decl f46 :: (-> S8 Int S7)
+decl f449 :: (-> S178 S3 S8)
+decl f20 :: S3
+#66 := f20
+decl f450 :: (-> S179 S3 S178)
+decl f9 :: S3
+#33 := f9
+decl f451 :: (-> S180 S10 S179)
+decl f452 :: S180
+#4695 := f452
+#4696 := (f451 f452 f444)
+#4974 := (f450 #4696 f9)
+#4975 := (f449 #4974 f20)
+#4976 := (f46 #4975 f472)
+#4977 := (f45 #4976 f33)
+#4978 := (= #4977 f1)
+#11705 := (not #4978)
+decl f464 :: Int
+#4790 := f464
+#12740 := (+ f464 #12677)
+#12739 := (= #12740 -1::Int)
+#12743 := (not #12739)
+#13726 := 4294967294::Int
+#13727 := (<= f464 4294967294::Int)
+#17212 := (not #13727)
+#12660 := (>= f464 -1::Int)
+#17209 := (not #12660)
+#21927 := (or #17209 #17212 #12743 #11705 #21026 #21027 #21924)
+#21930 := (not #21927)
+#21933 := (or #17209 #17212 #21930)
+#21936 := (not #21933)
+#968 := 1::Int
+#12639 := (>= f464 1::Int)
+#12777 := (not #12639)
+#4965 := (= f471 f464)
+#11751 := (not #4965)
+decl f469 :: Int
+#4949 := f469
+#4963 := (= f470 f469)
+#11760 := (not #4963)
+decl f10 :: S3
+#36 := f10
+decl f8 :: S3
+#30 := f8
+#4956 := (f450 #4696 f8)
+#4957 := (f449 #4956 f10)
+#4958 := (f46 #4957 f464)
+#4959 := (f45 #4958 f33)
+#4960 := (= #4959 f1)
+#11785 := (not #4960)
+decl f17 :: S3
+#57 := f17
+decl f11 :: S3
+#39 := f11
+#4951 := (f450 #4696 f11)
+#4952 := (f449 #4951 f17)
+#4953 := (f46 #4952 f469)
+#4954 := (f45 #4953 f35)
+#4955 := (= #4954 f1)
+#11794 := (not #4955)
+#4936 := (f140 #4734 f464)
+#4937 := (f139 #4936 f35)
+#4947 := (f55 #4748 #4937)
+#4950 := (= f469 #4947)
+#11803 := (not #4950)
+decl f71 :: (-> S27 S11 S1)
+decl f80 :: (-> S31 S10 S27)
+decl f157 :: S31
+#1372 := f157
+#4743 := (f80 f157 f444)
+#4944 := (f71 #4743 #4937)
+#4945 := (= #4944 f1)
+#17180 := (not #4945)
+decl f118 :: (-> S51 S11 S7)
+decl f123 :: S51
+#1172 := f123
+#4938 := (f118 f123 #4937)
+#4939 := (f45 #4938 f35)
+#4940 := (= #4939 f1)
+#17171 := (not #4940)
+#21939 := (or #17171 #17180 #11803 #11794 #11785 #11760 #11751 #12777 #21027 #21936)
+#21942 := (not #21939)
+#21945 := (or #17171 #17180 #21942)
+#21948 := (not #21945)
+decl f244 :: (-> S93 S4 Int)
+decl f245 :: S93
+#2903 := f245
+#4624 := (f244 f245 f35)
+#25629 := (* #4624 f464)
+#4735 := (f140 #4734 0::Int)
+#4736 := (f139 #4735 f35)
+decl f206 :: S17
+#2483 := f206
+#23971 := (f55 f206 #4736)
+#23991 := (f87 #4654 #23971)
+#24379 := (f55 f206 #23991)
+#25632 := (+ #24379 #25629)
+#25639 := (f87 #4654 #25632)
+decl f358 :: S31
+#3975 := f358
+#24190 := (f80 f358 f444)
+#25878 := (f71 #24190 #25639)
+#25879 := (= #25878 f1)
+decl f85 :: S11
+#1075 := f85
+decl f82 :: (-> S32 S11 S11)
+decl f83 :: (-> S33 S10 S32)
+decl f84 :: S33
+#1072 := f84
+#4661 := (f83 f84 f444)
+#25876 := (f82 #4661 #25639)
+#25877 := (= #25876 f85)
+#25880 := (or #25877 #25879)
+#25881 := (not #25880)
+decl f3 :: S2
+#7 := f3
+decl f61 :: (-> S4 S2)
+decl f62 :: (-> S22 S11 S4)
+decl f63 :: S22
+#999 := f63
+#25851 := (f62 f63 #25639)
+#25852 := (f61 #25851)
+#25853 := (= #25852 f3)
+#25882 := (or #25853 #25881)
+#25883 := (not #25882)
+decl f235 :: (-> S88 S56 S11)
+decl f134 :: (-> S55 S11 S56)
+decl f135 :: (-> S57 S58 S55)
+decl f137 :: (-> S59 S10 S58)
+decl f138 :: S59
+#1302 := f138
+#4882 := (f137 f138 f444)
+decl f136 :: S57
+#1301 := f136
+#4883 := (f135 f136 #4882)
+#25855 := (f134 #4883 #25639)
+decl f236 :: S88
+#2770 := f236
+#25859 := (f235 f236 #25855)
+#25870 := (f71 #24190 #25859)
+#25871 := (= #25870 f1)
+#25868 := (f82 #4661 #25859)
+#25869 := (= #25868 f85)
+#25872 := (or #25869 #25871)
+#25873 := (not #25872)
+#25865 := (f62 f63 #25859)
+#25866 := (f61 #25865)
+#25867 := (= #25866 f3)
+decl f86 :: S31
+#1078 := f86
+#4650 := (f80 f86 f444)
+#25860 := (f71 #4650 #25859)
+#25861 := (= #25860 f1)
+#25862 := (not #25861)
+decl f155 :: (-> S69 S56 S1)
+decl f237 :: S69
+#2777 := f237
+#25856 := (f155 f237 #25855)
+#25857 := (= #25856 f1)
+#25858 := (not #25857)
+#25863 := (or #25858 #25862)
+#25864 := (not #25863)
+#25854 := (not #25853)
+#25874 := (or #25854 #25864 #25867 #25873)
+#25875 := (not #25874)
+#25884 := (or #25875 #25883)
+#25885 := (not #25884)
+decl f81 :: S31
+#1068 := f81
+#4667 := (f80 f81 f444)
+#25848 := (f71 #4667 #25639)
+#25849 := (= #25848 f1)
+#4941 := (f71 #4667 #4937)
+#4942 := (= #4941 f1)
+#24580 := (f134 #4883 #4937)
+#25782 := (f155 f237 #24580)
+#25783 := (= #25782 f1)
+#17174 := (not #4942)
+#25784 := (or #17174 #25783)
+#25785 := (not #25784)
+#25916 := [hypothesis]: #25784
+decl f50 :: (-> S13 S12 S1)
+decl f65 :: (-> S23 S11 S12)
+#4657 := (f55 f206 #4656)
+decl f215 :: (-> S78 Int S4)
+decl f216 :: (-> S79 S4 S78)
+decl f217 :: S79
+#2593 := f217
+#4651 := (f216 f217 f35)
+#4652 := (f215 #4651 f443)
+#4653 := (f113 f114 #4652)
+#4658 := (f87 #4653 #4657)
+#22490 := (f55 f206 #4658)
+#23413 := (f87 #4653 #22490)
+decl f66 :: (-> S24 S10 S23)
+decl f67 :: S24
+#1018 := f67
+#23811 := (f66 f67 f444)
+#23819 := (f65 #23811 #23413)
+decl f51 :: (-> S14 S11 S13)
+#24084 := (f87 #4653 f445)
+decl f64 :: S14
+#1003 := f64
+#24085 := (f51 f64 #24084)
+#24086 := (f50 #24085 #23819)
+#24087 := (= #24086 f1)
+#23810 := (f51 f64 #23413)
+#23820 := (f50 #23810 #23819)
+#23821 := (= #23820 f1)
+decl f129 :: S24
+#1228 := f129
+#23570 := (f66 f129 f444)
+#23825 := (f65 #23570 #23413)
+#996 := (:var 0 S11)
+#1004 := (f51 f64 #996)
+#23826 := (f50 #1004 #23825)
+#23835 := (pattern #23826)
+decl f329 :: (-> S126 S19 S12)
+decl f58 :: (-> S20 S11 S19)
+decl f59 :: (-> S21 S10 S20)
+decl f60 :: S21
+#991 := f60
+#23829 := (f59 f60 f444)
+#23830 := (f58 #23829 #23413)
+decl f330 :: S126
+#3503 := f330
+#23831 := (f329 f330 #23830)
+decl f254 :: S14
+#2955 := f254
+#3762 := (f51 f254 #996)
+#23832 := (f50 #3762 #23831)
+#23833 := (= #23832 f1)
+#23827 := (= #23826 f1)
+#23828 := (not #23827)
+#23693 := (f62 f63 #23413)
+decl f337 :: S7
+#3683 := f337
+#23823 := (f45 f337 #23693)
+#23824 := (= #23823 f1)
+#23834 := (or #23824 #23828 #23833)
+#23836 := (forall (vars (?v3 S11)) (:pat #23835) #23834)
+#23837 := (not #23836)
+#23704 := (f71 #4650 #23413)
+#23705 := (= #23704 f1)
+#23730 := (not #23705)
+#23822 := (not #23821)
+#23838 := (or #23822 #23730 #23837)
+#23839 := (not #23838)
+decl f125 :: (-> S54 S11 S27)
+decl f334 :: (-> S128 S10 S54)
+decl f336 :: S128
+#3670 := f336
+#23786 := (f334 f336 f444)
+#23787 := (f125 #23786 #23413)
+#23788 := (f71 #23787 #23413)
+#23789 := (= #23788 f1)
+decl f338 :: (-> S130 S129 S1)
+decl f460 :: S129
+#4731 := f460
+decl f339 :: (-> S131 S11 S130)
+decl f340 :: (-> S132 S11 S131)
+decl f341 :: (-> S133 S10 S132)
+decl f345 :: S133
+#3792 := f345
+#4728 := (f341 f345 f444)
+#23775 := (f340 #4728 #23413)
+#23776 := (f339 #23775 #23413)
+#23784 := (f338 #23776 f460)
+#23785 := (= #23784 f1)
+#23790 := (iff #23785 #23789)
+#3776 := (:var 0 S129)
+#984 := (:var 1 S11)
+#993 := (:var 2 S11)
+#980 := (:var 3 S10)
+#3793 := (f341 f345 #980)
+#3794 := (f340 #3793 #993)
+#3795 := (f339 #3794 #984)
+#3796 := (f338 #3795 #3776)
+#3797 := (pattern #3796)
+#3720 := (f334 f336 #980)
+#3799 := (f125 #3720 #993)
+#3800 := (f71 #3799 #984)
+#3801 := (= #3800 f1)
+#3798 := (= #3796 f1)
+#3802 := (iff #3798 #3801)
+#3803 := (forall (vars (?v0 S10) (?v1 S11) (?v2 S11) (?v3 S129)) (:pat #3797) #3802)
+#16565 := (~ #3803 #3803)
+#16563 := (~ #3802 #3802)
+#16564 := [refl]: #16563
+#16566 := [nnf-pos #16564]: #16565
+#10696 := [asserted]: #3803
+#16567 := [mp~ #10696 #16566]: #3803
+#23799 := (not #3803)
+#23801 := (or #23799 #23790)
+#23802 := [quant-inst #4649 #23413 #23413 #4731]: #23801
+#23945 := [unit-resolution #23802 #16567]: #23790
+#4729 := (f340 #4728 #4658)
+#4730 := (f339 #4729 #4658)
+#4732 := (f338 #4730 f460)
+#4733 := (= #4732 f1)
+#23865 := (f61 #23693)
+#23866 := (= #23865 f3)
+#23954 := (not #23866)
+decl f6 :: S2
+#14 := f6
+#15 := (= f3 f6)
+#16 := (not #15)
+#23955 := (iff #16 #23954)
+#23952 := (iff #15 #23866)
+#23950 := (iff #23866 #15)
+#23928 := (= f6 f3)
+#23948 := (iff #23928 #15)
+#23949 := [commutativity]: #23948
+#23929 := (iff #23866 #23928)
+#23939 := (= #23865 f6)
+#4670 := (f61 #4652)
+#23587 := (= #4670 f6)
+decl f248 :: S7
+#2922 := f248
+#23515 := (f45 f248 #4652)
+#23516 := (= #23515 f1)
+#23588 := (iff #23516 #23587)
+#1287 := (:var 0 S4)
+#3295 := (f45 f248 #1287)
+#4521 := (pattern #3295)
+#4530 := (f61 #1287)
+#4534 := (= #4530 f6)
+#3297 := (= #3295 f1)
+#4535 := (iff #3297 #4534)
+#4536 := (forall (vars (?v0 S4)) (:pat #4521) #4535)
+#17000 := (~ #4536 #4536)
+#16998 := (~ #4535 #4535)
+#16999 := [refl]: #16998
+#17001 := [nnf-pos #16999]: #17000
+#11185 := [asserted]: #4536
+#17002 := [mp~ #11185 #17001]: #4536
+#23597 := (not #4536)
+#23598 := (or #23597 #23588)
+#23599 := [quant-inst #4652]: #23598
+#23796 := [unit-resolution #23599 #17002]: #23588
+#23600 := (not #23588)
+#23798 := (or #23600 #23587)
+#1426 := (:var 1 S4)
+#2594 := (f216 f217 #1426)
+#2595 := (f215 #2594 #1197)
+#2917 := (pattern #2595)
+#2923 := (f45 f248 #2595)
+#2924 := (= #2923 f1)
+#2925 := (forall (vars (?v0 S4) (?v1 Int)) (:pat #2917) #2924)
+#16099 := (~ #2925 #2925)
+#16097 := (~ #2924 #2924)
+#16098 := [refl]: #16097
+#16100 := [nnf-pos #16098]: #16099
+#9874 := [asserted]: #2925
+#16101 := [mp~ #9874 #16100]: #2925
+#23522 := (not #2925)
+#23523 := (or #23522 #23516)
+#23524 := [quant-inst #356 #4646]: #23523
+#24987 := [unit-resolution #23524 #16101]: #23516
+#23604 := (not #23516)
+#23605 := (or #23600 #23604 #23587)
+#23606 := [def-axiom]: #23605
+#23914 := [unit-resolution #23606 #24987]: #23798
+#23915 := [unit-resolution #23914 #23796]: #23587
+#23937 := (= #23865 #4670)
+#23935 := (= #23693 #4652)
+#23428 := (f62 f63 #4658)
+#23429 := (= #23428 #4652)
+#2667 := (f113 f114 #1426)
+#4356 := (f87 #2667 #1197)
+#21823 := (pattern #4356)
+#4360 := (f62 f63 #4356)
+#4361 := (= #4360 #1426)
+#21830 := (forall (vars (?v0 S4) (?v1 Int)) (:pat #21823) #4361)
+#4362 := (forall (vars (?v0 S4) (?v1 Int)) #4361)
+#21833 := (iff #4362 #21830)
+#21831 := (iff #4361 #4361)
+#21832 := [refl]: #21831
+#21834 := [quant-intro #21832]: #21833
+#16915 := (~ #4362 #4362)
+#16913 := (~ #4361 #4361)
+#16914 := [refl]: #16913
+#16916 := [nnf-pos #16914]: #16915
+#11104 := [asserted]: #4362
+#16917 := [mp~ #11104 #16916]: #4362
+#21835 := [mp #16917 #21834]: #21830
+#23455 := (not #21830)
+#23494 := (or #23455 #23429)
+#23495 := [quant-inst #4652 #4657]: #23494
+#23916 := [unit-resolution #23495 #21835]: #23429
+#23933 := (= #23693 #23428)
+#23931 := (= #23413 #4658)
+#23426 := (= #4658 #23413)
+#4664 := (f118 f123 #4658)
+#4665 := (f45 #4664 #4652)
+#4666 := (= #4665 f1)
+decl f79 :: S7
+#1064 := f79
+#4673 := (f45 f79 #4652)
+#4674 := (= #4673 f1)
+#4671 := (= #4670 f3)
+#4672 := (not #4671)
+#4668 := (f71 #4667 #4658)
+#4669 := (= #4668 f1)
+#4662 := (f82 #4661 #4658)
+#4663 := (= #4662 f85)
+#4659 := (f71 #4650 #4658)
+#4660 := (= #4659 f1)
+#13334 := (and #4660 #4663 #4666 #4669 #4672 #4674)
+decl f468 :: Int
+#4819 := f468
+#4826 := (= #4775 f468)
+#12568 := (* -1::Int f443)
+#12951 := (+ #1197 #12568)
+#12950 := (>= #12951 0::Int)
+#12952 := (not #12950)
+decl f168 :: Int
+#1519 := f168
+#6888 := (* -1::Int f168)
+#6889 := (+ #1197 #6888)
+#6890 := (<= #6889 0::Int)
+#12993 := (and #6706 #6890 #12952 #4826)
+#12998 := (exists (vars (?v0 Int)) #12993)
+#12962 := (* -1::Int f468)
+#12963 := (+ #4775 #12962)
+#12964 := (<= #12963 0::Int)
+#6897 := (and #6706 #6890)
+#7910 := (not #6897)
+#12973 := (or #7910 #12950 #12964)
+#12978 := (forall (vars (?v0 Int)) #12973)
+#12981 := (not #12978)
+#13001 := (or #12981 #12998)
+#13004 := (and #12978 #13001)
+decl f462 :: Int
+#4782 := f462
+#4820 := (= f468 f462)
+#11361 := (not #4820)
+decl f463 :: Int
+#4786 := f463
+decl f467 :: Int
+#4817 := f467
+#4818 := (= f467 f463)
+#11370 := (not #4818)
+decl f466 :: Int
+#4815 := f466
+#4816 := (= f466 f464)
+#11379 := (not #4816)
+decl f465 :: Int
+#4813 := f465
+#4814 := (= f465 f462)
+#11388 := (not #4814)
+#12642 := (>= f463 0::Int)
+#12644 := (and #12639 #12642)
+#12647 := (not #12644)
+decl f367 :: S1
+#4071 := f367
+#4072 := (= f367 f1)
+#11436 := (not #4072)
+#13031 := (or #11436 #12647 #11388 #11379 #11370 #11361 #13004)
+#13036 := (and #4072 #13031)
+#12663 := (* -1::Int f464)
+#12921 := (+ f443 #12663)
+#12922 := (<= #12921 0::Int)
+#12923 := (not #12922)
+#13061 := (or #12647 #12923 #13036)
+#12721 := (not #12720)
+#12724 := (and #12721 #4991)
+#12707 := (or #7910 #12684 #12698)
+#12712 := (forall (vars (?v0 Int)) #12707)
+#12715 := (not #12712)
+#12727 := (or #12715 #12724)
+#12730 := (and #12712 #12727)
+#12733 := (or #12681 #12730)
+#12736 := (and #12676 #12733)
+#12670 := (and #12668 #12651)
+#12673 := (not #12670)
+#12664 := (+ f168 #12663)
+#12662 := (>= #12664 1::Int)
+#12746 := (and #12660 #12662)
+#12749 := (not #12746)
+#12764 := (or #12749 #12743 #11705 #12673 #12736)
+#12772 := (and #12660 #12662 #12764)
+#12653 := (and #12639 #12651)
+#12656 := (not #12653)
+#5027 := (= f471 f463)
+#11883 := (not #5027)
+#5026 := (= f470 f462)
+#11892 := (not #5026)
+#12830 := (* -1::Int #4947)
+#12831 := (+ f462 #12830)
+#12829 := (>= #12831 0::Int)
+#12828 := (not #12829)
+#12883 := (or #12647 #12828 #11892 #11883 #12656 #12772)
+#4946 := (and #4940 #4945)
+#11812 := (not #4946)
+#12804 := (or #11812 #11803 #11794 #11785 #12777 #11760 #11751 #12656 #12772)
+#12812 := (and #4940 #4945 #12804)
+#4943 := (and #4940 #4942)
+#11824 := (not #4943)
+#12817 := (or #11824 #12812)
+#12823 := (and #4940 #4942 #12817)
+#12853 := (or #12647 #12829 #12823)
+#12888 := (and #12853 #12883)
+#12897 := (or #11812 #12647 #12888)
+#12905 := (and #4940 #4945 #12897)
+#12910 := (or #11824 #12905)
+#12916 := (and #4940 #4942 #12910)
+#12945 := (or #12647 #12922 #12916)
+#13066 := (and #12945 #13061)
+decl f48 :: (-> S9 S4 S4)
+decl f49 :: S9
+#976 := f49
+#977 := (f48 f49 f35)
+decl f453 :: (-> S181 S3 S51)
+decl f19 :: S3
+#63 := f19
+decl f454 :: (-> S182 S3 S181)
+decl f13 :: S3
+#45 := f13
+decl f455 :: (-> S183 S10 S182)
+decl f456 :: S183
+#4703 := f456
+#4704 := (f455 f456 f444)
+#4926 := (f454 #4704 f13)
+#4927 := (f453 #4926 f19)
+#4928 := (f118 #4927 #4656)
+#4929 := (f45 #4928 #977)
+#4930 := (= #4929 f1)
+decl f89 :: S17
+#1094 := f89
+#4699 := (f55 f89 #4656)
+#4905 := (f450 #4696 f13)
+#4922 := (f449 #4905 f19)
+#4923 := (f46 #4922 #4699)
+#4924 := (f45 #4923 #977)
+#4925 := (= #4924 f1)
+#4931 := (and #4925 #4930)
+#12108 := (not #4931)
+decl f18 :: S3
+#60 := f18
+#4918 := (f449 #4905 f18)
+#4919 := (f46 #4918 f443)
+#4920 := (f45 #4919 f33)
+#4921 := (= #4920 f1)
+#12117 := (not #4921)
+#4914 := (f449 #4905 f17)
+#4915 := (f46 #4914 f462)
+#4916 := (f45 #4915 f35)
+#4917 := (= #4916 f1)
+#12126 := (not #4917)
+#4910 := (f449 #4905 f10)
+#4911 := (f46 #4910 f463)
+#4912 := (f45 #4911 f33)
+#4913 := (= #4912 f1)
+#12135 := (not #4913)
+#4906 := (f449 #4905 f20)
+#4907 := (f46 #4906 f464)
+#4908 := (f45 #4907 f33)
+#4909 := (= #4908 f1)
+#12144 := (not #4909)
+decl f115 :: (-> S50 S10 S1)
+decl f131 :: S50
+#1279 := f131
+#4685 := (f115 f131 f444)
+#4686 := (= #4685 f1)
+decl f348 :: (-> S136 S3 S50)
+decl f349 :: S136
+#3828 := f349
+#4809 := (f348 f349 f13)
+#4810 := (f115 #4809 f444)
+#4811 := (= #4810 f1)
+#4812 := (and #4811 #4686)
+#11471 := (not #4812)
+decl f304 :: (-> S115 S10 S50)
+decl f305 :: S115
+#3261 := f305
+#4896 := (f304 f305 f444)
+#4897 := (f115 #4896 f444)
+#4898 := (= #4897 f1)
+#13090 := (not #4898)
+#4803 := (f140 #4734 f463)
+#4804 := (f139 #4803 f35)
+#4805 := (f55 #4748 #4804)
+#4806 := (= #4805 f462)
+#13093 := (* -1::Int f463)
+#13094 := (+ f443 #13093)
+#13095 := (<= #13094 0::Int)
+#13096 := (not #13095)
+#13099 := (and #13096 #4806)
+#13102 := (not #13099)
+#13117 := (* -1::Int f462)
+#13118 := (+ #4775 #13117)
+#13119 := (<= #13118 0::Int)
+#13106 := (+ #1197 #12663)
+#13105 := (>= #13106 0::Int)
+#13128 := (or #7910 #13105 #13119)
+#13133 := (forall (vars (?v0 Int)) #13128)
+#13136 := (not #13133)
+#13139 := (>= #12921 0::Int)
+#13142 := (not #13139)
+#13148 := (>= #12664 0::Int)
+#13145 := (>= f464 0::Int)
+#13151 := (and #13145 #13148)
+#13154 := (not #13151)
+#13158 := (+ f168 #13093)
+#13157 := (>= #13158 0::Int)
+#13161 := (and #12642 #13157)
+#13164 := (not #13161)
+decl f170 :: Int
+#1539 := f170
+#13171 := (+ f170 #13117)
+#13170 := (>= #13171 0::Int)
+#13167 := (>= f462 0::Int)
+#13174 := (and #13167 #13170)
+#13177 := (not #13174)
+decl f461 :: Int
+#4747 := f461
+#4749 := (f55 #4748 #4736)
+#4780 := (= #4749 f461)
+#12634 := (<= f443 0::Int)
+#12635 := (not #12634)
+#13180 := (and #12635 #4780)
+#13183 := (not #13180)
+#13249 := (or #13183 #13177 #13164 #13154 #12647 #13142 #13136 #13102 #13090 #11471 #12144 #12135 #12126 #12117 #12108 #13066)
+#13257 := (and #12635 #4780 #13249)
+#12614 := (* -1::Int #4775)
+#12615 := (+ f461 #12614)
+#12613 := (>= #12615 0::Int)
+#12601 := (>= #1197 1::Int)
+#12623 := (or #7910 #12601 #12613)
+#12628 := (forall (vars (?v0 Int)) #12623)
+#12631 := (not #12628)
+#13262 := (or #12631 #13257)
+#13265 := (and #12628 #13262)
+#12595 := (>= f443 1::Int)
+#12598 := (not #12595)
+#13268 := (or #12598 #13265)
+#13271 := (and #12595 #13268)
+decl f12 :: S3
+#42 := f12
+#4761 := (f450 #4696 f12)
+#4762 := (f449 #4761 f20)
+#4763 := (f46 #4762 1::Int)
+#4764 := (f45 #4763 f33)
+#4765 := (= #4764 f1)
+#12352 := (not #4765)
+decl f14 :: S3
+#48 := f14
+#4756 := (f450 #4696 f14)
+#4757 := (f449 #4756 f10)
+#4758 := (f46 #4757 0::Int)
+#4759 := (f45 #4758 f33)
+#4760 := (= #4759 f1)
+#12361 := (not #4760)
+decl f15 :: S3
+#51 := f15
+#4751 := (f450 #4696 f15)
+#4752 := (f449 #4751 f17)
+#4753 := (f46 #4752 f461)
+#4754 := (f45 #4753 f35)
+#4755 := (= #4754 f1)
+#12370 := (not #4755)
+#4750 := (= f461 #4749)
+#12379 := (not #4750)
+#4744 := (f71 #4743 #4736)
+#4745 := (= #4744 f1)
+#4737 := (f118 f123 #4736)
+#4738 := (f45 #4737 f35)
+#4739 := (= #4738 f1)
+#4746 := (and #4739 #4745)
+#12388 := (not #4746)
+#13292 := (or #12388 #12379 #12370 #12361 #12352 #13271)
+#13300 := (and #4739 #4745 #13292)
+#4740 := (f71 #4667 #4736)
+#4741 := (= #4740 f1)
+#4742 := (and #4739 #4741)
+#12400 := (not #4742)
+#13305 := (or #12400 #13300)
+#13311 := (and #4739 #4741 #13305)
+#12412 := (not #4733)
+#13316 := (or #12412 #13311)
+#13319 := (and #4733 #13316)
+#12569 := (+ f168 #12568)
+#12567 := (>= #12569 0::Int)
+#12565 := (>= f443 0::Int)
+#12572 := (and #12565 #12567)
+#12575 := (not #12572)
+decl f458 :: (-> S184 Int S27)
+decl f457 :: Int
+#4715 := f457
+decl f459 :: S184
+#4718 := f459
+#4719 := (f458 f459 f457)
+#4720 := (f71 #4719 #996)
+#4721 := (pattern #4720)
+#4722 := (= #4720 f1)
+#11269 := (not #4722)
+#11272 := (forall (vars (?v0 S11)) (:pat #4721) #11269)
+#12433 := (not #11272)
+decl f292 :: (-> S108 S10 Int)
+decl f293 :: S108
+#3194 := f293
+#4716 := (f292 f293 f444)
+#4717 := (= f457 #4716)
+#12442 := (not #4717)
+decl f16 :: S3
+#54 := f16
+#4697 := (f450 #4696 f16)
+#4711 := (f449 #4697 f18)
+#4712 := (f46 #4711 f443)
+#4713 := (f45 #4712 f33)
+#4714 := (= #4713 f1)
+#12451 := (not #4714)
+#4705 := (f454 #4704 f16)
+#4706 := (f453 #4705 f19)
+#4707 := (f118 #4706 #4656)
+#4708 := (f45 #4707 #977)
+#4709 := (= #4708 f1)
+#4698 := (f449 #4697 f19)
+#4700 := (f46 #4698 #4699)
+#4701 := (f45 #4700 #977)
+#4702 := (= #4701 f1)
+#4710 := (and #4702 #4709)
+#12460 := (not #4710)
+decl f446 :: (-> S177 S176 Int)
+#4689 := (:var 0 S176)
+decl f447 :: S177
+#4688 := f447
+#4690 := (f446 f447 #4689)
+#4691 := (pattern #4690)
+decl f448 :: Int
+#4692 := f448
+#13324 := (* -1::Int f448)
+#13325 := (+ #4690 #13324)
+#13323 := (>= #13325 0::Int)
+#13322 := (not #13323)
+#13328 := (forall (vars (?v0 S176)) (:pat #4691) #13322)
+#13331 := (not #13328)
+#4682 := (f348 f349 f16)
+#4683 := (f115 #4682 f444)
+#4684 := (= #4683 f1)
+#4687 := (and #4684 #4686)
+#12478 := (not #4687)
+decl f350 :: S50
+#3847 := f350
+#4680 := (f115 f350 f444)
+#4681 := (= #4680 f1)
+#12487 := (not #4681)
+#13337 := (not #13334)
+#2248 := 1099511627776::Int
+#13347 := (>= f443 1099511627776::Int)
+decl f442 :: Int
+#4642 := f442
+#13362 := (* -1::Int f442)
+#13363 := (+ f168 #13362)
+#13361 := (>= #13363 0::Int)
+#13359 := (>= f442 0::Int)
+#13366 := (and #13359 #13361)
+#13369 := (not #13366)
+decl f441 :: Int
+#4638 := f441
+#13376 := (* -1::Int f441)
+#13377 := (+ f168 #13376)
+#13375 := (>= #13377 0::Int)
+#13373 := (>= f441 0::Int)
+#13380 := (and #13373 #13375)
+#13383 := (not #13380)
+decl f440 :: Int
+#4634 := f440
+#13390 := (* -1::Int f440)
+#13391 := (+ f170 #13390)
+#13389 := (>= #13391 0::Int)
+#13387 := (>= f440 0::Int)
+#13394 := (and #13387 #13389)
+#13397 := (not #13394)
+#13442 := (or #13397 #13383 #13369 #13347 #12634 #13337 #12487 #12478 #13331 #12460 #12451 #12442 #12433 #12575 #13319)
+#13447 := (not #13442)
+#1 := true
+#4821 := (< #1197 f443)
+#4827 := (and #4821 #4826)
+#1521 := (<= #1197 f168)
+#4828 := (and #1521 #4827)
+#1363 := (<= 0::Int #1197)
+#4829 := (and #1363 #4828)
+#4830 := (exists (vars (?v0 Int)) #4829)
+#4831 := (implies #4830 true)
+#4832 := (and #4830 #4831)
+#4822 := (<= #4775 f468)
+#4823 := (implies #4821 #4822)
+#1522 := (and #1363 #1521)
+#4824 := (implies #1522 #4823)
+#4825 := (forall (vars (?v0 Int)) #4824)
+#4833 := (implies #4825 #4832)
+#4834 := (and #4825 #4833)
+#4835 := (implies true #4834)
+#4836 := (implies #4820 #4835)
+#4837 := (implies #4818 #4836)
+#4838 := (implies #4816 #4837)
+#4839 := (implies #4814 #4838)
+#4840 := (implies true #4839)
+#4787 := (<= 0::Int f463)
+#4794 := (<= 1::Int f464)
+#4795 := (and #4794 #4787)
+#4841 := (implies #4795 #4840)
+#4842 := (implies #4795 #4841)
+#4843 := (implies true #4842)
+#4844 := (implies #4795 #4843)
+#4845 := (implies #4072 #4844)
+#4846 := (and #4072 #4845)
+#4847 := (implies #4795 #4846)
+#4848 := (implies true #4847)
+#4849 := (implies #4795 #4848)
+#5051 := (implies #4795 #4849)
+#5052 := (implies true #5051)
+#5053 := (implies #4795 #5052)
+#5050 := (<= f443 f464)
+#5054 := (implies #5050 #5053)
+#5055 := (implies #4795 #5054)
+#5056 := (implies true #5055)
+#4993 := (implies false true)
+#4987 := (< f471 f443)
+#4992 := (and #4987 #4991)
+#4994 := (implies #4992 #4993)
+#4995 := (and #4992 #4994)
+#4983 := (<= #4775 f470)
+#4982 := (< #1197 f472)
+#4984 := (implies #4982 #4983)
+#4985 := (implies #1522 #4984)
+#4986 := (forall (vars (?v0 Int)) #4985)
+#4996 := (implies #4986 #4995)
+#4997 := (and #4986 #4996)
+#4981 := (<= f472 f443)
+#4998 := (implies #4981 #4997)
+#4999 := (and #4981 #4998)
+#5000 := (implies true #4999)
+#4966 := (<= 0::Int f471)
+#4979 := (<= 2::Int f472)
+#4980 := (and #4979 #4966)
+#5001 := (implies #4980 #5000)
+#5002 := (implies #4978 #5001)
+#4968 := (+ f464 1::Int)
+#4973 := (= f472 #4968)
+#5003 := (implies #4973 #5002)
+#4970 := (<= #4968 f168)
+#4969 := (<= 0::Int #4968)
+#4971 := (and #4969 #4970)
+#5004 := (implies #4971 #5003)
+#5005 := (and #4971 #5004)
+#4967 := (and #4794 #4966)
+#5006 := (implies #4967 #5005)
+#5007 := (implies true #5006)
+#5028 := (implies #5027 #5007)
+#5029 := (implies #5026 #5028)
+#5030 := (implies true #5029)
+#5031 := (implies #4795 #5030)
+#5032 := (implies #4795 #5031)
+#5033 := (implies true #5032)
+#5034 := (implies #4795 #5033)
+#5025 := (<= #4947 f462)
+#5035 := (implies #5025 #5034)
+#5036 := (implies #4795 #5035)
+#5037 := (implies true #5036)
+#5008 := (implies #4965 #5007)
+#5009 := (implies #4963 #5008)
+#5010 := (implies true #5009)
+#4961 := (and #4794 #4794)
+#5011 := (implies #4961 #5010)
+#5012 := (implies #4960 #5011)
+#5013 := (implies #4955 #5012)
+#5014 := (implies #4950 #5013)
+#5015 := (implies #4946 #5014)
+#5016 := (and #4946 #5015)
+#5017 := (implies #4943 #5016)
+#5018 := (and #4943 #5017)
+#5019 := (implies #4795 #5018)
+#5020 := (implies true #5019)
+#5021 := (implies #4795 #5020)
+#4948 := (< f462 #4947)
+#5022 := (implies #4948 #5021)
+#5023 := (implies #4795 #5022)
+#5024 := (implies true #5023)
+#5038 := (and #5024 #5037)
+#5039 := (implies #4795 #5038)
+#5040 := (implies #4946 #5039)
+#5041 := (and #4946 #5040)
+#5042 := (implies #4943 #5041)
+#5043 := (and #4943 #5042)
+#5044 := (implies #4795 #5043)
+#5045 := (implies true #5044)
+#5046 := (implies #4795 #5045)
+#4935 := (< f464 f443)
+#5047 := (implies #4935 #5046)
+#5048 := (implies #4795 #5047)
+#5049 := (implies true #5048)
+#5057 := (and #5049 #5056)
+#5058 := (implies #4795 #5057)
+decl f110 :: (-> S48 S10 S47)
+decl f111 :: S48
+#1128 := f111
+#4857 := (f110 f111 f444)
+#4933 := (= #4857 #4857)
+#4932 := (= #4882 #4882)
+#4934 := (and #4932 #4933)
+#5059 := (implies #4934 #5058)
+#5060 := (implies #4931 #5059)
+#5061 := (implies #4921 #5060)
+#5062 := (implies #4917 #5061)
+#5063 := (implies #4913 #5062)
+#5064 := (implies #4909 #5063)
+#5065 := (implies #4812 #5064)
+decl f291 :: S42
+#3191 := f291
+#4891 := (f103 f291 f444)
+#4892 := (f55 #4891 #996)
+#4893 := (pattern #4892)
+#4894 := (<= #4892 #4892)
+#4895 := (forall (vars (?v0 S11)) (:pat #4893) #4894)
+#4899 := (and #4895 #4898)
+#4890 := (<= #4716 #4716)
+#4900 := (and #4890 #4899)
+#5066 := (implies #4900 #5065)
+#4884 := (f134 #4883 #996)
+#4885 := (pattern #4884)
+#4872 := (f71 #4743 #996)
+#4873 := (= #4872 f1)
+#4886 := (= #4884 #4884)
+#4887 := (and #4886 #4873)
+#4888 := (implies #4873 #4887)
+#4889 := (forall (vars (?v0 S11)) (:pat #4885) #4888)
+#4901 := (and #4889 #4900)
+decl f107 :: (-> S45 S11 S44)
+decl f108 :: (-> S46 S47 S45)
+decl f109 :: S46
+#1127 := f109
+#4858 := (f108 f109 #4857)
+#4859 := (f107 #4858 #996)
+#4860 := (pattern #4859)
+#4878 := (= #4859 #4859)
+#4879 := (and #4878 #4873)
+#4880 := (implies #4873 #4879)
+#4881 := (forall (vars (?v0 S11)) (:pat #4860) #4880)
+#4902 := (and #4881 #4901)
+decl f73 :: (-> S28 S29 S17)
+decl f75 :: (-> S30 S10 S29)
+decl f76 :: S30
+#1039 := f76
+#4868 := (f75 f76 f444)
+decl f74 :: S28
+#1038 := f74
+#4869 := (f73 f74 #4868)
+#4870 := (f55 #4869 #996)
+#4871 := (pattern #4870)
+#4874 := (= #4870 #4870)
+#4875 := (and #4874 #4873)
+#4876 := (implies #4873 #4875)
+#4877 := (forall (vars (?v0 S11)) (:pat #4871) #4876)
+#4903 := (and #4877 #4902)
+decl f5 :: S2
+#11 := f5
+#4861 := (f82 #4661 #996)
+#4862 := (f62 f63 #4861)
+#4863 := (f61 #4862)
+#4864 := (= #4863 f5)
+#4865 := (not #4864)
+#4866 := (implies #4865 #4865)
+#4867 := (forall (vars (?v0 S11)) (:pat #4860) #4866)
+#4904 := (and #4867 #4903)
+#5067 := (implies #4904 #5066)
+#5068 := (implies #4795 #5067)
+#5069 := (implies true #5068)
+#5070 := (implies #4795 #5069)
+#5071 := (implies true #5070)
+#5072 := (implies #4795 #5071)
+#5073 := (implies true #5072)
+#4850 := (implies #4812 #4849)
+#4851 := (implies #4795 #4850)
+#4852 := (implies true #4851)
+#4853 := (implies #4795 #4852)
+#4808 := (not true)
+#4854 := (implies #4808 #4853)
+#4855 := (implies #4795 #4854)
+#4856 := (implies true #4855)
+#5074 := (and #4856 #5073)
+#5075 := (implies #4795 #5074)
+#4802 := (< f463 f443)
+#4807 := (and #4802 #4806)
+#5076 := (implies #4807 #5075)
+#4798 := (<= #4775 f462)
+#4797 := (< #1197 f464)
+#4799 := (implies #4797 #4798)
+#4800 := (implies #1522 #4799)
+#4801 := (forall (vars (?v0 Int)) #4800)
+#5077 := (implies #4801 #5076)
+#4796 := (<= f464 f443)
+#5078 := (implies #4796 #5077)
+#5079 := (implies #4795 #5078)
+#4792 := (<= f464 f168)
+#4791 := (<= 0::Int f464)
+#4793 := (and #4791 #4792)
+#5080 := (implies #4793 #5079)
+#4788 := (<= f463 f168)
+#4789 := (and #4787 #4788)
+#5081 := (implies #4789 #5080)
+#4784 := (<= f462 f170)
+#4783 := (<= 0::Int f462)
+#4785 := (and #4783 #4784)
+#5082 := (implies #4785 #5081)
+#5083 := (implies true #5082)
+#4648 := (< 0::Int f443)
+#4781 := (and #4648 #4780)
+#5084 := (implies #4781 #5083)
+#5085 := (and #4781 #5084)
+#4776 := (<= #4775 f461)
+#4772 := (< #1197 1::Int)
+#4777 := (implies #4772 #4776)
+#4778 := (implies #1522 #4777)
+#4779 := (forall (vars (?v0 Int)) #4778)
+#5086 := (implies #4779 #5085)
+#5087 := (and #4779 #5086)
+#4771 := (<= 1::Int f443)
+#5088 := (implies #4771 #5087)
+#5089 := (and #4771 #5088)
+#4767 := (<= 0::Int 0::Int)
+#4768 := (and #4767 #4767)
+#4766 := (<= 1::Int 1::Int)
+#4769 := (and #4766 #4768)
+#4770 := (and #4766 #4769)
+#5090 := (implies #4770 #5089)
+#5091 := (implies #4765 #5090)
+#5092 := (implies #4760 #5091)
+#5093 := (implies #4755 #5092)
+#5094 := (implies #4750 #5093)
+#5095 := (implies #4746 #5094)
+#5096 := (and #4746 #5095)
+#5097 := (implies #4742 #5096)
+#5098 := (and #4742 #5097)
+#5099 := (implies #4733 #5098)
+#5100 := (and #4733 #5099)
+#4726 := (<= f443 f168)
+#4725 := (<= 0::Int f443)
+#4727 := (and #4725 #4726)
+#5101 := (implies #4727 #5100)
+#4723 := (iff #4722 false)
+#4724 := (forall (vars (?v0 S11)) (:pat #4721) #4723)
+#5102 := (implies #4724 #5101)
+#5103 := (implies #4717 #5102)
+#5104 := (implies #4714 #5103)
+#5105 := (implies #4710 #5104)
+#4693 := (< #4690 f448)
+#4694 := (forall (vars (?v0 S176)) (:pat #4691) #4693)
+#5106 := (implies #4694 #5105)
+#5107 := (implies #4687 #5106)
+#5108 := (implies #4681 #5107)
+#5109 := (implies true #5108)
+#4675 := (and #4672 #4674)
+#4676 := (and #4669 #4675)
+#4677 := (and #4666 #4676)
+#4678 := (and #4663 #4677)
+#4679 := (and #4660 #4678)
+#5110 := (implies #4679 #5109)
+#5111 := (implies #4648 #5110)
+#4647 := (< f443 1099511627776::Int)
+#5112 := (implies #4647 #5111)
+#4644 := (<= f442 f168)
+#4643 := (<= 0::Int f442)
+#4645 := (and #4643 #4644)
+#5113 := (implies #4645 #5112)
+#4640 := (<= f441 f168)
+#4639 := (<= 0::Int f441)
+#4641 := (and #4639 #4640)
+#5114 := (implies #4641 #5113)
+#4636 := (<= f440 f170)
+#4635 := (<= 0::Int f440)
+#4637 := (and #4635 #4636)
+#5115 := (implies #4637 #5114)
+#5116 := (implies true #5115)
+#5117 := (not #5116)
+#13450 := (iff #5117 #13447)
+#11315 := (not #4821)
+#11316 := (or #11315 #4822)
+#7903 := (not #1522)
+#11322 := (or #7903 #11316)
+#11327 := (forall (vars (?v0 Int)) #11322)
+#11342 := (not #11327)
+#11343 := (or #11342 #4830)
+#11348 := (and #11327 #11343)
+#11362 := (or #11361 #11348)
+#11371 := (or #11370 #11362)
+#11380 := (or #11379 #11371)
+#11389 := (or #11388 #11380)
+#11404 := (not #4795)
+#11405 := (or #11404 #11389)
+#11413 := (or #11404 #11405)
+#11428 := (or #11404 #11413)
+#11437 := (or #11436 #11428)
+#11442 := (and #4072 #11437)
+#11448 := (or #11404 #11442)
+#11463 := (or #11404 #11448)
+#12043 := (or #11404 #11463)
+#12058 := (or #11404 #12043)
+#12066 := (not #5050)
+#12067 := (or #12066 #12058)
+#12075 := (or #11404 #12067)
+#11631 := (not #4982)
+#11632 := (or #11631 #4983)
+#11638 := (or #7903 #11632)
+#11643 := (forall (vars (?v0 Int)) #11638)
+#11665 := (not #11643)
+#11666 := (or #11665 #4992)
+#11671 := (and #11643 #11666)
+#11677 := (not #4981)
+#11678 := (or #11677 #11671)
+#11683 := (and #4981 #11678)
+#11696 := (not #4980)
+#11697 := (or #11696 #11683)
+#11706 := (or #11705 #11697)
+#11616 := (+ 1::Int f464)
+#11628 := (= f472 #11616)
+#11714 := (not #11628)
+#11715 := (or #11714 #11706)
+#11622 := (<= #11616 f168)
+#11619 := (<= 0::Int #11616)
+#11625 := (and #11619 #11622)
+#11723 := (not #11625)
+#11724 := (or #11723 #11715)
+#11729 := (and #11625 #11724)
+#11735 := (not #4967)
+#11736 := (or #11735 #11729)
+#11884 := (or #11883 #11736)
+#11893 := (or #11892 #11884)
+#11908 := (or #11404 #11893)
+#11916 := (or #11404 #11908)
+#11931 := (or #11404 #11916)
+#11939 := (not #5025)
+#11940 := (or #11939 #11931)
+#11948 := (or #11404 #11940)
+#11752 := (or #11751 #11736)
+#11761 := (or #11760 #11752)
+#11776 := (not #4794)
+#11777 := (or #11776 #11761)
+#11786 := (or #11785 #11777)
+#11795 := (or #11794 #11786)
+#11804 := (or #11803 #11795)
+#11813 := (or #11812 #11804)
+#11818 := (and #4946 #11813)
+#11825 := (or #11824 #11818)
+#11830 := (and #4943 #11825)
+#11836 := (or #11404 #11830)
+#11851 := (or #11404 #11836)
+#11859 := (not #4948)
+#11860 := (or #11859 #11851)
+#11868 := (or #11404 #11860)
+#11960 := (and #11868 #11948)
+#11966 := (or #11404 #11960)
+#11974 := (or #11812 #11966)
+#11979 := (and #4946 #11974)
+#11985 := (or #11824 #11979)
+#11990 := (and #4943 #11985)
+#11996 := (or #11404 #11990)
+#12011 := (or #11404 #11996)
+#12019 := (not #4935)
+#12020 := (or #12019 #12011)
+#12028 := (or #11404 #12020)
+#12087 := (and #12028 #12075)
+#12093 := (or #11404 #12087)
+#12109 := (or #12108 #12093)
+#12118 := (or #12117 #12109)
+#12127 := (or #12126 #12118)
+#12136 := (or #12135 #12127)
+#12145 := (or #12144 #12136)
+#12153 := (or #11471 #12145)
+#12161 := (not #4900)
+#12162 := (or #12161 #12153)
+#12170 := (or #12161 #12162)
+#12178 := (or #11404 #12170)
+#12193 := (or #11404 #12178)
+#12208 := (or #11404 #12193)
+#12230 := (or #11404 #12208)
+#12238 := (not #4807)
+#12239 := (or #12238 #12230)
+#11298 := (not #4797)
+#11299 := (or #11298 #4798)
+#11305 := (or #7903 #11299)
+#11310 := (forall (vars (?v0 Int)) #11305)
+#12247 := (not #11310)
+#12248 := (or #12247 #12239)
+#12256 := (not #4796)
+#12257 := (or #12256 #12248)
+#12265 := (or #11404 #12257)
+#12273 := (not #4793)
+#12274 := (or #12273 #12265)
+#12282 := (not #4789)
+#12283 := (or #12282 #12274)
+#12291 := (not #4785)
+#12292 := (or #12291 #12283)
+#12307 := (not #4781)
+#12308 := (or #12307 #12292)
+#12313 := (and #4781 #12308)
+#11283 := (not #4772)
+#11284 := (or #11283 #4776)
+#11290 := (or #7903 #11284)
+#11295 := (forall (vars (?v0 Int)) #11290)
+#12319 := (not #11295)
+#12320 := (or #12319 #12313)
+#12325 := (and #11295 #12320)
+#12331 := (not #4771)
+#12332 := (or #12331 #12325)
+#12337 := (and #4771 #12332)
+#11277 := (and #4766 #4767)
+#11280 := (and #4766 #11277)
+#12343 := (not #11280)
+#12344 := (or #12343 #12337)
+#12353 := (or #12352 #12344)
+#12362 := (or #12361 #12353)
+#12371 := (or #12370 #12362)
+#12380 := (or #12379 #12371)
+#12389 := (or #12388 #12380)
+#12394 := (and #4746 #12389)
+#12401 := (or #12400 #12394)
+#12406 := (and #4742 #12401)
+#12413 := (or #12412 #12406)
+#12418 := (and #4733 #12413)
+#12424 := (not #4727)
+#12425 := (or #12424 #12418)
+#12434 := (or #12433 #12425)
+#12443 := (or #12442 #12434)
+#12452 := (or #12451 #12443)
+#12461 := (or #12460 #12452)
+#12469 := (not #4694)
+#12470 := (or #12469 #12461)
+#12479 := (or #12478 #12470)
+#12488 := (or #12487 #12479)
+#12503 := (not #4679)
+#12504 := (or #12503 #12488)
+#12512 := (not #4648)
+#12513 := (or #12512 #12504)
+#12521 := (not #4647)
+#12522 := (or #12521 #12513)
+#12530 := (not #4645)
+#12531 := (or #12530 #12522)
+#12539 := (not #4641)
+#12540 := (or #12539 #12531)
+#12548 := (not #4637)
+#12549 := (or #12548 #12540)
+#12561 := (not #12549)
+#13448 := (iff #12561 #13447)
+#13445 := (iff #12549 #13442)
+#13400 := (or #12575 #13319)
+#13403 := (or #12433 #13400)
+#13406 := (or #12442 #13403)
+#13409 := (or #12451 #13406)
+#13412 := (or #12460 #13409)
+#13415 := (or #13331 #13412)
+#13418 := (or #12478 #13415)
+#13421 := (or #12487 #13418)
+#13424 := (or #13337 #13421)
+#13427 := (or #12634 #13424)
+#13430 := (or #13347 #13427)
+#13433 := (or #13369 #13430)
+#13436 := (or #13383 #13433)
+#13439 := (or #13397 #13436)
+#13443 := (iff #13439 #13442)
+#13444 := [rewrite]: #13443
+#13440 := (iff #12549 #13439)
+#13437 := (iff #12540 #13436)
+#13434 := (iff #12531 #13433)
+#13431 := (iff #12522 #13430)
+#13428 := (iff #12513 #13427)
+#13425 := (iff #12504 #13424)
+#13422 := (iff #12488 #13421)
+#13419 := (iff #12479 #13418)
+#13416 := (iff #12470 #13415)
+#13413 := (iff #12461 #13412)
+#13410 := (iff #12452 #13409)
+#13407 := (iff #12443 #13406)
+#13404 := (iff #12434 #13403)
+#13401 := (iff #12425 #13400)
+#13320 := (iff #12418 #13319)
+#13317 := (iff #12413 #13316)
+#13314 := (iff #12406 #13311)
+#13308 := (and #4742 #13305)
+#13312 := (iff #13308 #13311)
+#13313 := [rewrite]: #13312
+#13309 := (iff #12406 #13308)
+#13306 := (iff #12401 #13305)
+#13303 := (iff #12394 #13300)
+#13297 := (and #4746 #13292)
+#13301 := (iff #13297 #13300)
+#13302 := [rewrite]: #13301
+#13298 := (iff #12394 #13297)
+#13295 := (iff #12389 #13292)
+#13274 := (or false #13271)
+#13277 := (or #12352 #13274)
+#13280 := (or #12361 #13277)
+#13283 := (or #12370 #13280)
+#13286 := (or #12379 #13283)
+#13289 := (or #12388 #13286)
+#13293 := (iff #13289 #13292)
+#13294 := [rewrite]: #13293
+#13290 := (iff #12389 #13289)
+#13287 := (iff #12380 #13286)
+#13284 := (iff #12371 #13283)
+#13281 := (iff #12362 #13280)
+#13278 := (iff #12353 #13277)
+#13275 := (iff #12344 #13274)
+#13272 := (iff #12337 #13271)
+#13269 := (iff #12332 #13268)
+#13266 := (iff #12325 #13265)
+#13263 := (iff #12320 #13262)
+#13260 := (iff #12313 #13257)
+#13254 := (and #13180 #13249)
+#13258 := (iff #13254 #13257)
+#13259 := [rewrite]: #13258
+#13255 := (iff #12313 #13254)
+#13252 := (iff #12308 #13249)
+#13186 := (or #12647 #13066)
+#13189 := (or #12108 #13186)
+#13192 := (or #12117 #13189)
+#13195 := (or #12126 #13192)
+#13198 := (or #12135 #13195)
+#13201 := (or #12144 #13198)
+#13204 := (or #11471 #13201)
+#13207 := (or #13090 #13204)
+#13210 := (or #13090 #13207)
+#13213 := (or #12647 #13210)
+#13216 := (or #12647 #13213)
+#13219 := (or #12647 #13216)
+#13222 := (or #12647 #13219)
+#13225 := (or #13102 #13222)
+#13228 := (or #13136 #13225)
+#13231 := (or #13142 #13228)
+#13234 := (or #12647 #13231)
+#13237 := (or #13154 #13234)
+#13240 := (or #13164 #13237)
+#13243 := (or #13177 #13240)
+#13246 := (or #13183 #13243)
+#13250 := (iff #13246 #13249)
+#13251 := [rewrite]: #13250
+#13247 := (iff #12308 #13246)
+#13244 := (iff #12292 #13243)
+#13241 := (iff #12283 #13240)
+#13238 := (iff #12274 #13237)
+#13235 := (iff #12265 #13234)
+#13232 := (iff #12257 #13231)
+#13229 := (iff #12248 #13228)
+#13226 := (iff #12239 #13225)
+#13223 := (iff #12230 #13222)
+#13220 := (iff #12208 #13219)
+#13217 := (iff #12193 #13216)
+#13214 := (iff #12178 #13213)
+#13211 := (iff #12170 #13210)
+#13208 := (iff #12162 #13207)
+#13205 := (iff #12153 #13204)
+#13202 := (iff #12145 #13201)
+#13199 := (iff #12136 #13198)
+#13196 := (iff #12127 #13195)
+#13193 := (iff #12118 #13192)
+#13190 := (iff #12109 #13189)
+#13187 := (iff #12093 #13186)
+#13067 := (iff #12087 #13066)
+#13064 := (iff #12075 #13061)
+#13043 := (or #12647 #13036)
+#13046 := (or #12647 #13043)
+#13049 := (or #12647 #13046)
+#13052 := (or #12647 #13049)
+#13055 := (or #12923 #13052)
+#13058 := (or #12647 #13055)
+#13062 := (iff #13058 #13061)
+#13063 := [rewrite]: #13062
+#13059 := (iff #12075 #13058)
+#13056 := (iff #12067 #13055)
+#13053 := (iff #12058 #13052)
+#13050 := (iff #12043 #13049)
+#13047 := (iff #11463 #13046)
+#13044 := (iff #11448 #13043)
+#13037 := (iff #11442 #13036)
+#13034 := (iff #11437 #13031)
+#13007 := (or #11361 #13004)
+#13010 := (or #11370 #13007)
+#13013 := (or #11379 #13010)
+#13016 := (or #11388 #13013)
+#13019 := (or #12647 #13016)
+#13022 := (or #12647 #13019)
+#13025 := (or #12647 #13022)
+#13028 := (or #11436 #13025)
+#13032 := (iff #13028 #13031)
+#13033 := [rewrite]: #13032
+#13029 := (iff #11437 #13028)
+#13026 := (iff #11428 #13025)
+#13023 := (iff #11413 #13022)
+#13020 := (iff #11405 #13019)
+#13017 := (iff #11389 #13016)
+#13014 := (iff #11380 #13013)
+#13011 := (iff #11371 #13010)
+#13008 := (iff #11362 #13007)
+#13005 := (iff #11348 #13004)
+#13002 := (iff #11343 #13001)
+#12999 := (iff #4830 #12998)
+#12996 := (iff #4829 #12993)
+#12984 := (and #12952 #4826)
+#12987 := (and #6890 #12984)
+#12990 := (and #6706 #12987)
+#12994 := (iff #12990 #12993)
+#12995 := [rewrite]: #12994
+#12991 := (iff #4829 #12990)
+#12988 := (iff #4828 #12987)
+#12985 := (iff #4827 #12984)
+#12953 := (iff #4821 #12952)
+#12954 := [rewrite]: #12953
+#12986 := [monotonicity #12954]: #12985
+#6891 := (iff #1521 #6890)
+#6892 := [rewrite]: #6891
+#12989 := [monotonicity #6892 #12986]: #12988
+#6704 := (iff #1363 #6706)
+#6705 := [rewrite]: #6704
+#12992 := [monotonicity #6705 #12989]: #12991
+#12997 := [trans #12992 #12995]: #12996
+#13000 := [quant-intro #12997]: #12999
+#12982 := (iff #11342 #12981)
+#12979 := (iff #11327 #12978)
+#12976 := (iff #11322 #12973)
+#12967 := (or #12950 #12964)
+#12970 := (or #7910 #12967)
+#12974 := (iff #12970 #12973)
+#12975 := [rewrite]: #12974
+#12971 := (iff #11322 #12970)
+#12968 := (iff #11316 #12967)
+#12965 := (iff #4822 #12964)
+#12966 := [rewrite]: #12965
+#12960 := (iff #11315 #12950)
+#12955 := (not #12952)
+#12958 := (iff #12955 #12950)
+#12959 := [rewrite]: #12958
+#12956 := (iff #11315 #12955)
+#12957 := [monotonicity #12954]: #12956
+#12961 := [trans #12957 #12959]: #12960
+#12969 := [monotonicity #12961 #12966]: #12968
+#7911 := (iff #7903 #7910)
+#6898 := (iff #1522 #6897)
+#6899 := [monotonicity #6705 #6892]: #6898
+#7912 := [monotonicity #6899]: #7911
+#12972 := [monotonicity #7912 #12969]: #12971
+#12977 := [trans #12972 #12975]: #12976
+#12980 := [quant-intro #12977]: #12979
+#12983 := [monotonicity #12980]: #12982
+#13003 := [monotonicity #12983 #13000]: #13002
+#13006 := [monotonicity #12980 #13003]: #13005
+#13009 := [monotonicity #13006]: #13008
+#13012 := [monotonicity #13009]: #13011
+#13015 := [monotonicity #13012]: #13014
+#13018 := [monotonicity #13015]: #13017
+#12648 := (iff #11404 #12647)
+#12645 := (iff #4795 #12644)
+#12641 := (iff #4787 #12642)
+#12643 := [rewrite]: #12641
+#12638 := (iff #4794 #12639)
+#12640 := [rewrite]: #12638
+#12646 := [monotonicity #12640 #12643]: #12645
+#12649 := [monotonicity #12646]: #12648
+#13021 := [monotonicity #12649 #13018]: #13020
+#13024 := [monotonicity #12649 #13021]: #13023
+#13027 := [monotonicity #12649 #13024]: #13026
+#13030 := [monotonicity #13027]: #13029
+#13035 := [trans #13030 #13033]: #13034
+#13038 := [monotonicity #13035]: #13037
+#13045 := [monotonicity #12649 #13038]: #13044
+#13048 := [monotonicity #12649 #13045]: #13047
+#13051 := [monotonicity #12649 #13048]: #13050
+#13054 := [monotonicity #12649 #13051]: #13053
+#13041 := (iff #12066 #12923)
+#13039 := (iff #5050 #12922)
+#13040 := [rewrite]: #13039
+#13042 := [monotonicity #13040]: #13041
+#13057 := [monotonicity #13042 #13054]: #13056
+#13060 := [monotonicity #12649 #13057]: #13059
+#13065 := [trans #13060 #13063]: #13064
+#12948 := (iff #12028 #12945)
+#12933 := (or #12647 #12916)
+#12936 := (or #12647 #12933)
+#12939 := (or #12922 #12936)
+#12942 := (or #12647 #12939)
+#12946 := (iff #12942 #12945)
+#12947 := [rewrite]: #12946
+#12943 := (iff #12028 #12942)
+#12940 := (iff #12020 #12939)
+#12937 := (iff #12011 #12936)
+#12934 := (iff #11996 #12933)
+#12919 := (iff #11990 #12916)
+#12913 := (and #4943 #12910)
+#12917 := (iff #12913 #12916)
+#12918 := [rewrite]: #12917
+#12914 := (iff #11990 #12913)
+#12911 := (iff #11985 #12910)
+#12908 := (iff #11979 #12905)
+#12902 := (and #4946 #12897)
+#12906 := (iff #12902 #12905)
+#12907 := [rewrite]: #12906
+#12903 := (iff #11979 #12902)
+#12900 := (iff #11974 #12897)
+#12891 := (or #12647 #12888)
+#12894 := (or #11812 #12891)
+#12898 := (iff #12894 #12897)
+#12899 := [rewrite]: #12898
+#12895 := (iff #11974 #12894)
+#12892 := (iff #11966 #12891)
+#12889 := (iff #11960 #12888)
+#12886 := (iff #11948 #12883)
+#12780 := (or #12656 #12772)
+#12862 := (or #11883 #12780)
+#12865 := (or #11892 #12862)
+#12868 := (or #12647 #12865)
+#12871 := (or #12647 #12868)
+#12874 := (or #12647 #12871)
+#12877 := (or #12828 #12874)
+#12880 := (or #12647 #12877)
+#12884 := (iff #12880 #12883)
+#12885 := [rewrite]: #12884
+#12881 := (iff #11948 #12880)
+#12878 := (iff #11940 #12877)
+#12875 := (iff #11931 #12874)
+#12872 := (iff #11916 #12871)
+#12869 := (iff #11908 #12868)
+#12866 := (iff #11893 #12865)
+#12863 := (iff #11884 #12862)
+#12781 := (iff #11736 #12780)
+#12775 := (iff #11729 #12772)
+#12769 := (and #12746 #12764)
+#12773 := (iff #12769 #12772)
+#12774 := [rewrite]: #12773
+#12770 := (iff #11729 #12769)
+#12767 := (iff #11724 #12764)
+#12752 := (or #12673 #12736)
+#12755 := (or #11705 #12752)
+#12758 := (or #12743 #12755)
+#12761 := (or #12749 #12758)
+#12765 := (iff #12761 #12764)
+#12766 := [rewrite]: #12765
+#12762 := (iff #11724 #12761)
+#12759 := (iff #11715 #12758)
+#12756 := (iff #11706 #12755)
+#12753 := (iff #11697 #12752)
+#12737 := (iff #11683 #12736)
+#12734 := (iff #11678 #12733)
+#12731 := (iff #11671 #12730)
+#12728 := (iff #11666 #12727)
+#12725 := (iff #4992 #12724)
+#12722 := (iff #4987 #12721)
+#12723 := [rewrite]: #12722
+#12726 := [monotonicity #12723]: #12725
+#12716 := (iff #11665 #12715)
+#12713 := (iff #11643 #12712)
+#12710 := (iff #11638 #12707)
+#12701 := (or #12684 #12698)
+#12704 := (or #7910 #12701)
+#12708 := (iff #12704 #12707)
+#12709 := [rewrite]: #12708
+#12705 := (iff #11638 #12704)
+#12702 := (iff #11632 #12701)
+#12699 := (iff #4983 #12698)
+#12700 := [rewrite]: #12699
+#12694 := (iff #11631 #12684)
+#12686 := (not #12684)
+#12689 := (not #12686)
+#12692 := (iff #12689 #12684)
+#12693 := [rewrite]: #12692
+#12690 := (iff #11631 #12689)
+#12687 := (iff #4982 #12686)
+#12688 := [rewrite]: #12687
+#12691 := [monotonicity #12688]: #12690
+#12695 := [trans #12691 #12693]: #12694
+#12703 := [monotonicity #12695 #12700]: #12702
+#12706 := [monotonicity #7912 #12703]: #12705
+#12711 := [trans #12706 #12709]: #12710
+#12714 := [quant-intro #12711]: #12713
+#12717 := [monotonicity #12714]: #12716
+#12729 := [monotonicity #12717 #12726]: #12728
+#12732 := [monotonicity #12714 #12729]: #12731
+#12682 := (iff #11677 #12681)
+#12679 := (iff #4981 #12676)
+#12680 := [rewrite]: #12679
+#12683 := [monotonicity #12680]: #12682
+#12735 := [monotonicity #12683 #12732]: #12734
+#12738 := [monotonicity #12680 #12735]: #12737
+#12674 := (iff #11696 #12673)
+#12671 := (iff #4980 #12670)
+#12650 := (iff #4966 #12651)
+#12652 := [rewrite]: #12650
+#12667 := (iff #4979 #12668)
+#12669 := [rewrite]: #12667
+#12672 := [monotonicity #12669 #12652]: #12671
+#12675 := [monotonicity #12672]: #12674
+#12754 := [monotonicity #12675 #12738]: #12753
+#12757 := [monotonicity #12754]: #12756
+#12744 := (iff #11714 #12743)
+#12741 := (iff #11628 #12739)
+#12742 := [rewrite]: #12741
+#12745 := [monotonicity #12742]: #12744
+#12760 := [monotonicity #12745 #12757]: #12759
+#12750 := (iff #11723 #12749)
+#12747 := (iff #11625 #12746)
+#12665 := (iff #11622 #12662)
+#12666 := [rewrite]: #12665
+#12659 := (iff #11619 #12660)
+#12661 := [rewrite]: #12659
+#12748 := [monotonicity #12661 #12666]: #12747
+#12751 := [monotonicity #12748]: #12750
+#12763 := [monotonicity #12751 #12760]: #12762
+#12768 := [trans #12763 #12766]: #12767
+#12771 := [monotonicity #12748 #12768]: #12770
+#12776 := [trans #12771 #12774]: #12775
+#12657 := (iff #11735 #12656)
+#12654 := (iff #4967 #12653)
+#12655 := [monotonicity #12640 #12652]: #12654
+#12658 := [monotonicity #12655]: #12657
+#12782 := [monotonicity #12658 #12776]: #12781
+#12864 := [monotonicity #12782]: #12863
+#12867 := [monotonicity #12864]: #12866
+#12870 := [monotonicity #12649 #12867]: #12869
+#12873 := [monotonicity #12649 #12870]: #12872
+#12876 := [monotonicity #12649 #12873]: #12875
+#12860 := (iff #11939 #12828)
+#12858 := (iff #5025 #12829)
+#12859 := [rewrite]: #12858
+#12861 := [monotonicity #12859]: #12860
+#12879 := [monotonicity #12861 #12876]: #12878
+#12882 := [monotonicity #12649 #12879]: #12881
+#12887 := [trans #12882 #12885]: #12886
+#12856 := (iff #11868 #12853)
+#12841 := (or #12647 #12823)
+#12844 := (or #12647 #12841)
+#12847 := (or #12829 #12844)
+#12850 := (or #12647 #12847)
+#12854 := (iff #12850 #12853)
+#12855 := [rewrite]: #12854
+#12851 := (iff #11868 #12850)
+#12848 := (iff #11860 #12847)
+#12845 := (iff #11851 #12844)
+#12842 := (iff #11836 #12841)
+#12826 := (iff #11830 #12823)
+#12820 := (and #4943 #12817)
+#12824 := (iff #12820 #12823)
+#12825 := [rewrite]: #12824
+#12821 := (iff #11830 #12820)
+#12818 := (iff #11825 #12817)
+#12815 := (iff #11818 #12812)
+#12809 := (and #4946 #12804)
+#12813 := (iff #12809 #12812)
+#12814 := [rewrite]: #12813
+#12810 := (iff #11818 #12809)
+#12807 := (iff #11813 #12804)
+#12783 := (or #11751 #12780)
+#12786 := (or #11760 #12783)
+#12789 := (or #12777 #12786)
+#12792 := (or #11785 #12789)
+#12795 := (or #11794 #12792)
+#12798 := (or #11803 #12795)
+#12801 := (or #11812 #12798)
+#12805 := (iff #12801 #12804)
+#12806 := [rewrite]: #12805
+#12802 := (iff #11813 #12801)
+#12799 := (iff #11804 #12798)
+#12796 := (iff #11795 #12795)
+#12793 := (iff #11786 #12792)
+#12790 := (iff #11777 #12789)
+#12787 := (iff #11761 #12786)
+#12784 := (iff #11752 #12783)
+#12785 := [monotonicity #12782]: #12784
+#12788 := [monotonicity #12785]: #12787
+#12778 := (iff #11776 #12777)
+#12779 := [monotonicity #12640]: #12778
+#12791 := [monotonicity #12779 #12788]: #12790
+#12794 := [monotonicity #12791]: #12793
+#12797 := [monotonicity #12794]: #12796
+#12800 := [monotonicity #12797]: #12799
+#12803 := [monotonicity #12800]: #12802
+#12808 := [trans #12803 #12806]: #12807
+#12811 := [monotonicity #12808]: #12810
+#12816 := [trans #12811 #12814]: #12815
+#12819 := [monotonicity #12816]: #12818
+#12822 := [monotonicity #12819]: #12821
+#12827 := [trans #12822 #12825]: #12826
+#12843 := [monotonicity #12649 #12827]: #12842
+#12846 := [monotonicity #12649 #12843]: #12845
+#12839 := (iff #11859 #12829)
+#12834 := (not #12828)
+#12837 := (iff #12834 #12829)
+#12838 := [rewrite]: #12837
+#12835 := (iff #11859 #12834)
+#12832 := (iff #4948 #12828)
+#12833 := [rewrite]: #12832
+#12836 := [monotonicity #12833]: #12835
+#12840 := [trans #12836 #12838]: #12839
+#12849 := [monotonicity #12840 #12846]: #12848
+#12852 := [monotonicity #12649 #12849]: #12851
+#12857 := [trans #12852 #12855]: #12856
+#12890 := [monotonicity #12857 #12887]: #12889
+#12893 := [monotonicity #12649 #12890]: #12892
+#12896 := [monotonicity #12893]: #12895
+#12901 := [trans #12896 #12899]: #12900
+#12904 := [monotonicity #12901]: #12903
+#12909 := [trans #12904 #12907]: #12908
+#12912 := [monotonicity #12909]: #12911
+#12915 := [monotonicity #12912]: #12914
+#12920 := [trans #12915 #12918]: #12919
+#12935 := [monotonicity #12649 #12920]: #12934
+#12938 := [monotonicity #12649 #12935]: #12937
+#12931 := (iff #12019 #12922)
+#12926 := (not #12923)
+#12929 := (iff #12926 #12922)
+#12930 := [rewrite]: #12929
+#12927 := (iff #12019 #12926)
+#12924 := (iff #4935 #12923)
+#12925 := [rewrite]: #12924
+#12928 := [monotonicity #12925]: #12927
+#12932 := [trans #12928 #12930]: #12931
+#12941 := [monotonicity #12932 #12938]: #12940
+#12944 := [monotonicity #12649 #12941]: #12943
+#12949 := [trans #12944 #12947]: #12948
+#13068 := [monotonicity #12949 #13065]: #13067
+#13188 := [monotonicity #12649 #13068]: #13187
+#13191 := [monotonicity #13188]: #13190
+#13194 := [monotonicity #13191]: #13193
+#13197 := [monotonicity #13194]: #13196
+#13200 := [monotonicity #13197]: #13199
+#13203 := [monotonicity #13200]: #13202
+#13206 := [monotonicity #13203]: #13205
+#13091 := (iff #12161 #13090)
+#13088 := (iff #4900 #4898)
+#13080 := (and true #4898)
+#13083 := (and true #13080)
+#13086 := (iff #13083 #4898)
+#13087 := [rewrite]: #13086
+#13084 := (iff #4900 #13083)
+#13081 := (iff #4899 #13080)
+#13076 := (iff #4895 true)
+#13071 := (forall (vars (?v0 S11)) (:pat #4893) true)
+#13074 := (iff #13071 true)
+#13075 := [elim-unused]: #13074
+#13072 := (iff #4895 #13071)
+#13069 := (iff #4894 true)
+#13070 := [rewrite]: #13069
+#13073 := [quant-intro #13070]: #13072
+#13077 := [trans #13073 #13075]: #13076
+#13082 := [monotonicity #13077]: #13081
+#13078 := (iff #4890 true)
+#13079 := [rewrite]: #13078
+#13085 := [monotonicity #13079 #13082]: #13084
+#13089 := [trans #13085 #13087]: #13088
+#13092 := [monotonicity #13089]: #13091
+#13209 := [monotonicity #13092 #13206]: #13208
+#13212 := [monotonicity #13092 #13209]: #13211
+#13215 := [monotonicity #12649 #13212]: #13214
+#13218 := [monotonicity #12649 #13215]: #13217
+#13221 := [monotonicity #12649 #13218]: #13220
+#13224 := [monotonicity #12649 #13221]: #13223
+#13103 := (iff #12238 #13102)
+#13100 := (iff #4807 #13099)
+#13097 := (iff #4802 #13096)
+#13098 := [rewrite]: #13097
+#13101 := [monotonicity #13098]: #13100
+#13104 := [monotonicity #13101]: #13103
+#13227 := [monotonicity #13104 #13224]: #13226
+#13137 := (iff #12247 #13136)
+#13134 := (iff #11310 #13133)
+#13131 := (iff #11305 #13128)
+#13122 := (or #13105 #13119)
+#13125 := (or #7910 #13122)
+#13129 := (iff #13125 #13128)
+#13130 := [rewrite]: #13129
+#13126 := (iff #11305 #13125)
+#13123 := (iff #11299 #13122)
+#13120 := (iff #4798 #13119)
+#13121 := [rewrite]: #13120
+#13115 := (iff #11298 #13105)
+#13107 := (not #13105)
+#13110 := (not #13107)
+#13113 := (iff #13110 #13105)
+#13114 := [rewrite]: #13113
+#13111 := (iff #11298 #13110)
+#13108 := (iff #4797 #13107)
+#13109 := [rewrite]: #13108
+#13112 := [monotonicity #13109]: #13111
+#13116 := [trans #13112 #13114]: #13115
+#13124 := [monotonicity #13116 #13121]: #13123
+#13127 := [monotonicity #7912 #13124]: #13126
+#13132 := [trans #13127 #13130]: #13131
+#13135 := [quant-intro #13132]: #13134
+#13138 := [monotonicity #13135]: #13137
+#13230 := [monotonicity #13138 #13227]: #13229
+#13143 := (iff #12256 #13142)
+#13140 := (iff #4796 #13139)
+#13141 := [rewrite]: #13140
+#13144 := [monotonicity #13141]: #13143
+#13233 := [monotonicity #13144 #13230]: #13232
+#13236 := [monotonicity #12649 #13233]: #13235
+#13155 := (iff #12273 #13154)
+#13152 := (iff #4793 #13151)
+#13149 := (iff #4792 #13148)
+#13150 := [rewrite]: #13149
+#13146 := (iff #4791 #13145)
+#13147 := [rewrite]: #13146
+#13153 := [monotonicity #13147 #13150]: #13152
+#13156 := [monotonicity #13153]: #13155
+#13239 := [monotonicity #13156 #13236]: #13238
+#13165 := (iff #12282 #13164)
+#13162 := (iff #4789 #13161)
+#13159 := (iff #4788 #13157)
+#13160 := [rewrite]: #13159
+#13163 := [monotonicity #12643 #13160]: #13162
+#13166 := [monotonicity #13163]: #13165
+#13242 := [monotonicity #13166 #13239]: #13241
+#13178 := (iff #12291 #13177)
+#13175 := (iff #4785 #13174)
+#13172 := (iff #4784 #13170)
+#13173 := [rewrite]: #13172
+#13168 := (iff #4783 #13167)
+#13169 := [rewrite]: #13168
+#13176 := [monotonicity #13169 #13173]: #13175
+#13179 := [monotonicity #13176]: #13178
+#13245 := [monotonicity #13179 #13242]: #13244
+#13184 := (iff #12307 #13183)
+#13181 := (iff #4781 #13180)
+#12636 := (iff #4648 #12635)
+#12637 := [rewrite]: #12636
+#13182 := [monotonicity #12637]: #13181
+#13185 := [monotonicity #13182]: #13184
+#13248 := [monotonicity #13185 #13245]: #13247
+#13253 := [trans #13248 #13251]: #13252
+#13256 := [monotonicity #13182 #13253]: #13255
+#13261 := [trans #13256 #13259]: #13260
+#12632 := (iff #12319 #12631)
+#12629 := (iff #11295 #12628)
+#12626 := (iff #11290 #12623)
+#12617 := (or #12601 #12613)
+#12620 := (or #7910 #12617)
+#12624 := (iff #12620 #12623)
+#12625 := [rewrite]: #12624
+#12621 := (iff #11290 #12620)
+#12618 := (iff #11284 #12617)
+#12612 := (iff #4776 #12613)
+#12616 := [rewrite]: #12612
+#12610 := (iff #11283 #12601)
+#12602 := (not #12601)
+#12605 := (not #12602)
+#12608 := (iff #12605 #12601)
+#12609 := [rewrite]: #12608
+#12606 := (iff #11283 #12605)
+#12603 := (iff #4772 #12602)
+#12604 := [rewrite]: #12603
+#12607 := [monotonicity #12604]: #12606
+#12611 := [trans #12607 #12609]: #12610
+#12619 := [monotonicity #12611 #12616]: #12618
+#12622 := [monotonicity #7912 #12619]: #12621
+#12627 := [trans #12622 #12625]: #12626
+#12630 := [quant-intro #12627]: #12629
+#12633 := [monotonicity #12630]: #12632
+#13264 := [monotonicity #12633 #13261]: #13263
+#13267 := [monotonicity #12630 #13264]: #13266
+#12599 := (iff #12331 #12598)
+#12596 := (iff #4771 #12595)
+#12597 := [rewrite]: #12596
+#12600 := [monotonicity #12597]: #12599
+#13270 := [monotonicity #12600 #13267]: #13269
+#13273 := [monotonicity #12597 #13270]: #13272
+#12593 := (iff #12343 false)
+#11313 := (iff #4808 false)
+#11314 := [rewrite]: #11313
+#12591 := (iff #12343 #4808)
+#12589 := (iff #11280 true)
+#11607 := (and true true)
+#12584 := (and true #11607)
+#12587 := (iff #12584 true)
+#12588 := [rewrite]: #12587
+#12585 := (iff #11280 #12584)
+#12582 := (iff #11277 #11607)
+#12580 := (iff #4767 true)
+#12581 := [rewrite]: #12580
+#12578 := (iff #4766 true)
+#12579 := [rewrite]: #12578
+#12583 := [monotonicity #12579 #12581]: #12582
+#12586 := [monotonicity #12579 #12583]: #12585
+#12590 := [trans #12586 #12588]: #12589
+#12592 := [monotonicity #12590]: #12591
+#12594 := [trans #12592 #11314]: #12593
+#13276 := [monotonicity #12594 #13273]: #13275
+#13279 := [monotonicity #13276]: #13278
+#13282 := [monotonicity #13279]: #13281
+#13285 := [monotonicity #13282]: #13284
+#13288 := [monotonicity #13285]: #13287
+#13291 := [monotonicity #13288]: #13290
+#13296 := [trans #13291 #13294]: #13295
+#13299 := [monotonicity #13296]: #13298
+#13304 := [trans #13299 #13302]: #13303
+#13307 := [monotonicity #13304]: #13306
+#13310 := [monotonicity #13307]: #13309
+#13315 := [trans #13310 #13313]: #13314
+#13318 := [monotonicity #13315]: #13317
+#13321 := [monotonicity #13318]: #13320
+#12576 := (iff #12424 #12575)
+#12573 := (iff #4727 #12572)
+#12570 := (iff #4726 #12567)
+#12571 := [rewrite]: #12570
+#12564 := (iff #4725 #12565)
+#12566 := [rewrite]: #12564
+#12574 := [monotonicity #12566 #12571]: #12573
+#12577 := [monotonicity #12574]: #12576
+#13402 := [monotonicity #12577 #13321]: #13401
+#13405 := [monotonicity #13402]: #13404
+#13408 := [monotonicity #13405]: #13407
+#13411 := [monotonicity #13408]: #13410
+#13414 := [monotonicity #13411]: #13413
+#13332 := (iff #12469 #13331)
+#13329 := (iff #4694 #13328)
+#13326 := (iff #4693 #13322)
+#13327 := [rewrite]: #13326
+#13330 := [quant-intro #13327]: #13329
+#13333 := [monotonicity #13330]: #13332
+#13417 := [monotonicity #13333 #13414]: #13416
+#13420 := [monotonicity #13417]: #13419
+#13423 := [monotonicity #13420]: #13422
+#13338 := (iff #12503 #13337)
+#13335 := (iff #4679 #13334)
+#13336 := [rewrite]: #13335
+#13339 := [monotonicity #13336]: #13338
+#13426 := [monotonicity #13339 #13423]: #13425
+#13345 := (iff #12512 #12634)
+#13340 := (not #12635)
+#13343 := (iff #13340 #12634)
+#13344 := [rewrite]: #13343
+#13341 := (iff #12512 #13340)
+#13342 := [monotonicity #12637]: #13341
+#13346 := [trans #13342 #13344]: #13345
+#13429 := [monotonicity #13346 #13426]: #13428
+#13356 := (iff #12521 #13347)
+#13348 := (not #13347)
+#13351 := (not #13348)
+#13354 := (iff #13351 #13347)
+#13355 := [rewrite]: #13354
+#13352 := (iff #12521 #13351)
+#13349 := (iff #4647 #13348)
+#13350 := [rewrite]: #13349
+#13353 := [monotonicity #13350]: #13352
+#13357 := [trans #13353 #13355]: #13356
+#13432 := [monotonicity #13357 #13429]: #13431
+#13370 := (iff #12530 #13369)
+#13367 := (iff #4645 #13366)
+#13364 := (iff #4644 #13361)
+#13365 := [rewrite]: #13364
+#13358 := (iff #4643 #13359)
+#13360 := [rewrite]: #13358
+#13368 := [monotonicity #13360 #13365]: #13367
+#13371 := [monotonicity #13368]: #13370
+#13435 := [monotonicity #13371 #13432]: #13434
+#13384 := (iff #12539 #13383)
+#13381 := (iff #4641 #13380)
+#13378 := (iff #4640 #13375)
+#13379 := [rewrite]: #13378
+#13372 := (iff #4639 #13373)
+#13374 := [rewrite]: #13372
+#13382 := [monotonicity #13374 #13379]: #13381
+#13385 := [monotonicity #13382]: #13384
+#13438 := [monotonicity #13385 #13435]: #13437
+#13398 := (iff #12548 #13397)
+#13395 := (iff #4637 #13394)
+#13392 := (iff #4636 #13389)
+#13393 := [rewrite]: #13392
+#13386 := (iff #4635 #13387)
+#13388 := [rewrite]: #13386
+#13396 := [monotonicity #13388 #13393]: #13395
+#13399 := [monotonicity #13396]: #13398
+#13441 := [monotonicity #13399 #13438]: #13440
+#13446 := [trans #13441 #13444]: #13445
+#13449 := [monotonicity #13446]: #13448
+#12562 := (iff #5117 #12561)
+#12559 := (iff #5116 #12549)
+#12554 := (implies true #12549)
+#12557 := (iff #12554 #12549)
+#12558 := [rewrite]: #12557
+#12555 := (iff #5116 #12554)
+#12552 := (iff #5115 #12549)
+#12545 := (implies #4637 #12540)
+#12550 := (iff #12545 #12549)
+#12551 := [rewrite]: #12550
+#12546 := (iff #5115 #12545)
+#12543 := (iff #5114 #12540)
+#12536 := (implies #4641 #12531)
+#12541 := (iff #12536 #12540)
+#12542 := [rewrite]: #12541
+#12537 := (iff #5114 #12536)
+#12534 := (iff #5113 #12531)
+#12527 := (implies #4645 #12522)
+#12532 := (iff #12527 #12531)
+#12533 := [rewrite]: #12532
+#12528 := (iff #5113 #12527)
+#12525 := (iff #5112 #12522)
+#12518 := (implies #4647 #12513)
+#12523 := (iff #12518 #12522)
+#12524 := [rewrite]: #12523
+#12519 := (iff #5112 #12518)
+#12516 := (iff #5111 #12513)
+#12509 := (implies #4648 #12504)
+#12514 := (iff #12509 #12513)
+#12515 := [rewrite]: #12514
+#12510 := (iff #5111 #12509)
+#12507 := (iff #5110 #12504)
+#12500 := (implies #4679 #12488)
+#12505 := (iff #12500 #12504)
+#12506 := [rewrite]: #12505
+#12501 := (iff #5110 #12500)
+#12498 := (iff #5109 #12488)
+#12493 := (implies true #12488)
+#12496 := (iff #12493 #12488)
+#12497 := [rewrite]: #12496
+#12494 := (iff #5109 #12493)
+#12491 := (iff #5108 #12488)
+#12484 := (implies #4681 #12479)
+#12489 := (iff #12484 #12488)
+#12490 := [rewrite]: #12489
+#12485 := (iff #5108 #12484)
+#12482 := (iff #5107 #12479)
+#12475 := (implies #4687 #12470)
+#12480 := (iff #12475 #12479)
+#12481 := [rewrite]: #12480
+#12476 := (iff #5107 #12475)
+#12473 := (iff #5106 #12470)
+#12466 := (implies #4694 #12461)
+#12471 := (iff #12466 #12470)
+#12472 := [rewrite]: #12471
+#12467 := (iff #5106 #12466)
+#12464 := (iff #5105 #12461)
+#12457 := (implies #4710 #12452)
+#12462 := (iff #12457 #12461)
+#12463 := [rewrite]: #12462
+#12458 := (iff #5105 #12457)
+#12455 := (iff #5104 #12452)
+#12448 := (implies #4714 #12443)
+#12453 := (iff #12448 #12452)
+#12454 := [rewrite]: #12453
+#12449 := (iff #5104 #12448)
+#12446 := (iff #5103 #12443)
+#12439 := (implies #4717 #12434)
+#12444 := (iff #12439 #12443)
+#12445 := [rewrite]: #12444
+#12440 := (iff #5103 #12439)
+#12437 := (iff #5102 #12434)
+#12430 := (implies #11272 #12425)
+#12435 := (iff #12430 #12434)
+#12436 := [rewrite]: #12435
+#12431 := (iff #5102 #12430)
+#12428 := (iff #5101 #12425)
+#12421 := (implies #4727 #12418)
+#12426 := (iff #12421 #12425)
+#12427 := [rewrite]: #12426
+#12422 := (iff #5101 #12421)
+#12419 := (iff #5100 #12418)
+#12416 := (iff #5099 #12413)
+#12409 := (implies #4733 #12406)
+#12414 := (iff #12409 #12413)
+#12415 := [rewrite]: #12414
+#12410 := (iff #5099 #12409)
+#12407 := (iff #5098 #12406)
+#12404 := (iff #5097 #12401)
+#12397 := (implies #4742 #12394)
+#12402 := (iff #12397 #12401)
+#12403 := [rewrite]: #12402
+#12398 := (iff #5097 #12397)
+#12395 := (iff #5096 #12394)
+#12392 := (iff #5095 #12389)
+#12385 := (implies #4746 #12380)
+#12390 := (iff #12385 #12389)
+#12391 := [rewrite]: #12390
+#12386 := (iff #5095 #12385)
+#12383 := (iff #5094 #12380)
+#12376 := (implies #4750 #12371)
+#12381 := (iff #12376 #12380)
+#12382 := [rewrite]: #12381
+#12377 := (iff #5094 #12376)
+#12374 := (iff #5093 #12371)
+#12367 := (implies #4755 #12362)
+#12372 := (iff #12367 #12371)
+#12373 := [rewrite]: #12372
+#12368 := (iff #5093 #12367)
+#12365 := (iff #5092 #12362)
+#12358 := (implies #4760 #12353)
+#12363 := (iff #12358 #12362)
+#12364 := [rewrite]: #12363
+#12359 := (iff #5092 #12358)
+#12356 := (iff #5091 #12353)
+#12349 := (implies #4765 #12344)
+#12354 := (iff #12349 #12353)
+#12355 := [rewrite]: #12354
+#12350 := (iff #5091 #12349)
+#12347 := (iff #5090 #12344)
+#12340 := (implies #11280 #12337)
+#12345 := (iff #12340 #12344)
+#12346 := [rewrite]: #12345
+#12341 := (iff #5090 #12340)
+#12338 := (iff #5089 #12337)
+#12335 := (iff #5088 #12332)
+#12328 := (implies #4771 #12325)
+#12333 := (iff #12328 #12332)
+#12334 := [rewrite]: #12333
+#12329 := (iff #5088 #12328)
+#12326 := (iff #5087 #12325)
+#12323 := (iff #5086 #12320)
+#12316 := (implies #11295 #12313)
+#12321 := (iff #12316 #12320)
+#12322 := [rewrite]: #12321
+#12317 := (iff #5086 #12316)
+#12314 := (iff #5085 #12313)
+#12311 := (iff #5084 #12308)
+#12304 := (implies #4781 #12292)
+#12309 := (iff #12304 #12308)
+#12310 := [rewrite]: #12309
+#12305 := (iff #5084 #12304)
+#12302 := (iff #5083 #12292)
+#12297 := (implies true #12292)
+#12300 := (iff #12297 #12292)
+#12301 := [rewrite]: #12300
+#12298 := (iff #5083 #12297)
+#12295 := (iff #5082 #12292)
+#12288 := (implies #4785 #12283)
+#12293 := (iff #12288 #12292)
+#12294 := [rewrite]: #12293
+#12289 := (iff #5082 #12288)
+#12286 := (iff #5081 #12283)
+#12279 := (implies #4789 #12274)
+#12284 := (iff #12279 #12283)
+#12285 := [rewrite]: #12284
+#12280 := (iff #5081 #12279)
+#12277 := (iff #5080 #12274)
+#12270 := (implies #4793 #12265)
+#12275 := (iff #12270 #12274)
+#12276 := [rewrite]: #12275
+#12271 := (iff #5080 #12270)
+#12268 := (iff #5079 #12265)
+#12262 := (implies #4795 #12257)
+#12266 := (iff #12262 #12265)
+#12267 := [rewrite]: #12266
+#12263 := (iff #5079 #12262)
+#12260 := (iff #5078 #12257)
+#12253 := (implies #4796 #12248)
+#12258 := (iff #12253 #12257)
+#12259 := [rewrite]: #12258
+#12254 := (iff #5078 #12253)
+#12251 := (iff #5077 #12248)
+#12244 := (implies #11310 #12239)
+#12249 := (iff #12244 #12248)
+#12250 := [rewrite]: #12249
+#12245 := (iff #5077 #12244)
+#12242 := (iff #5076 #12239)
+#12235 := (implies #4807 #12230)
+#12240 := (iff #12235 #12239)
+#12241 := [rewrite]: #12240
+#12236 := (iff #5076 #12235)
+#12233 := (iff #5075 #12230)
+#12227 := (implies #4795 #12208)
+#12231 := (iff #12227 #12230)
+#12232 := [rewrite]: #12231
+#12228 := (iff #5075 #12227)
+#12225 := (iff #5074 #12208)
+#12220 := (and true #12208)
+#12223 := (iff #12220 #12208)
+#12224 := [rewrite]: #12223
+#12221 := (iff #5074 #12220)
+#12218 := (iff #5073 #12208)
+#12213 := (implies true #12208)
+#12216 := (iff #12213 #12208)
+#12217 := [rewrite]: #12216
+#12214 := (iff #5073 #12213)
+#12211 := (iff #5072 #12208)
+#12205 := (implies #4795 #12193)
+#12209 := (iff #12205 #12208)
+#12210 := [rewrite]: #12209
+#12206 := (iff #5072 #12205)
+#12203 := (iff #5071 #12193)
+#12198 := (implies true #12193)
+#12201 := (iff #12198 #12193)
+#12202 := [rewrite]: #12201
+#12199 := (iff #5071 #12198)
+#12196 := (iff #5070 #12193)
+#12190 := (implies #4795 #12178)
+#12194 := (iff #12190 #12193)
+#12195 := [rewrite]: #12194
+#12191 := (iff #5070 #12190)
+#12188 := (iff #5069 #12178)
+#12183 := (implies true #12178)
+#12186 := (iff #12183 #12178)
+#12187 := [rewrite]: #12186
+#12184 := (iff #5069 #12183)
+#12181 := (iff #5068 #12178)
+#12175 := (implies #4795 #12170)
+#12179 := (iff #12175 #12178)
+#12180 := [rewrite]: #12179
+#12176 := (iff #5068 #12175)
+#12173 := (iff #5067 #12170)
+#12167 := (implies #4900 #12162)
+#12171 := (iff #12167 #12170)
+#12172 := [rewrite]: #12171
+#12168 := (iff #5067 #12167)
+#12165 := (iff #5066 #12162)
+#12158 := (implies #4900 #12153)
+#12163 := (iff #12158 #12162)
+#12164 := [rewrite]: #12163
+#12159 := (iff #5066 #12158)
+#12156 := (iff #5065 #12153)
+#12150 := (implies #4812 #12145)
+#12154 := (iff #12150 #12153)
+#12155 := [rewrite]: #12154
+#12151 := (iff #5065 #12150)
+#12148 := (iff #5064 #12145)
+#12141 := (implies #4909 #12136)
+#12146 := (iff #12141 #12145)
+#12147 := [rewrite]: #12146
+#12142 := (iff #5064 #12141)
+#12139 := (iff #5063 #12136)
+#12132 := (implies #4913 #12127)
+#12137 := (iff #12132 #12136)
+#12138 := [rewrite]: #12137
+#12133 := (iff #5063 #12132)
+#12130 := (iff #5062 #12127)
+#12123 := (implies #4917 #12118)
+#12128 := (iff #12123 #12127)
+#12129 := [rewrite]: #12128
+#12124 := (iff #5062 #12123)
+#12121 := (iff #5061 #12118)
+#12114 := (implies #4921 #12109)
+#12119 := (iff #12114 #12118)
+#12120 := [rewrite]: #12119
+#12115 := (iff #5061 #12114)
+#12112 := (iff #5060 #12109)
+#12105 := (implies #4931 #12093)
+#12110 := (iff #12105 #12109)
+#12111 := [rewrite]: #12110
+#12106 := (iff #5060 #12105)
+#12103 := (iff #5059 #12093)
+#12098 := (implies true #12093)
+#12101 := (iff #12098 #12093)
+#12102 := [rewrite]: #12101
+#12099 := (iff #5059 #12098)
+#12096 := (iff #5058 #12093)
+#12090 := (implies #4795 #12087)
+#12094 := (iff #12090 #12093)
+#12095 := [rewrite]: #12094
+#12091 := (iff #5058 #12090)
+#12088 := (iff #5057 #12087)
+#12085 := (iff #5056 #12075)
+#12080 := (implies true #12075)
+#12083 := (iff #12080 #12075)
+#12084 := [rewrite]: #12083
+#12081 := (iff #5056 #12080)
+#12078 := (iff #5055 #12075)
+#12072 := (implies #4795 #12067)
+#12076 := (iff #12072 #12075)
+#12077 := [rewrite]: #12076
+#12073 := (iff #5055 #12072)
+#12070 := (iff #5054 #12067)
+#12063 := (implies #5050 #12058)
+#12068 := (iff #12063 #12067)
+#12069 := [rewrite]: #12068
+#12064 := (iff #5054 #12063)
+#12061 := (iff #5053 #12058)
+#12055 := (implies #4795 #12043)
+#12059 := (iff #12055 #12058)
+#12060 := [rewrite]: #12059
+#12056 := (iff #5053 #12055)
+#12053 := (iff #5052 #12043)
+#12048 := (implies true #12043)
+#12051 := (iff #12048 #12043)
+#12052 := [rewrite]: #12051
+#12049 := (iff #5052 #12048)
+#12046 := (iff #5051 #12043)
+#12040 := (implies #4795 #11463)
+#12044 := (iff #12040 #12043)
+#12045 := [rewrite]: #12044
+#12041 := (iff #5051 #12040)
+#11466 := (iff #4849 #11463)
+#11460 := (implies #4795 #11448)
+#11464 := (iff #11460 #11463)
+#11465 := [rewrite]: #11464
+#11461 := (iff #4849 #11460)
+#11458 := (iff #4848 #11448)
+#11453 := (implies true #11448)
+#11456 := (iff #11453 #11448)
+#11457 := [rewrite]: #11456
+#11454 := (iff #4848 #11453)
+#11451 := (iff #4847 #11448)
+#11445 := (implies #4795 #11442)
+#11449 := (iff #11445 #11448)
+#11450 := [rewrite]: #11449
+#11446 := (iff #4847 #11445)
+#11443 := (iff #4846 #11442)
+#11440 := (iff #4845 #11437)
+#11433 := (implies #4072 #11428)
+#11438 := (iff #11433 #11437)
+#11439 := [rewrite]: #11438
+#11434 := (iff #4845 #11433)
+#11431 := (iff #4844 #11428)
+#11425 := (implies #4795 #11413)
+#11429 := (iff #11425 #11428)
+#11430 := [rewrite]: #11429
+#11426 := (iff #4844 #11425)
+#11423 := (iff #4843 #11413)
+#11418 := (implies true #11413)
+#11421 := (iff #11418 #11413)
+#11422 := [rewrite]: #11421
+#11419 := (iff #4843 #11418)
+#11416 := (iff #4842 #11413)
+#11410 := (implies #4795 #11405)
+#11414 := (iff #11410 #11413)
+#11415 := [rewrite]: #11414
+#11411 := (iff #4842 #11410)
+#11408 := (iff #4841 #11405)
+#11401 := (implies #4795 #11389)
+#11406 := (iff #11401 #11405)
+#11407 := [rewrite]: #11406
+#11402 := (iff #4841 #11401)
+#11399 := (iff #4840 #11389)
+#11394 := (implies true #11389)
+#11397 := (iff #11394 #11389)
+#11398 := [rewrite]: #11397
+#11395 := (iff #4840 #11394)
+#11392 := (iff #4839 #11389)
+#11385 := (implies #4814 #11380)
+#11390 := (iff #11385 #11389)
+#11391 := [rewrite]: #11390
+#11386 := (iff #4839 #11385)
+#11383 := (iff #4838 #11380)
+#11376 := (implies #4816 #11371)
+#11381 := (iff #11376 #11380)
+#11382 := [rewrite]: #11381
+#11377 := (iff #4838 #11376)
+#11374 := (iff #4837 #11371)
+#11367 := (implies #4818 #11362)
+#11372 := (iff #11367 #11371)
+#11373 := [rewrite]: #11372
+#11368 := (iff #4837 #11367)
+#11365 := (iff #4836 #11362)
+#11358 := (implies #4820 #11348)
+#11363 := (iff #11358 #11362)
+#11364 := [rewrite]: #11363
+#11359 := (iff #4836 #11358)
+#11356 := (iff #4835 #11348)
+#11351 := (implies true #11348)
+#11354 := (iff #11351 #11348)
+#11355 := [rewrite]: #11354
+#11352 := (iff #4835 #11351)
+#11349 := (iff #4834 #11348)
+#11346 := (iff #4833 #11343)
+#11339 := (implies #11327 #4830)
+#11344 := (iff #11339 #11343)
+#11345 := [rewrite]: #11344
+#11340 := (iff #4833 #11339)
+#11337 := (iff #4832 #4830)
+#11332 := (and #4830 true)
+#11335 := (iff #11332 #4830)
+#11336 := [rewrite]: #11335
+#11333 := (iff #4832 #11332)
+#11330 := (iff #4831 true)
+#11331 := [rewrite]: #11330
+#11334 := [monotonicity #11331]: #11333
+#11338 := [trans #11334 #11336]: #11337
+#11328 := (iff #4825 #11327)
+#11325 := (iff #4824 #11322)
+#11319 := (implies #1522 #11316)
+#11323 := (iff #11319 #11322)
+#11324 := [rewrite]: #11323
+#11320 := (iff #4824 #11319)
+#11317 := (iff #4823 #11316)
+#11318 := [rewrite]: #11317
+#11321 := [monotonicity #11318]: #11320
+#11326 := [trans #11321 #11324]: #11325
+#11329 := [quant-intro #11326]: #11328
+#11341 := [monotonicity #11329 #11338]: #11340
+#11347 := [trans #11341 #11345]: #11346
+#11350 := [monotonicity #11329 #11347]: #11349
+#11353 := [monotonicity #11350]: #11352
+#11357 := [trans #11353 #11355]: #11356
+#11360 := [monotonicity #11357]: #11359
+#11366 := [trans #11360 #11364]: #11365
+#11369 := [monotonicity #11366]: #11368
+#11375 := [trans #11369 #11373]: #11374
+#11378 := [monotonicity #11375]: #11377
+#11384 := [trans #11378 #11382]: #11383
+#11387 := [monotonicity #11384]: #11386
+#11393 := [trans #11387 #11391]: #11392
+#11396 := [monotonicity #11393]: #11395
+#11400 := [trans #11396 #11398]: #11399
+#11403 := [monotonicity #11400]: #11402
+#11409 := [trans #11403 #11407]: #11408
+#11412 := [monotonicity #11409]: #11411
+#11417 := [trans #11412 #11415]: #11416
+#11420 := [monotonicity #11417]: #11419
+#11424 := [trans #11420 #11422]: #11423
+#11427 := [monotonicity #11424]: #11426
+#11432 := [trans #11427 #11430]: #11431
+#11435 := [monotonicity #11432]: #11434
+#11441 := [trans #11435 #11439]: #11440
+#11444 := [monotonicity #11441]: #11443
+#11447 := [monotonicity #11444]: #11446
+#11452 := [trans #11447 #11450]: #11451
+#11455 := [monotonicity #11452]: #11454
+#11459 := [trans #11455 #11457]: #11458
+#11462 := [monotonicity #11459]: #11461
+#11467 := [trans #11462 #11465]: #11466
+#12042 := [monotonicity #11467]: #12041
+#12047 := [trans #12042 #12045]: #12046
+#12050 := [monotonicity #12047]: #12049
+#12054 := [trans #12050 #12052]: #12053
+#12057 := [monotonicity #12054]: #12056
+#12062 := [trans #12057 #12060]: #12061
+#12065 := [monotonicity #12062]: #12064
+#12071 := [trans #12065 #12069]: #12070
+#12074 := [monotonicity #12071]: #12073
+#12079 := [trans #12074 #12077]: #12078
+#12082 := [monotonicity #12079]: #12081
+#12086 := [trans #12082 #12084]: #12085
+#12038 := (iff #5049 #12028)
+#12033 := (implies true #12028)
+#12036 := (iff #12033 #12028)
+#12037 := [rewrite]: #12036
+#12034 := (iff #5049 #12033)
+#12031 := (iff #5048 #12028)
+#12025 := (implies #4795 #12020)
+#12029 := (iff #12025 #12028)
+#12030 := [rewrite]: #12029
+#12026 := (iff #5048 #12025)
+#12023 := (iff #5047 #12020)
+#12016 := (implies #4935 #12011)
+#12021 := (iff #12016 #12020)
+#12022 := [rewrite]: #12021
+#12017 := (iff #5047 #12016)
+#12014 := (iff #5046 #12011)
+#12008 := (implies #4795 #11996)
+#12012 := (iff #12008 #12011)
+#12013 := [rewrite]: #12012
+#12009 := (iff #5046 #12008)
+#12006 := (iff #5045 #11996)
+#12001 := (implies true #11996)
+#12004 := (iff #12001 #11996)
+#12005 := [rewrite]: #12004
+#12002 := (iff #5045 #12001)
+#11999 := (iff #5044 #11996)
+#11993 := (implies #4795 #11990)
+#11997 := (iff #11993 #11996)
+#11998 := [rewrite]: #11997
+#11994 := (iff #5044 #11993)
+#11991 := (iff #5043 #11990)
+#11988 := (iff #5042 #11985)
+#11982 := (implies #4943 #11979)
+#11986 := (iff #11982 #11985)
+#11987 := [rewrite]: #11986
+#11983 := (iff #5042 #11982)
+#11980 := (iff #5041 #11979)
+#11977 := (iff #5040 #11974)
+#11971 := (implies #4946 #11966)
+#11975 := (iff #11971 #11974)
+#11976 := [rewrite]: #11975
+#11972 := (iff #5040 #11971)
+#11969 := (iff #5039 #11966)
+#11963 := (implies #4795 #11960)
+#11967 := (iff #11963 #11966)
+#11968 := [rewrite]: #11967
+#11964 := (iff #5039 #11963)
+#11961 := (iff #5038 #11960)
+#11958 := (iff #5037 #11948)
+#11953 := (implies true #11948)
+#11956 := (iff #11953 #11948)
+#11957 := [rewrite]: #11956
+#11954 := (iff #5037 #11953)
+#11951 := (iff #5036 #11948)
+#11945 := (implies #4795 #11940)
+#11949 := (iff #11945 #11948)
+#11950 := [rewrite]: #11949
+#11946 := (iff #5036 #11945)
+#11943 := (iff #5035 #11940)
+#11936 := (implies #5025 #11931)
+#11941 := (iff #11936 #11940)
+#11942 := [rewrite]: #11941
+#11937 := (iff #5035 #11936)
+#11934 := (iff #5034 #11931)
+#11928 := (implies #4795 #11916)
+#11932 := (iff #11928 #11931)
+#11933 := [rewrite]: #11932
+#11929 := (iff #5034 #11928)
+#11926 := (iff #5033 #11916)
+#11921 := (implies true #11916)
+#11924 := (iff #11921 #11916)
+#11925 := [rewrite]: #11924
+#11922 := (iff #5033 #11921)
+#11919 := (iff #5032 #11916)
+#11913 := (implies #4795 #11908)
+#11917 := (iff #11913 #11916)
+#11918 := [rewrite]: #11917
+#11914 := (iff #5032 #11913)
+#11911 := (iff #5031 #11908)
+#11905 := (implies #4795 #11893)
+#11909 := (iff #11905 #11908)
+#11910 := [rewrite]: #11909
+#11906 := (iff #5031 #11905)
+#11903 := (iff #5030 #11893)
+#11898 := (implies true #11893)
+#11901 := (iff #11898 #11893)
+#11902 := [rewrite]: #11901
+#11899 := (iff #5030 #11898)
+#11896 := (iff #5029 #11893)
+#11889 := (implies #5026 #11884)
+#11894 := (iff #11889 #11893)
+#11895 := [rewrite]: #11894
+#11890 := (iff #5029 #11889)
+#11887 := (iff #5028 #11884)
+#11880 := (implies #5027 #11736)
+#11885 := (iff #11880 #11884)
+#11886 := [rewrite]: #11885
+#11881 := (iff #5028 #11880)
+#11746 := (iff #5007 #11736)
+#11741 := (implies true #11736)
+#11744 := (iff #11741 #11736)
+#11745 := [rewrite]: #11744
+#11742 := (iff #5007 #11741)
+#11739 := (iff #5006 #11736)
+#11732 := (implies #4967 #11729)
+#11737 := (iff #11732 #11736)
+#11738 := [rewrite]: #11737
+#11733 := (iff #5006 #11732)
+#11730 := (iff #5005 #11729)
+#11727 := (iff #5004 #11724)
+#11720 := (implies #11625 #11715)
+#11725 := (iff #11720 #11724)
+#11726 := [rewrite]: #11725
+#11721 := (iff #5004 #11720)
+#11718 := (iff #5003 #11715)
+#11711 := (implies #11628 #11706)
+#11716 := (iff #11711 #11715)
+#11717 := [rewrite]: #11716
+#11712 := (iff #5003 #11711)
+#11709 := (iff #5002 #11706)
+#11702 := (implies #4978 #11697)
+#11707 := (iff #11702 #11706)
+#11708 := [rewrite]: #11707
+#11703 := (iff #5002 #11702)
+#11700 := (iff #5001 #11697)
+#11693 := (implies #4980 #11683)
+#11698 := (iff #11693 #11697)
+#11699 := [rewrite]: #11698
+#11694 := (iff #5001 #11693)
+#11691 := (iff #5000 #11683)
+#11686 := (implies true #11683)
+#11689 := (iff #11686 #11683)
+#11690 := [rewrite]: #11689
+#11687 := (iff #5000 #11686)
+#11684 := (iff #4999 #11683)
+#11681 := (iff #4998 #11678)
+#11674 := (implies #4981 #11671)
+#11679 := (iff #11674 #11678)
+#11680 := [rewrite]: #11679
+#11675 := (iff #4998 #11674)
+#11672 := (iff #4997 #11671)
+#11669 := (iff #4996 #11666)
+#11662 := (implies #11643 #4992)
+#11667 := (iff #11662 #11666)
+#11668 := [rewrite]: #11667
+#11663 := (iff #4996 #11662)
+#11660 := (iff #4995 #4992)
+#11655 := (and #4992 true)
+#11658 := (iff #11655 #4992)
+#11659 := [rewrite]: #11658
+#11656 := (iff #4995 #11655)
+#11653 := (iff #4994 true)
+#11648 := (implies #4992 true)
+#11651 := (iff #11648 true)
+#11652 := [rewrite]: #11651
+#11649 := (iff #4994 #11648)
+#11646 := (iff #4993 true)
+#11647 := [rewrite]: #11646
+#11650 := [monotonicity #11647]: #11649
+#11654 := [trans #11650 #11652]: #11653
+#11657 := [monotonicity #11654]: #11656
+#11661 := [trans #11657 #11659]: #11660
+#11644 := (iff #4986 #11643)
+#11641 := (iff #4985 #11638)
+#11635 := (implies #1522 #11632)
+#11639 := (iff #11635 #11638)
+#11640 := [rewrite]: #11639
+#11636 := (iff #4985 #11635)
+#11633 := (iff #4984 #11632)
+#11634 := [rewrite]: #11633
+#11637 := [monotonicity #11634]: #11636
+#11642 := [trans #11637 #11640]: #11641
+#11645 := [quant-intro #11642]: #11644
+#11664 := [monotonicity #11645 #11661]: #11663
+#11670 := [trans #11664 #11668]: #11669
+#11673 := [monotonicity #11645 #11670]: #11672
+#11676 := [monotonicity #11673]: #11675
+#11682 := [trans #11676 #11680]: #11681
+#11685 := [monotonicity #11682]: #11684
+#11688 := [monotonicity #11685]: #11687
+#11692 := [trans #11688 #11690]: #11691
+#11695 := [monotonicity #11692]: #11694
+#11701 := [trans #11695 #11699]: #11700
+#11704 := [monotonicity #11701]: #11703
+#11710 := [trans #11704 #11708]: #11709
+#11629 := (iff #4973 #11628)
+#11617 := (= #4968 #11616)
+#11618 := [rewrite]: #11617
+#11630 := [monotonicity #11618]: #11629
+#11713 := [monotonicity #11630 #11710]: #11712
+#11719 := [trans #11713 #11717]: #11718
+#11626 := (iff #4971 #11625)
+#11623 := (iff #4970 #11622)
+#11624 := [monotonicity #11618]: #11623
+#11620 := (iff #4969 #11619)
+#11621 := [monotonicity #11618]: #11620
+#11627 := [monotonicity #11621 #11624]: #11626
+#11722 := [monotonicity #11627 #11719]: #11721
+#11728 := [trans #11722 #11726]: #11727
+#11731 := [monotonicity #11627 #11728]: #11730
+#11734 := [monotonicity #11731]: #11733
+#11740 := [trans #11734 #11738]: #11739
+#11743 := [monotonicity #11740]: #11742
+#11747 := [trans #11743 #11745]: #11746
+#11882 := [monotonicity #11747]: #11881
+#11888 := [trans #11882 #11886]: #11887
+#11891 := [monotonicity #11888]: #11890
+#11897 := [trans #11891 #11895]: #11896
+#11900 := [monotonicity #11897]: #11899
+#11904 := [trans #11900 #11902]: #11903
+#11907 := [monotonicity #11904]: #11906
+#11912 := [trans #11907 #11910]: #11911
+#11915 := [monotonicity #11912]: #11914
+#11920 := [trans #11915 #11918]: #11919
+#11923 := [monotonicity #11920]: #11922
+#11927 := [trans #11923 #11925]: #11926
+#11930 := [monotonicity #11927]: #11929
+#11935 := [trans #11930 #11933]: #11934
+#11938 := [monotonicity #11935]: #11937
+#11944 := [trans #11938 #11942]: #11943
+#11947 := [monotonicity #11944]: #11946
+#11952 := [trans #11947 #11950]: #11951
+#11955 := [monotonicity #11952]: #11954
+#11959 := [trans #11955 #11957]: #11958
+#11878 := (iff #5024 #11868)
+#11873 := (implies true #11868)
+#11876 := (iff #11873 #11868)
+#11877 := [rewrite]: #11876
+#11874 := (iff #5024 #11873)
+#11871 := (iff #5023 #11868)
+#11865 := (implies #4795 #11860)
+#11869 := (iff #11865 #11868)
+#11870 := [rewrite]: #11869
+#11866 := (iff #5023 #11865)
+#11863 := (iff #5022 #11860)
+#11856 := (implies #4948 #11851)
+#11861 := (iff #11856 #11860)
+#11862 := [rewrite]: #11861
+#11857 := (iff #5022 #11856)
+#11854 := (iff #5021 #11851)
+#11848 := (implies #4795 #11836)
+#11852 := (iff #11848 #11851)
+#11853 := [rewrite]: #11852
+#11849 := (iff #5021 #11848)
+#11846 := (iff #5020 #11836)
+#11841 := (implies true #11836)
+#11844 := (iff #11841 #11836)
+#11845 := [rewrite]: #11844
+#11842 := (iff #5020 #11841)
+#11839 := (iff #5019 #11836)
+#11833 := (implies #4795 #11830)
+#11837 := (iff #11833 #11836)
+#11838 := [rewrite]: #11837
+#11834 := (iff #5019 #11833)
+#11831 := (iff #5018 #11830)
+#11828 := (iff #5017 #11825)
+#11821 := (implies #4943 #11818)
+#11826 := (iff #11821 #11825)
+#11827 := [rewrite]: #11826
+#11822 := (iff #5017 #11821)
+#11819 := (iff #5016 #11818)
+#11816 := (iff #5015 #11813)
+#11809 := (implies #4946 #11804)
+#11814 := (iff #11809 #11813)
+#11815 := [rewrite]: #11814
+#11810 := (iff #5015 #11809)
+#11807 := (iff #5014 #11804)
+#11800 := (implies #4950 #11795)
+#11805 := (iff #11800 #11804)
+#11806 := [rewrite]: #11805
+#11801 := (iff #5014 #11800)
+#11798 := (iff #5013 #11795)
+#11791 := (implies #4955 #11786)
+#11796 := (iff #11791 #11795)
+#11797 := [rewrite]: #11796
+#11792 := (iff #5013 #11791)
+#11789 := (iff #5012 #11786)
+#11782 := (implies #4960 #11777)
+#11787 := (iff #11782 #11786)
+#11788 := [rewrite]: #11787
+#11783 := (iff #5012 #11782)
+#11780 := (iff #5011 #11777)
+#11773 := (implies #4794 #11761)
+#11778 := (iff #11773 #11777)
+#11779 := [rewrite]: #11778
+#11774 := (iff #5011 #11773)
+#11771 := (iff #5010 #11761)
+#11766 := (implies true #11761)
+#11769 := (iff #11766 #11761)
+#11770 := [rewrite]: #11769
+#11767 := (iff #5010 #11766)
+#11764 := (iff #5009 #11761)
+#11757 := (implies #4963 #11752)
+#11762 := (iff #11757 #11761)
+#11763 := [rewrite]: #11762
+#11758 := (iff #5009 #11757)
+#11755 := (iff #5008 #11752)
+#11748 := (implies #4965 #11736)
+#11753 := (iff #11748 #11752)
+#11754 := [rewrite]: #11753
+#11749 := (iff #5008 #11748)
+#11750 := [monotonicity #11747]: #11749
+#11756 := [trans #11750 #11754]: #11755
+#11759 := [monotonicity #11756]: #11758
+#11765 := [trans #11759 #11763]: #11764
+#11768 := [monotonicity #11765]: #11767
+#11772 := [trans #11768 #11770]: #11771
+#11614 := (iff #4961 #4794)
+#11615 := [rewrite]: #11614
+#11775 := [monotonicity #11615 #11772]: #11774
+#11781 := [trans #11775 #11779]: #11780
+#11784 := [monotonicity #11781]: #11783
+#11790 := [trans #11784 #11788]: #11789
+#11793 := [monotonicity #11790]: #11792
+#11799 := [trans #11793 #11797]: #11798
+#11802 := [monotonicity #11799]: #11801
+#11808 := [trans #11802 #11806]: #11807
+#11811 := [monotonicity #11808]: #11810
+#11817 := [trans #11811 #11815]: #11816
+#11820 := [monotonicity #11817]: #11819
+#11823 := [monotonicity #11820]: #11822
+#11829 := [trans #11823 #11827]: #11828
+#11832 := [monotonicity #11829]: #11831
+#11835 := [monotonicity #11832]: #11834
+#11840 := [trans #11835 #11838]: #11839
+#11843 := [monotonicity #11840]: #11842
+#11847 := [trans #11843 #11845]: #11846
+#11850 := [monotonicity #11847]: #11849
+#11855 := [trans #11850 #11853]: #11854
+#11858 := [monotonicity #11855]: #11857
+#11864 := [trans #11858 #11862]: #11863
+#11867 := [monotonicity #11864]: #11866
+#11872 := [trans #11867 #11870]: #11871
+#11875 := [monotonicity #11872]: #11874
+#11879 := [trans #11875 #11877]: #11878
+#11962 := [monotonicity #11879 #11959]: #11961
+#11965 := [monotonicity #11962]: #11964
+#11970 := [trans #11965 #11968]: #11969
+#11973 := [monotonicity #11970]: #11972
+#11978 := [trans #11973 #11976]: #11977
+#11981 := [monotonicity #11978]: #11980
+#11984 := [monotonicity #11981]: #11983
+#11989 := [trans #11984 #11987]: #11988
+#11992 := [monotonicity #11989]: #11991
+#11995 := [monotonicity #11992]: #11994
+#12000 := [trans #11995 #11998]: #11999
+#12003 := [monotonicity #12000]: #12002
+#12007 := [trans #12003 #12005]: #12006
+#12010 := [monotonicity #12007]: #12009
+#12015 := [trans #12010 #12013]: #12014
+#12018 := [monotonicity #12015]: #12017
+#12024 := [trans #12018 #12022]: #12023
+#12027 := [monotonicity #12024]: #12026
+#12032 := [trans #12027 #12030]: #12031
+#12035 := [monotonicity #12032]: #12034
+#12039 := [trans #12035 #12037]: #12038
+#12089 := [monotonicity #12039 #12086]: #12088
+#12092 := [monotonicity #12089]: #12091
+#12097 := [trans #12092 #12095]: #12096
+#11612 := (iff #4934 true)
+#11610 := (iff #11607 true)
+#11611 := [rewrite]: #11610
+#11608 := (iff #4934 #11607)
+#11605 := (iff #4933 true)
+#11606 := [rewrite]: #11605
+#11603 := (iff #4932 true)
+#11604 := [rewrite]: #11603
+#11609 := [monotonicity #11604 #11606]: #11608
+#11613 := [trans #11609 #11611]: #11612
+#12100 := [monotonicity #11613 #12097]: #12099
+#12104 := [trans #12100 #12102]: #12103
+#12107 := [monotonicity #12104]: #12106
+#12113 := [trans #12107 #12111]: #12112
+#12116 := [monotonicity #12113]: #12115
+#12122 := [trans #12116 #12120]: #12121
+#12125 := [monotonicity #12122]: #12124
+#12131 := [trans #12125 #12129]: #12130
+#12134 := [monotonicity #12131]: #12133
+#12140 := [trans #12134 #12138]: #12139
+#12143 := [monotonicity #12140]: #12142
+#12149 := [trans #12143 #12147]: #12148
+#12152 := [monotonicity #12149]: #12151
+#12157 := [trans #12152 #12155]: #12156
+#12160 := [monotonicity #12157]: #12159
+#12166 := [trans #12160 #12164]: #12165
+#11601 := (iff #4904 #4900)
+#11584 := (and true #4900)
+#11587 := (iff #11584 #4900)
+#11588 := [rewrite]: #11587
+#11599 := (iff #4904 #11584)
+#11597 := (iff #4903 #4900)
+#11595 := (iff #4903 #11584)
+#11593 := (iff #4902 #4900)
+#11591 := (iff #4902 #11584)
+#11589 := (iff #4901 #4900)
+#11585 := (iff #4901 #11584)
+#11582 := (iff #4889 true)
+#11577 := (forall (vars (?v0 S11)) (:pat #4885) true)
+#11580 := (iff #11577 true)
+#11581 := [elim-unused]: #11580
+#11578 := (iff #4889 #11577)
+#11575 := (iff #4888 true)
+#11539 := (implies #4873 #4873)
+#11542 := (iff #11539 true)
+#11543 := [rewrite]: #11542
+#11573 := (iff #4888 #11539)
+#11571 := (iff #4887 #4873)
+#11532 := (and true #4873)
+#11535 := (iff #11532 #4873)
+#11536 := [rewrite]: #11535
+#11569 := (iff #4887 #11532)
+#11567 := (iff #4886 true)
+#11568 := [rewrite]: #11567
+#11570 := [monotonicity #11568]: #11569
+#11572 := [trans #11570 #11536]: #11571
+#11574 := [monotonicity #11572]: #11573
+#11576 := [trans #11574 #11543]: #11575
+#11579 := [quant-intro #11576]: #11578
+#11583 := [trans #11579 #11581]: #11582
+#11586 := [monotonicity #11583]: #11585
+#11590 := [trans #11586 #11588]: #11589
+#11565 := (iff #4881 true)
+#11523 := (forall (vars (?v0 S11)) (:pat #4860) true)
+#11526 := (iff #11523 true)
+#11527 := [elim-unused]: #11526
+#11563 := (iff #4881 #11523)
+#11561 := (iff #4880 true)
+#11559 := (iff #4880 #11539)
+#11557 := (iff #4879 #4873)
+#11555 := (iff #4879 #11532)
+#11553 := (iff #4878 true)
+#11554 := [rewrite]: #11553
+#11556 := [monotonicity #11554]: #11555
+#11558 := [trans #11556 #11536]: #11557
+#11560 := [monotonicity #11558]: #11559
+#11562 := [trans #11560 #11543]: #11561
+#11564 := [quant-intro #11562]: #11563
+#11566 := [trans #11564 #11527]: #11565
+#11592 := [monotonicity #11566 #11590]: #11591
+#11594 := [trans #11592 #11588]: #11593
+#11551 := (iff #4877 true)
+#11546 := (forall (vars (?v0 S11)) (:pat #4871) true)
+#11549 := (iff #11546 true)
+#11550 := [elim-unused]: #11549
+#11547 := (iff #4877 #11546)
+#11544 := (iff #4876 true)
+#11540 := (iff #4876 #11539)
+#11537 := (iff #4875 #4873)
+#11533 := (iff #4875 #11532)
+#11530 := (iff #4874 true)
+#11531 := [rewrite]: #11530
+#11534 := [monotonicity #11531]: #11533
+#11538 := [trans #11534 #11536]: #11537
+#11541 := [monotonicity #11538]: #11540
+#11545 := [trans #11541 #11543]: #11544
+#11548 := [quant-intro #11545]: #11547
+#11552 := [trans #11548 #11550]: #11551
+#11596 := [monotonicity #11552 #11594]: #11595
+#11598 := [trans #11596 #11588]: #11597
+#11528 := (iff #4867 true)
+#11524 := (iff #4867 #11523)
+#11521 := (iff #4866 true)
+#11522 := [rewrite]: #11521
+#11525 := [quant-intro #11522]: #11524
+#11529 := [trans #11525 #11527]: #11528
+#11600 := [monotonicity #11529 #11598]: #11599
+#11602 := [trans #11600 #11588]: #11601
+#12169 := [monotonicity #11602 #12166]: #12168
+#12174 := [trans #12169 #12172]: #12173
+#12177 := [monotonicity #12174]: #12176
+#12182 := [trans #12177 #12180]: #12181
+#12185 := [monotonicity #12182]: #12184
+#12189 := [trans #12185 #12187]: #12188
+#12192 := [monotonicity #12189]: #12191
+#12197 := [trans #12192 #12195]: #12196
+#12200 := [monotonicity #12197]: #12199
+#12204 := [trans #12200 #12202]: #12203
+#12207 := [monotonicity #12204]: #12206
+#12212 := [trans #12207 #12210]: #12211
+#12215 := [monotonicity #12212]: #12214
+#12219 := [trans #12215 #12217]: #12218
+#11519 := (iff #4856 true)
+#11514 := (implies true true)
+#11517 := (iff #11514 true)
+#11518 := [rewrite]: #11517
+#11515 := (iff #4856 #11514)
+#11512 := (iff #4855 true)
+#11507 := (implies #4795 true)
+#11510 := (iff #11507 true)
+#11511 := [rewrite]: #11510
+#11508 := (iff #4855 #11507)
+#11505 := (iff #4854 true)
+#11472 := (or #11471 #11463)
+#11480 := (or #11404 #11472)
+#11495 := (or #11404 #11480)
+#11500 := (implies false #11495)
+#11503 := (iff #11500 true)
+#11504 := [rewrite]: #11503
+#11501 := (iff #4854 #11500)
+#11498 := (iff #4853 #11495)
+#11492 := (implies #4795 #11480)
+#11496 := (iff #11492 #11495)
+#11497 := [rewrite]: #11496
+#11493 := (iff #4853 #11492)
+#11490 := (iff #4852 #11480)
+#11485 := (implies true #11480)
+#11488 := (iff #11485 #11480)
+#11489 := [rewrite]: #11488
+#11486 := (iff #4852 #11485)
+#11483 := (iff #4851 #11480)
+#11477 := (implies #4795 #11472)
+#11481 := (iff #11477 #11480)
+#11482 := [rewrite]: #11481
+#11478 := (iff #4851 #11477)
+#11475 := (iff #4850 #11472)
+#11468 := (implies #4812 #11463)
+#11473 := (iff #11468 #11472)
+#11474 := [rewrite]: #11473
+#11469 := (iff #4850 #11468)
+#11470 := [monotonicity #11467]: #11469
+#11476 := [trans #11470 #11474]: #11475
+#11479 := [monotonicity #11476]: #11478
+#11484 := [trans #11479 #11482]: #11483
+#11487 := [monotonicity #11484]: #11486
+#11491 := [trans #11487 #11489]: #11490
+#11494 := [monotonicity #11491]: #11493
+#11499 := [trans #11494 #11497]: #11498
+#11502 := [monotonicity #11314 #11499]: #11501
+#11506 := [trans #11502 #11504]: #11505
+#11509 := [monotonicity #11506]: #11508
+#11513 := [trans #11509 #11511]: #11512
+#11516 := [monotonicity #11513]: #11515
+#11520 := [trans #11516 #11518]: #11519
+#12222 := [monotonicity #11520 #12219]: #12221
+#12226 := [trans #12222 #12224]: #12225
+#12229 := [monotonicity #12226]: #12228
+#12234 := [trans #12229 #12232]: #12233
+#12237 := [monotonicity #12234]: #12236
+#12243 := [trans #12237 #12241]: #12242
+#11311 := (iff #4801 #11310)
+#11308 := (iff #4800 #11305)
+#11302 := (implies #1522 #11299)
+#11306 := (iff #11302 #11305)
+#11307 := [rewrite]: #11306
+#11303 := (iff #4800 #11302)
+#11300 := (iff #4799 #11299)
+#11301 := [rewrite]: #11300
+#11304 := [monotonicity #11301]: #11303
+#11309 := [trans #11304 #11307]: #11308
+#11312 := [quant-intro #11309]: #11311
+#12246 := [monotonicity #11312 #12243]: #12245
+#12252 := [trans #12246 #12250]: #12251
+#12255 := [monotonicity #12252]: #12254
+#12261 := [trans #12255 #12259]: #12260
+#12264 := [monotonicity #12261]: #12263
+#12269 := [trans #12264 #12267]: #12268
+#12272 := [monotonicity #12269]: #12271
+#12278 := [trans #12272 #12276]: #12277
+#12281 := [monotonicity #12278]: #12280
+#12287 := [trans #12281 #12285]: #12286
+#12290 := [monotonicity #12287]: #12289
+#12296 := [trans #12290 #12294]: #12295
+#12299 := [monotonicity #12296]: #12298
+#12303 := [trans #12299 #12301]: #12302
+#12306 := [monotonicity #12303]: #12305
+#12312 := [trans #12306 #12310]: #12311
+#12315 := [monotonicity #12312]: #12314
+#11296 := (iff #4779 #11295)
+#11293 := (iff #4778 #11290)
+#11287 := (implies #1522 #11284)
+#11291 := (iff #11287 #11290)
+#11292 := [rewrite]: #11291
+#11288 := (iff #4778 #11287)
+#11285 := (iff #4777 #11284)
+#11286 := [rewrite]: #11285
+#11289 := [monotonicity #11286]: #11288
+#11294 := [trans #11289 #11292]: #11293
+#11297 := [quant-intro #11294]: #11296
+#12318 := [monotonicity #11297 #12315]: #12317
+#12324 := [trans #12318 #12322]: #12323
+#12327 := [monotonicity #11297 #12324]: #12326
+#12330 := [monotonicity #12327]: #12329
+#12336 := [trans #12330 #12334]: #12335
+#12339 := [monotonicity #12336]: #12338
+#11281 := (iff #4770 #11280)
+#11278 := (iff #4769 #11277)
+#11275 := (iff #4768 #4767)
+#11276 := [rewrite]: #11275
+#11279 := [monotonicity #11276]: #11278
+#11282 := [monotonicity #11279]: #11281
+#12342 := [monotonicity #11282 #12339]: #12341
+#12348 := [trans #12342 #12346]: #12347
+#12351 := [monotonicity #12348]: #12350
+#12357 := [trans #12351 #12355]: #12356
+#12360 := [monotonicity #12357]: #12359
+#12366 := [trans #12360 #12364]: #12365
+#12369 := [monotonicity #12366]: #12368
+#12375 := [trans #12369 #12373]: #12374
+#12378 := [monotonicity #12375]: #12377
+#12384 := [trans #12378 #12382]: #12383
+#12387 := [monotonicity #12384]: #12386
+#12393 := [trans #12387 #12391]: #12392
+#12396 := [monotonicity #12393]: #12395
+#12399 := [monotonicity #12396]: #12398
+#12405 := [trans #12399 #12403]: #12404
+#12408 := [monotonicity #12405]: #12407
+#12411 := [monotonicity #12408]: #12410
+#12417 := [trans #12411 #12415]: #12416
+#12420 := [monotonicity #12417]: #12419
+#12423 := [monotonicity #12420]: #12422
+#12429 := [trans #12423 #12427]: #12428
+#11273 := (iff #4724 #11272)
+#11270 := (iff #4723 #11269)
+#11271 := [rewrite]: #11270
+#11274 := [quant-intro #11271]: #11273
+#12432 := [monotonicity #11274 #12429]: #12431
+#12438 := [trans #12432 #12436]: #12437
+#12441 := [monotonicity #12438]: #12440
+#12447 := [trans #12441 #12445]: #12446
+#12450 := [monotonicity #12447]: #12449
+#12456 := [trans #12450 #12454]: #12455
+#12459 := [monotonicity #12456]: #12458
+#12465 := [trans #12459 #12463]: #12464
+#12468 := [monotonicity #12465]: #12467
+#12474 := [trans #12468 #12472]: #12473
+#12477 := [monotonicity #12474]: #12476
+#12483 := [trans #12477 #12481]: #12482
+#12486 := [monotonicity #12483]: #12485
+#12492 := [trans #12486 #12490]: #12491
+#12495 := [monotonicity #12492]: #12494
+#12499 := [trans #12495 #12497]: #12498
+#12502 := [monotonicity #12499]: #12501
+#12508 := [trans #12502 #12506]: #12507
+#12511 := [monotonicity #12508]: #12510
+#12517 := [trans #12511 #12515]: #12516
+#12520 := [monotonicity #12517]: #12519
+#12526 := [trans #12520 #12524]: #12525
+#12529 := [monotonicity #12526]: #12528
+#12535 := [trans #12529 #12533]: #12534
+#12538 := [monotonicity #12535]: #12537
+#12544 := [trans #12538 #12542]: #12543
+#12547 := [monotonicity #12544]: #12546
+#12553 := [trans #12547 #12551]: #12552
+#12556 := [monotonicity #12553]: #12555
+#12560 := [trans #12556 #12558]: #12559
+#12563 := [monotonicity #12560]: #12562
+#13451 := [trans #12563 #13449]: #13450
+#11268 := [asserted]: #5117
+#13452 := [mp #11268 #13451]: #13447
+#13464 := [not-or-elim #13452]: #13334
+#13467 := [and-elim #13464]: #4666
+#1254 := (f118 f123 #984)
+#4317 := (f45 #1254 #1287)
+#4318 := (pattern #4317)
+#2484 := (f55 f206 #984)
+#1329 := (f113 f114 #1287)
+#4320 := (f87 #1329 #2484)
+#4321 := (= #984 #4320)
+#4319 := (= #4317 f1)
+#11084 := (not #4319)
+#11085 := (or #11084 #4321)
+#11088 := (forall (vars (?v0 S11) (?v1 S4)) (:pat #4318) #11085)
+#16890 := (~ #11088 #11088)
+#16888 := (~ #11085 #11085)
+#16889 := [refl]: #16888
+#16891 := [nnf-pos #16889]: #16890
+#4322 := (implies #4319 #4321)
+#4323 := (forall (vars (?v0 S11) (?v1 S4)) (:pat #4318) #4322)
+#11089 := (iff #4323 #11088)
+#11086 := (iff #4322 #11085)
+#11087 := [rewrite]: #11086
+#11090 := [quant-intro #11087]: #11089
+#11083 := [asserted]: #4323
+#11093 := [mp #11083 #11090]: #11088
+#16892 := [mp~ #11093 #16891]: #11088
+#23412 := (not #4666)
+#23430 := (not #11088)
+#23431 := (or #23430 #23412 #23426)
+#23427 := (or #23412 #23426)
+#23432 := (or #23430 #23427)
+#23434 := (iff #23432 #23431)
+#23435 := [rewrite]: #23434
+#23433 := [quant-inst #4658 #4652]: #23432
+#23436 := [mp #23433 #23435]: #23431
+#24979 := [unit-resolution #23436 #16892 #13467]: #23426
+#23932 := [symm #24979]: #23931
+#23934 := [monotonicity #23932]: #23933
+#23936 := [trans #23934 #23916]: #23935
+#23938 := [monotonicity #23936]: #23937
+#23940 := [trans #23938 #23915]: #23939
+#23930 := [monotonicity #23940]: #23929
+#23951 := [trans #23930 #23949]: #23950
+#23953 := [symm #23951]: #23952
+#23956 := [monotonicity #23953]: #23955
+#21 := (= f5 f6)
+#22 := (not #21)
+decl f4 :: S2
+#8 := f4
+#19 := (= f4 f6)
+#20 := (not #19)
+#17 := (= f4 f5)
+#18 := (not #17)
+#12 := (= f3 f5)
+#13 := (not #12)
+#9 := (= f3 f4)
+#10 := (not #9)
+#5155 := (and #10 #13 #16 #18 #20 #22)
+#23 := (and #22 true)
+#24 := (and #20 #23)
+#25 := (and #18 #24)
+#26 := (and #16 #25)
+#27 := (and #13 #26)
+#28 := (and #10 #27)
+#5158 := (iff #28 #5155)
+#5140 := (and #20 #22)
+#5143 := (and #18 #5140)
+#5146 := (and #16 #5143)
+#5149 := (and #13 #5146)
+#5152 := (and #10 #5149)
+#5156 := (iff #5152 #5155)
+#5157 := [rewrite]: #5156
+#5153 := (iff #28 #5152)
+#5150 := (iff #27 #5149)
+#5147 := (iff #26 #5146)
+#5144 := (iff #25 #5143)
+#5141 := (iff #24 #5140)
+#5138 := (iff #23 #22)
+#5139 := [rewrite]: #5138
+#5142 := [monotonicity #5139]: #5141
+#5145 := [monotonicity #5142]: #5144
+#5148 := [monotonicity #5145]: #5147
+#5151 := [monotonicity #5148]: #5150
+#5154 := [monotonicity #5151]: #5153
+#5159 := [trans #5154 #5157]: #5158
+#5137 := [asserted]: #28
+#5160 := [mp #5137 #5159]: #5155
+#5163 := [and-elim #5160]: #16
+#23957 := [mp #5163 #23956]: #23954
+#23797 := (not #23789)
+#23800 := (not #23785)
+#23982 := (iff #12412 #23800)
+#23980 := (iff #4733 #23785)
+#23967 := (iff #23785 #4733)
+#23965 := (= #23784 #4732)
+#23960 := (= #23776 #4730)
+#23947 := (= #23775 #4729)
+#23959 := [monotonicity #23932]: #23947
+#23961 := [monotonicity #23959 #23932]: #23960
+#23966 := [monotonicity #23961]: #23965
+#23979 := [monotonicity #23966]: #23967
+#23981 := [symm #23979]: #23980
+#23983 := [monotonicity #23981]: #23982
+#23946 := [hypothesis]: #12412
+#23984 := [mp #23946 #23983]: #23800
+#23803 := (not #23790)
+#23804 := (or #23803 #23785 #23797)
+#23805 := [def-axiom]: #23804
+#23985 := [unit-resolution #23805 #23984 #23945]: #23797
+#23862 := (f71 #4667 #23413)
+#23863 := (= #23862 f1)
+#13468 := [and-elim #13464]: #4669
+#23986 := (= #23862 #4668)
+#23987 := [monotonicity #23932]: #23986
+#23988 := [trans #23987 #13468]: #23863
+#23858 := (f118 f123 #23413)
+#23859 := (f45 #23858 #23693)
+#23860 := (= #23859 f1)
+#23973 := (= #23859 #4665)
+#23974 := (= #23858 #4664)
+#23975 := [monotonicity #23932]: #23974
+#23976 := [monotonicity #23975 #23936]: #23973
+#23977 := [trans #23976 #13467]: #23860
+#23864 := (not #23863)
+#23861 := (not #23860)
+#24002 := (or #23861 #23864 #23866 #23789)
+#23699 := (f82 #4661 #23413)
+#23841 := (= #23699 f85)
+#13466 := [and-elim #13464]: #4663
+#23978 := (= #23699 #4662)
+#23994 := [monotonicity #23932]: #23978
+#23995 := [trans #23994 #13466]: #23841
+#13465 := [and-elim #13464]: #4660
+#23993 := (= #23704 #4659)
+#23996 := [monotonicity #23932]: #23993
+#23997 := [trans #23996 #13465]: #23705
+#23694 := (f45 f79 #23693)
+#23697 := (= #23694 f1)
+#13470 := [and-elim #13464]: #4674
+#23998 := (= #23694 #4673)
+#24003 := [monotonicity #23936]: #23998
+#24004 := [trans #24003 #13470]: #23697
+#13472 := [not-or-elim #13452]: #4687
+#13474 := [and-elim #13472]: #4686
+#1029 := (:var 1 S10)
+#3740 := (f334 f336 #1029)
+#3741 := (f125 #3740 #996)
+#3742 := (f71 #3741 #996)
+#3743 := (pattern #3742)
+#3750 := (= #3742 f1)
+#1000 := (f62 f63 #996)
+#1065 := (f45 f79 #1000)
+#1066 := (= #1065 f1)
+#10922 := (not #1066)
+#1001 := (f61 #1000)
+#1002 := (= #1001 f3)
+#1176 := (f80 f81 #1029)
+#1177 := (f71 #1176 #996)
+#1178 := (= #1177 f1)
+#11048 := (not #1178)
+#1173 := (f118 f123 #996)
+#1174 := (f45 #1173 #1000)
+#1175 := (= #1174 f1)
+#17928 := (not #1175)
+#1169 := (f83 f84 #1029)
+#1170 := (f82 #1169 #996)
+#1171 := (= #1170 f85)
+#17927 := (not #1171)
+#1159 := (f80 f86 #1029)
+#1160 := (f71 #1159 #996)
+#1161 := (= #1160 f1)
+#4045 := (not #1161)
+#1280 := (f115 f131 #1029)
+#1282 := (= #1280 f1)
+#18054 := (not #1282)
+#20507 := (or #18054 #4045 #17927 #17928 #11048 #1002 #10922 #3750)
+#20512 := (forall (vars (?v0 S10) (?v1 S11)) (:pat #3743) #20507)
+#1036 := (not #1002)
+#10634 := (and #1282 #1161 #1171 #1175 #1178 #1036 #1066)
+#10637 := (not #10634)
+#10640 := (or #10637 #3750)
+#10643 := (forall (vars (?v0 S10) (?v1 S11)) (:pat #3743) #10640)
+#20513 := (iff #10643 #20512)
+#20510 := (iff #10640 #20507)
+#20493 := (or #18054 #4045 #17927 #17928 #11048 #1002 #10922)
+#20504 := (or #20493 #3750)
+#20508 := (iff #20504 #20507)
+#20509 := [rewrite]: #20508
+#20505 := (iff #10640 #20504)
+#20502 := (iff #10637 #20493)
+#20494 := (not #20493)
+#20497 := (not #20494)
+#20500 := (iff #20497 #20493)
+#20501 := [rewrite]: #20500
+#20498 := (iff #10637 #20497)
+#20495 := (iff #10634 #20494)
+#20496 := [rewrite]: #20495
+#20499 := [monotonicity #20496]: #20498
+#20503 := [trans #20499 #20501]: #20502
+#20506 := [monotonicity #20503]: #20505
+#20511 := [trans #20506 #20509]: #20510
+#20514 := [quant-intro #20511]: #20513
+#16533 := (~ #10643 #10643)
+#16531 := (~ #10640 #10640)
+#16532 := [refl]: #16531
+#16534 := [nnf-pos #16532]: #16533
+#3744 := (and #1036 #1066)
+#3745 := (and #1178 #3744)
+#3746 := (and #1175 #3745)
+#3747 := (and #1171 #3746)
+#3748 := (and #1161 #3747)
+#3749 := (and #1282 #3748)
+#3751 := (implies #3749 #3750)
+#3752 := (forall (vars (?v0 S10) (?v1 S11)) (:pat #3743) #3751)
+#10646 := (iff #3752 #10643)
+#10626 := (not #3749)
+#10628 := (or #10626 #3750)
+#10631 := (forall (vars (?v0 S10) (?v1 S11)) (:pat #3743) #10628)
+#10644 := (iff #10631 #10643)
+#10641 := (iff #10628 #10640)
+#10638 := (iff #10626 #10637)
+#10635 := (iff #3749 #10634)
+#10636 := [rewrite]: #10635
+#10639 := [monotonicity #10636]: #10638
+#10642 := [monotonicity #10639]: #10641
+#10645 := [quant-intro #10642]: #10644
+#10632 := (iff #3752 #10631)
+#10629 := (iff #3751 #10628)
+#10630 := [rewrite]: #10629
+#10633 := [quant-intro #10630]: #10632
+#10647 := [trans #10633 #10645]: #10646
+#10625 := [asserted]: #3752
+#10648 := [mp #10625 #10647]: #10643
+#16535 := [mp~ #10648 #16534]: #10643
+#20515 := [mp #16535 #20514]: #20512
+#23698 := (not #23697)
+#23842 := (not #23841)
+#22428 := (not #4686)
+#23846 := (not #20512)
+#23844 := (or #23846 #22428 #23730 #23842 #23861 #23864 #23866 #23698 #23789)
+#23867 := (or #22428 #23730 #23842 #23861 #23864 #23866 #23698 #23789)
+#23847 := (or #23846 #23867)
+#23849 := (iff #23847 #23844)
+#23870 := [rewrite]: #23849
+#23848 := [quant-inst #4649 #23413]: #23847
+#23872 := [mp #23848 #23870]: #23844
+#24005 := [unit-resolution #23872 #20515 #13474 #24004 #23997 #23995]: #24002
+#24006 := [unit-resolution #24005 #23977 #23988 #23985 #23957]: false
+#24007 := [lemma #24006]: #4733
+#24421 := [trans #23966 #24007]: #23785
+#23794 := (or #23803 #23800 #23789)
+#23795 := [def-axiom]: #23794
+#24422 := [unit-resolution #23795 #24421 #23945]: #23789
+#23840 := (or #23797 #23839)
+#982 := (:var 2 S10)
+#3671 := (f334 f336 #982)
+#3672 := (f125 #3671 #984)
+#3673 := (f71 #3672 #996)
+#3753 := (pattern #3673)
+#3713 := (f66 f129 #980)
+#3754 := (f65 #3713 #993)
+#3755 := (f50 #1004 #3754)
+#3756 := (pattern #3755)
+#992 := (f59 f60 #980)
+#3658 := (f58 #992 #984)
+#3763 := (f329 f330 #3658)
+#3764 := (f50 #3762 #3763)
+#3765 := (= #3764 f1)
+#3760 := (= #3755 f1)
+#20516 := (not #3760)
+#2628 := (f62 f63 #993)
+#3757 := (f45 f337 #2628)
+#3758 := (= #3757 f1)
+#20531 := (or #3758 #20516 #3765)
+#20536 := (forall (vars (?v3 S11)) (:pat #3756) #20531)
+#20542 := (not #20536)
+#1219 := (f80 f86 #982)
+#1220 := (f71 #1219 #984)
+#1225 := (= #1220 f1)
+#3930 := (not #1225)
+#1021 := (f66 f67 #982)
+#3645 := (f65 #1021 #996)
+#2942 := (f51 f64 #984)
+#3646 := (f50 #2942 #3645)
+#3651 := (= #3646 f1)
+#20351 := (not #3651)
+#20543 := (or #20351 #3930 #20542)
+#20544 := (not #20543)
+#3674 := (= #3673 f1)
+#10666 := (not #3674)
+#20549 := (or #10666 #20544)
+#20552 := (forall (vars (?v0 S10) (?v1 S11) (?v2 S11)) (:pat #3753) #20549)
+#3759 := (not #3758)
+#3761 := (and #3759 #3760)
+#10650 := (not #3761)
+#10651 := (or #10650 #3765)
+#10654 := (forall (vars (?v3 S11)) (:pat #3756) #10651)
+#10675 := (and #3651 #1225 #10654)
+#10678 := (or #10666 #10675)
+#10681 := (forall (vars (?v0 S10) (?v1 S11) (?v2 S11)) (:pat #3753) #10678)
+#20553 := (iff #10681 #20552)
+#20550 := (iff #10678 #20549)
+#20547 := (iff #10675 #20544)
+#20539 := (and #3651 #1225 #20536)
+#20545 := (iff #20539 #20544)
+#20546 := [rewrite]: #20545
+#20540 := (iff #10675 #20539)
+#20537 := (iff #10654 #20536)
+#20534 := (iff #10651 #20531)
+#20517 := (or #3758 #20516)
+#20528 := (or #20517 #3765)
+#20532 := (iff #20528 #20531)
+#20533 := [rewrite]: #20532
+#20529 := (iff #10651 #20528)
+#20526 := (iff #10650 #20517)
+#20518 := (not #20517)
+#20521 := (not #20518)
+#20524 := (iff #20521 #20517)
+#20525 := [rewrite]: #20524
+#20522 := (iff #10650 #20521)
+#20519 := (iff #3761 #20518)
+#20520 := [rewrite]: #20519
+#20523 := [monotonicity #20520]: #20522
+#20527 := [trans #20523 #20525]: #20526
+#20530 := [monotonicity #20527]: #20529
+#20535 := [trans #20530 #20533]: #20534
+#20538 := [quant-intro #20535]: #20537
+#20541 := [monotonicity #20538]: #20540
+#20548 := [trans #20541 #20546]: #20547
+#20551 := [monotonicity #20548]: #20550
+#20554 := [quant-intro #20551]: #20553
+#16550 := (~ #10681 #10681)
+#16548 := (~ #10678 #10678)
+#16546 := (~ #10675 #10675)
+#16544 := (~ #10654 #10654)
+#16542 := (~ #10651 #10651)
+#16543 := [refl]: #16542
+#16545 := [nnf-pos #16543]: #16544
+#16540 := (~ #1225 #1225)
+#16541 := [refl]: #16540
+#16538 := (~ #3651 #3651)
+#16539 := [refl]: #16538
+#16547 := [monotonicity #16539 #16541 #16545]: #16546
+#16536 := (~ #10666 #10666)
+#16537 := [refl]: #16536
+#16549 := [monotonicity #16537 #16547]: #16548
+#16551 := [nnf-pos #16549]: #16550
+#3766 := (implies #3761 #3765)
+#3767 := (forall (vars (?v3 S11)) (:pat #3756) #3766)
+#3768 := (and #1225 #3767)
+#3769 := (and #3651 #3768)
+#3770 := (implies #3674 #3769)
+#3771 := (forall (vars (?v0 S10) (?v1 S11) (?v2 S11)) (:pat #3753) #3770)
+#10684 := (iff #3771 #10681)
+#10657 := (and #1225 #10654)
+#10660 := (and #3651 #10657)
+#10667 := (or #10666 #10660)
+#10672 := (forall (vars (?v0 S10) (?v1 S11) (?v2 S11)) (:pat #3753) #10667)
+#10682 := (iff #10672 #10681)
+#10679 := (iff #10667 #10678)
+#10676 := (iff #10660 #10675)
+#10677 := [rewrite]: #10676
+#10680 := [monotonicity #10677]: #10679
+#10683 := [quant-intro #10680]: #10682
+#10673 := (iff #3771 #10672)
+#10670 := (iff #3770 #10667)
+#10663 := (implies #3674 #10660)
+#10668 := (iff #10663 #10667)
+#10669 := [rewrite]: #10668
+#10664 := (iff #3770 #10663)
+#10661 := (iff #3769 #10660)
+#10658 := (iff #3768 #10657)
+#10655 := (iff #3767 #10654)
+#10652 := (iff #3766 #10651)
+#10653 := [rewrite]: #10652
+#10656 := [quant-intro #10653]: #10655
+#10659 := [monotonicity #10656]: #10658
+#10662 := [monotonicity #10659]: #10661
+#10665 := [monotonicity #10662]: #10664
+#10671 := [trans #10665 #10669]: #10670
+#10674 := [quant-intro #10671]: #10673
+#10685 := [trans #10674 #10683]: #10684
+#10649 := [asserted]: #3771
+#10686 := [mp #10649 #10685]: #10681
+#16552 := [mp~ #10686 #16551]: #10681
+#20555 := [mp #16552 #20554]: #20552
+#23816 := (not #20552)
+#23817 := (or #23816 #23797 #23839)
+#23813 := (or #23816 #23840)
+#23850 := (iff #23813 #23817)
+#23851 := [rewrite]: #23850
+#23818 := [quant-inst #4649 #23413 #23413]: #23813
+#23873 := [mp #23818 #23851]: #23817
+#24409 := [unit-resolution #23873 #20555]: #23840
+#24410 := [unit-resolution #24409 #24422]: #23839
+#23874 := (or #23838 #23821)
+#23875 := [def-axiom]: #23874
+#24408 := [unit-resolution #23875 #24410]: #23821
+#24413 := (= #24086 #23820)
+#24411 := (= #24085 #23810)
+#24985 := (= #24084 #23413)
+#24983 := (= #24084 #4658)
+#24981 := (= f445 #4657)
+#23487 := (= #4657 f445)
+#4357 := (f55 f206 #4356)
+#4358 := (= #4357 #1197)
+#21824 := (forall (vars (?v0 S4) (?v1 Int)) (:pat #21823) #4358)
+#4359 := (forall (vars (?v0 S4) (?v1 Int)) #4358)
+#21827 := (iff #4359 #21824)
+#21825 := (iff #4358 #4358)
+#21826 := [refl]: #21825
+#21828 := [quant-intro #21826]: #21827
+#16910 := (~ #4359 #4359)
+#16908 := (~ #4358 #4358)
+#16909 := [refl]: #16908
+#16911 := [nnf-pos #16909]: #16910
+#11096 := [asserted]: #4359
+#16912 := [mp~ #11096 #16911]: #4359
+#21829 := [mp #16912 #21828]: #21824
+#23460 := (not #21824)
+#23492 := (or #23460 #23487)
+#23493 := [quant-inst #356 #4655]: #23492
+#24980 := [unit-resolution #23493 #21829]: #23487
+#24982 := [symm #24980]: #24981
+#24984 := [monotonicity #24982]: #24983
+#24986 := [trans #24984 #24979]: #24985
+#24412 := [monotonicity #24986]: #24411
+#24414 := [monotonicity #24412]: #24413
+#24415 := [trans #24414 #24408]: #24087
+#24088 := (not #24087)
+#24420 := [hypothesis]: #24088
+#24416 := [unit-resolution #24420 #24415]: false
+#24429 := [lemma #24416]: #24087
+#21067 := (not #12642)
+#21969 := (or #21067 #12828 #11892 #11883 #12777 #21027 #21936)
+#21972 := (not #21969)
+#21951 := (or #17171 #17174 #21948)
+#21954 := (not #21951)
+#21957 := (or #17171 #17174 #21954)
+#21960 := (not #21957)
+#21963 := (or #12777 #21067 #12829 #21960)
+#21966 := (not #21963)
+#21975 := (or #21966 #21972)
+#21978 := (not #21975)
+#21981 := (or #17171 #17180 #12777 #21067 #21978)
+#21984 := (not #21981)
+#21987 := (or #17171 #17180 #21984)
+#21990 := (not #21987)
+#21993 := (or #17171 #17174 #21990)
+#21996 := (not #21993)
+#21999 := (or #17171 #17174 #21996)
+#22002 := (not #21999)
+#22005 := (or #12777 #21067 #12922 #22002)
+#22008 := (not #22005)
+#21158 := (not #4826)
+#21159 := (or #7428 #18181 #12950 #21158)
+#22019 := (forall (vars (?v0 Int)) (:pat #21878) #21159)
+#22024 := (not #22019)
+#21150 := (or #7428 #18181 #12950 #12964)
+#22011 := (forall (vars (?v0 Int)) (:pat #21878) #21150)
+#22016 := (not #22011)
+#22027 := (or #22016 #22024)
+#22030 := (not #22027)
+decl ?v0!15 :: Int
+#17354 := ?v0!15
+#17361 := (f140 #4734 ?v0!15)
+#17362 := (f139 #17361 f35)
+#17363 := (f55 #4748 #17362)
+#17678 := (* -1::Int #17363)
+#17679 := (+ f468 #17678)
+#17680 := (>= #17679 0::Int)
+#17665 := (* -1::Int ?v0!15)
+#17666 := (+ f443 #17665)
+#17667 := (<= #17666 0::Int)
+#17356 := (<= ?v0!15 4294967295::Int)
+#21124 := (not #17356)
+#17355 := (>= ?v0!15 0::Int)
+#21123 := (not #17355)
+#21139 := (or #21123 #21124 #17667 #17680)
+#21144 := (not #21139)
+#22033 := (or #21144 #22030)
+#22036 := (not #22033)
+#22039 := (or #12923 #12777 #21067 #11388 #11379 #11370 #11361 #22036)
+#22042 := (not #22039)
+#22045 := (or #22008 #22042)
+#22048 := (not #22045)
+#21211 := (not #4930)
+#21210 := (not #4925)
+#15031 := (not #4811)
+#21209 := (not #4806)
+#20942 := (or #7428 #18181 #13105 #13119)
+#21887 := (forall (vars (?v0 Int)) (:pat #21878) #20942)
+#21892 := (not #21887)
+#13751 := (<= f464 4294967295::Int)
+#21207 := (not #13751)
+#21206 := (not #13145)
+#13766 := (<= f463 4294967295::Int)
+#21205 := (not #13766)
+#2561 := 255::Int
+#13785 := (<= f462 255::Int)
+#21204 := (not #13785)
+#21203 := (not #13167)
+#17117 := (not #4780)
+#22051 := (or #12634 #17117 #21203 #21204 #21205 #21206 #21207 #12777 #21067 #13142 #21892 #13095 #21209 #13090 #15031 #12144 #12135 #12126 #12117 #21210 #21211 #22048)
+#22054 := (not #22051)
+#25641 := (iff #4750 #4780)
+#25637 := (iff #4780 #4750)
+#25638 := [commutativity]: #25637
+#25642 := [symm #25638]: #25641
+#22057 := (or #12634 #17117 #22054)
+#22060 := (not #22057)
+#20931 := (or #7428 #18181 #12601 #12613)
+#21879 := (forall (vars (?v0 Int)) (:pat #21878) #20931)
+#21884 := (not #21879)
+#22063 := (or #21884 #22060)
+#22066 := (not #22063)
+decl ?v0!13 :: Int
+#17090 := ?v0!13
+#17096 := (f140 #4734 ?v0!13)
+#17097 := (f139 #17096 f35)
+#17098 := (f55 #4748 #17097)
+#17099 := (* -1::Int #17098)
+#17100 := (+ f461 #17099)
+#17101 := (>= #17100 0::Int)
+#17095 := (>= ?v0!13 1::Int)
+#17092 := (<= ?v0!13 4294967295::Int)
+#20905 := (not #17092)
+#17091 := (>= ?v0!13 0::Int)
+#20904 := (not #17091)
+#20920 := (or #20904 #20905 #17095 #17101)
+#20925 := (not #20920)
+#22069 := (or #20925 #22066)
+#22072 := (not #22069)
+#22075 := (or #12598 #22072)
+#22078 := (not #22075)
+#22081 := (or #12598 #22078)
+#22084 := (not #22081)
+#17067 := (not #4745)
+#17058 := (not #4739)
+#22087 := (or #17058 #17067 #12379 #12370 #12361 #12352 #22084)
+#22090 := (not #22087)
+#24199 := (f71 #24190 #23991)
+#24200 := (= #24199 f1)
+#24197 := (f82 #4661 #23991)
+#24198 := (= #24197 f85)
+#24201 := (or #24198 #24200)
+#24202 := (not #24201)
+#24171 := (f62 f63 #23991)
+#24172 := (f61 #24171)
+#24173 := (= #24172 f3)
+#24203 := (or #24173 #24202)
+#24204 := (not #24203)
+#24175 := (f134 #4883 #23991)
+#24179 := (f235 f236 #24175)
+#24191 := (f71 #24190 #24179)
+#24192 := (= #24191 f1)
+#24188 := (f82 #4661 #24179)
+#24189 := (= #24188 f85)
+#24193 := (or #24189 #24192)
+#24194 := (not #24193)
+#24185 := (f62 f63 #24179)
+#24186 := (f61 #24185)
+#24187 := (= #24186 f3)
+#24180 := (f71 #4650 #24179)
+#24181 := (= #24180 f1)
+#24182 := (not #24181)
+#24176 := (f155 f237 #24175)
+#24177 := (= #24176 f1)
+#24178 := (not #24177)
+#24183 := (or #24178 #24182)
+#24184 := (not #24183)
+#24174 := (not #24173)
+#24195 := (or #24174 #24184 #24187 #24194)
+#24196 := (not #24195)
+#24205 := (or #24196 #24204)
+#24206 := (not #24205)
+#24168 := (f71 #4667 #23991)
+#24169 := (= #24168 f1)
+#23963 := (f134 #4883 #4736)
+#24093 := (f155 f237 #23963)
+#24094 := (= #24093 f1)
+#17061 := (not #4741)
+#24095 := (or #17061 #24094)
+#24096 := (not #24095)
+#24430 := [hypothesis]: #24095
+#13463 := [not-or-elim #13452]: #12635
+decl f78 :: S7
+#1061 := f78
+#4476 := (f45 f78 f35)
+#4477 := (= #4476 f1)
+#11138 := [asserted]: #4477
+#1291 := (f45 f78 #1287)
+#1306 := (:var 1 Int)
+#1917 := (:var 4 Int)
+#3555 := (f87 #1329 #1917)
+#3556 := (f153 f154 #3555)
+#3557 := (f140 #3556 #1306)
+#3558 := (f139 #3557 #1287)
+#2614 := (:var 5 S10)
+#3576 := (f83 f84 #2614)
+#3577 := (f82 #3576 #3558)
+#2604 := (:var 3 S11)
+#3552 := (f66 f67 #2614)
+#3553 := (f65 #3552 #2604)
+#1336 := (:var 2 Int)
+#3547 := (f216 f217 #1287)
+#3548 := (f215 #3547 #1336)
+#3549 := (f113 f114 #3548)
+#3550 := (f87 #3549 #1917)
+#3551 := (f51 f64 #3550)
+#3554 := (f50 #3551 #3553)
+#3578 := (pattern #3554 #3577 #1291)
+#2858 := (f137 f138 #2614)
+#2859 := (f135 f136 #2858)
+#3574 := (f134 #2859 #3558)
+#3575 := (pattern #3554 #3574 #1291)
+#3581 := (f155 f237 #3574)
+#3582 := (= #3581 f1)
+#2871 := (f80 f81 #2614)
+#3579 := (f71 #2871 #3558)
+#3580 := (= #3579 f1)
+#20261 := (not #3580)
+#20262 := (or #20261 #3582)
+#20263 := (not #20262)
+#6710 := (* -1::Int #1336)
+#8256 := (+ #1306 #6710)
+#8810 := (>= #8256 0::Int)
+#6842 := (>= #1306 0::Int)
+#18148 := (not #6842)
+#3563 := (= #3554 f1)
+#20237 := (not #3563)
+#1292 := (= #1291 f1)
+#10761 := (not #1292)
+#3561 := (f115 f131 #2614)
+#3562 := (= #3561 f1)
+#20236 := (not #3562)
+#20269 := (or #20236 #10761 #20237 #18148 #8810 #20263)
+#20274 := (forall (vars (?v0 S10) (?v1 Int) (?v2 S11) (?v3 Int) (?v4 Int) (?v5 S4)) (:pat #3575 #3578) #20269)
+#3583 := (not #3582)
+#3584 := (and #3580 #3583)
+#9575 := (not #8810)
+#10400 := (and #3562 #1292 #3563 #6842 #9575)
+#10405 := (not #10400)
+#10424 := (or #10405 #3584)
+#10427 := (forall (vars (?v0 S10) (?v1 Int) (?v2 S11) (?v3 Int) (?v4 Int) (?v5 S4)) (:pat #3575 #3578) #10424)
+#20275 := (iff #10427 #20274)
+#20272 := (iff #10424 #20269)
+#20238 := (or #20236 #10761 #20237 #18148 #8810)
+#20266 := (or #20238 #20263)
+#20270 := (iff #20266 #20269)
+#20271 := [rewrite]: #20270
+#20267 := (iff #10424 #20266)
+#20264 := (iff #3584 #20263)
+#20265 := [rewrite]: #20264
+#20247 := (iff #10405 #20238)
+#20239 := (not #20238)
+#20242 := (not #20239)
+#20245 := (iff #20242 #20238)
+#20246 := [rewrite]: #20245
+#20243 := (iff #10405 #20242)
+#20240 := (iff #10400 #20239)
+#20241 := [rewrite]: #20240
+#20244 := [monotonicity #20241]: #20243
+#20248 := [trans #20244 #20246]: #20247
+#20268 := [monotonicity #20248 #20265]: #20267
+#20273 := [trans #20268 #20271]: #20272
+#20276 := [quant-intro #20273]: #20275
+#16441 := (~ #10427 #10427)
+#16439 := (~ #10424 #10424)
+#16440 := [refl]: #16439
+#16442 := [nnf-pos #16440]: #16441
+#2706 := (< #1306 #1336)
+#1507 := (<= 0::Int #1306)
+#2707 := (and #1507 #2706)
+#3564 := (and #3563 #2707)
+#3565 := (and #1292 #3564)
+#3566 := (and #3562 #3565)
+#3585 := (implies #3566 #3584)
+#3586 := (forall (vars (?v0 S10) (?v1 Int) (?v2 S11) (?v3 Int) (?v4 Int) (?v5 S4)) (:pat #3575 #3578) #3585)
+#10430 := (iff #3586 #10427)
+#10384 := (not #3566)
+#10418 := (or #10384 #3584)
+#10421 := (forall (vars (?v0 S10) (?v1 Int) (?v2 S11) (?v3 Int) (?v4 Int) (?v5 S4)) (:pat #3575 #3578) #10418)
+#10428 := (iff #10421 #10427)
+#10425 := (iff #10418 #10424)
+#10406 := (iff #10384 #10405)
+#10403 := (iff #3566 #10400)
+#9578 := (and #6842 #9575)
+#10391 := (and #3563 #9578)
+#10394 := (and #1292 #10391)
+#10397 := (and #3562 #10394)
+#10401 := (iff #10397 #10400)
+#10402 := [rewrite]: #10401
+#10398 := (iff #3566 #10397)
+#10395 := (iff #3565 #10394)
+#10392 := (iff #3564 #10391)
+#9579 := (iff #2707 #9578)
+#9576 := (iff #2706 #9575)
+#9577 := [rewrite]: #9576
+#6841 := (iff #1507 #6842)
+#6843 := [rewrite]: #6841
+#9580 := [monotonicity #6843 #9577]: #9579
+#10393 := [monotonicity #9580]: #10392
+#10396 := [monotonicity #10393]: #10395
+#10399 := [monotonicity #10396]: #10398
+#10404 := [trans #10399 #10402]: #10403
+#10407 := [monotonicity #10404]: #10406
+#10426 := [monotonicity #10407]: #10425
+#10429 := [quant-intro #10426]: #10428
+#10422 := (iff #3586 #10421)
+#10419 := (iff #3585 #10418)
+#10420 := [rewrite]: #10419
+#10423 := [quant-intro #10420]: #10422
+#10431 := [trans #10423 #10429]: #10430
+#10417 := [asserted]: #3586
+#10432 := [mp #10417 #10431]: #10427
+#16443 := [mp~ #10432 #16442]: #10427
+#20277 := [mp #16443 #20276]: #20274
+#22809 := (not #4477)
+#24348 := (not #20274)
+#24349 := (or #24348 #22428 #22809 #24088 #12634 #24096)
+#24091 := (+ 0::Int #12568)
+#24092 := (>= #24091 0::Int)
+#24089 := (>= 0::Int 0::Int)
+#24090 := (not #24089)
+#24097 := (or #22428 #22809 #24088 #24090 #24092 #24096)
+#24372 := (or #24348 #24097)
+#24365 := (iff #24372 #24349)
+#24116 := (or #22428 #22809 #24088 #12634 #24096)
+#24417 := (or #24348 #24116)
+#24344 := (iff #24417 #24349)
+#24364 := [rewrite]: #24344
+#24418 := (iff #24372 #24417)
+#24119 := (iff #24097 #24116)
+#24113 := (or #22428 #22809 #24088 false #12634 #24096)
+#24117 := (iff #24113 #24116)
+#24118 := [rewrite]: #24117
+#24114 := (iff #24097 #24113)
+#24111 := (iff #24092 #12634)
+#24106 := (>= #12568 0::Int)
+#24109 := (iff #24106 #12634)
+#24110 := [rewrite]: #24109
+#24107 := (iff #24092 #24106)
+#24104 := (= #24091 #12568)
+#24105 := [rewrite]: #24104
+#24108 := [monotonicity #24105]: #24107
+#24112 := [trans #24108 #24110]: #24111
+#24102 := (iff #24090 false)
+#24100 := (iff #24090 #4808)
+#24098 := (iff #24089 true)
+#24099 := [rewrite]: #24098
+#24101 := [monotonicity #24099]: #24100
+#24103 := [trans #24101 #11314]: #24102
+#24115 := [monotonicity #24103 #24112]: #24114
+#24120 := [trans #24115 #24118]: #24119
+#24419 := [monotonicity #24120]: #24418
+#24366 := [trans #24419 #24364]: #24365
+#24373 := [quant-inst #4649 #4655 #23413 #4646 #1138 #356]: #24372
+#24367 := [mp #24373 #24366]: #24349
+#24452 := [unit-resolution #24367 #20277 #11138 #13463 #13474 #24429 #24430]: false
+#24453 := [lemma #24452]: #24096
+#24325 := (or #24095 #4741)
+#24326 := [def-axiom]: #24325
+#25074 := [unit-resolution #24326 #24453]: #4741
+#25101 := (= #24168 #4740)
+#25097 := (= #23991 #4736)
+#23992 := (= #4736 #23991)
+#24000 := (f62 f63 #4736)
+#24001 := (= #24000 f35)
+#23483 := (f62 f63 #4656)
+#23484 := (= #23483 f35)
+#23489 := (or #23455 #23484)
+#23490 := [quant-inst #356 #4655]: #23489
+#24454 := [unit-resolution #23490 #21835]: #23484
+#24485 := (= #24000 #23483)
+#24459 := (= #4736 #4656)
+#24041 := (f87 #4654 #4657)
+#24457 := (= #24041 #4656)
+#24458 := [monotonicity #24980]: #24457
+#24044 := (= #4736 #24041)
+#24047 := (not #24044)
+decl f243 :: S54
+#2898 := f243
+#24009 := (f125 f243 #4736)
+#24010 := (f71 #24009 #4656)
+#24023 := (= #24010 f1)
+#24024 := (not #24023)
+#24050 := (or #24024 #24047)
+#24053 := (not #24050)
+#2626 := (f153 f154 #993)
+#2627 := (f140 #2626 #1306)
+#2896 := (f139 #2627 #1287)
+#2897 := (pattern #2896)
+#2904 := (f244 f245 #1287)
+#2905 := (* #1306 #2904)
+#2902 := (f55 f206 #993)
+#2906 := (+ #2902 #2905)
+#2907 := (f87 #1329 #2906)
+#2908 := (= #2896 #2907)
+#19805 := (not #2908)
+#2899 := (f125 f243 #2896)
+#2900 := (f71 #2899 #993)
+#2901 := (= #2900 f1)
+#19804 := (not #2901)
+#19806 := (or #19804 #19805)
+#19807 := (not #19806)
+#19810 := (forall (vars (?v0 S11) (?v1 Int) (?v2 S4)) (:pat #2897) #19807)
+#2909 := (and #2901 #2908)
+#2910 := (forall (vars (?v0 S11) (?v1 Int) (?v2 S4)) (:pat #2897) #2909)
+#19811 := (iff #2910 #19810)
+#19808 := (iff #2909 #19807)
+#19809 := [rewrite]: #19808
+#19812 := [quant-intro #19809]: #19811
+#16084 := (~ #2910 #2910)
+#16082 := (~ #2909 #2909)
+#16083 := [refl]: #16082
+#16085 := [nnf-pos #16083]: #16084
+#9870 := [asserted]: #2910
+#16086 := [mp~ #9870 #16085]: #2910
+#19813 := [mp #16086 #19812]: #19810
+#24299 := (not #19810)
+#24336 := (or #24299 #24053)
+#24025 := (* 0::Int #4624)
+#24026 := (+ #4657 #24025)
+#24027 := (f87 #4654 #24026)
+#24028 := (= #4736 #24027)
+#24029 := (not #24028)
+#24030 := (or #24024 #24029)
+#24031 := (not #24030)
+#24335 := (or #24299 #24031)
+#24337 := (iff #24335 #24336)
+#24301 := (iff #24336 #24336)
+#24339 := [rewrite]: #24301
+#24054 := (iff #24031 #24053)
+#24051 := (iff #24030 #24050)
+#24048 := (iff #24029 #24047)
+#24045 := (iff #24028 #24044)
+#24042 := (= #24027 #24041)
+#24039 := (= #24026 #4657)
+#24034 := (+ #4657 0::Int)
+#24037 := (= #24034 #4657)
+#24038 := [rewrite]: #24037
+#24035 := (= #24026 #24034)
+#24032 := (= #24025 0::Int)
+#24033 := [rewrite]: #24032
+#24036 := [monotonicity #24033]: #24035
+#24040 := [trans #24036 #24038]: #24039
+#24043 := [monotonicity #24040]: #24042
+#24046 := [monotonicity #24043]: #24045
+#24049 := [monotonicity #24046]: #24048
+#24052 := [monotonicity #24049]: #24051
+#24055 := [monotonicity #24052]: #24054
+#24338 := [monotonicity #24055]: #24337
+#24343 := [trans #24338 #24339]: #24337
+#24300 := [quant-inst #4656 #1138 #356]: #24335
+#24292 := [mp #24300 #24343]: #24336
+#24455 := [unit-resolution #24292 #19813]: #24053
+#24294 := (or #24050 #24044)
+#24350 := [def-axiom]: #24294
+#24456 := [unit-resolution #24350 #24455]: #24044
+#24484 := [trans #24456 #24458]: #24459
+#24486 := [monotonicity #24484]: #24485
+#24487 := [trans #24486 #24454]: #24001
+#24302 := (not #24001)
+#24008 := (iff #4739 #24001)
+#2640 := (f62 f63 #984)
+#3307 := (= #2640 #1287)
+#4324 := (iff #4319 #3307)
+#21817 := (forall (vars (?v0 S11) (?v1 S4)) (:pat #4318) #4324)
+#4325 := (forall (vars (?v0 S11) (?v1 S4)) #4324)
+#21820 := (iff #4325 #21817)
+#21818 := (iff #4324 #4324)
+#21819 := [refl]: #21818
+#21821 := [quant-intro #21819]: #21820
+#16895 := (~ #4325 #4325)
+#16893 := (~ #4324 #4324)
+#16894 := [refl]: #16893
+#16896 := [nnf-pos #16894]: #16895
+#11091 := [asserted]: #4325
+#16897 := [mp~ #11091 #16896]: #4325
+#21822 := [mp #16897 #21821]: #21817
+#23440 := (not #21817)
+#24334 := (or #23440 #24008)
+#24303 := [quant-inst #4736 #356]: #24334
+#24368 := [unit-resolution #24303 #21822]: #24008
+#24309 := (not #24008)
+#24358 := (or #24309 #24302)
+#24345 := [hypothesis]: #17058
+#24310 := (or #24309 #4739 #24302)
+#24323 := [def-axiom]: #24310
+#24361 := [unit-resolution #24323 #24345]: #24358
+#24451 := [unit-resolution #24361 #24368]: #24302
+#24488 := [unit-resolution #24451 #24487]: false
+#24483 := [lemma #24488]: #4739
+#24526 := (or #23430 #17058 #23992)
+#23999 := (or #17058 #23992)
+#24527 := (or #23430 #23999)
+#24529 := (iff #24527 #24526)
+#24530 := [rewrite]: #24529
+#24528 := [quant-inst #4736 #356]: #24527
+#24525 := [mp #24528 #24530]: #24526
+#25084 := [unit-resolution #24525 #16892 #24483]: #23992
+#25100 := [symm #25084]: #25097
+#25102 := [monotonicity #25100]: #25101
+#25104 := [trans #25102 #25074]: #24169
+#24170 := (not #24169)
+#24207 := (or #24170 #24206)
+#24208 := (not #24207)
+#24163 := (f71 #4743 #23991)
+#24164 := (= #24163 f1)
+#24209 := (iff #24164 #24208)
+#1373 := (f80 f157 #1029)
+#3957 := (f71 #1373 #996)
+#3958 := (pattern #3957)
+#3976 := (f80 f358 #1029)
+#3983 := (f71 #3976 #996)
+#3984 := (= #3983 f1)
+#3985 := (or #1171 #3984)
+#20658 := (not #3985)
+#20659 := (or #1002 #20658)
+#20660 := (not #20659)
+#1359 := (f137 f138 #1029)
+#1360 := (f135 f136 #1359)
+#3960 := (f134 #1360 #996)
+#3964 := (f235 f236 #3960)
+#3977 := (f71 #3976 #3964)
+#3978 := (= #3977 f1)
+#3973 := (f82 #1169 #3964)
+#3974 := (= #3973 f85)
+#3979 := (or #3974 #3978)
+#20653 := (not #3979)
+#3969 := (f62 f63 #3964)
+#3970 := (f61 #3969)
+#3971 := (= #3970 f3)
+#3965 := (f71 #1159 #3964)
+#3966 := (= #3965 f1)
+#3967 := (not #3966)
+#3961 := (f155 f237 #3960)
+#3962 := (= #3961 f1)
+#3963 := (not #3962)
+#3968 := (or #3963 #3967)
+#20652 := (not #3968)
+#20654 := (or #1036 #20652 #3971 #20653)
+#20655 := (not #20654)
+#20663 := (or #20655 #20660)
+#20669 := (not #20663)
+#20670 := (or #11048 #20669)
+#20671 := (not #20670)
+#3959 := (= #3957 f1)
+#20676 := (iff #3959 #20671)
+#20679 := (forall (vars (?v0 S10) (?v1 S11)) (:pat #3958) #20676)
+#3986 := (and #1036 #3985)
+#3972 := (not #3971)
+#10834 := (and #1002 #3968 #3972 #3979)
+#10837 := (or #10834 #3986)
+#10840 := (and #1178 #10837)
+#10843 := (iff #3959 #10840)
+#10846 := (forall (vars (?v0 S10) (?v1 S11)) (:pat #3958) #10843)
+#20680 := (iff #10846 #20679)
+#20677 := (iff #10843 #20676)
+#20674 := (iff #10840 #20671)
+#20666 := (and #1178 #20663)
+#20672 := (iff #20666 #20671)
+#20673 := [rewrite]: #20672
+#20667 := (iff #10840 #20666)
+#20664 := (iff #10837 #20663)
+#20661 := (iff #3986 #20660)
+#20662 := [rewrite]: #20661
+#20656 := (iff #10834 #20655)
+#20657 := [rewrite]: #20656
+#20665 := [monotonicity #20657 #20662]: #20664
+#20668 := [monotonicity #20665]: #20667
+#20675 := [trans #20668 #20673]: #20674
+#20678 := [monotonicity #20675]: #20677
+#20681 := [quant-intro #20678]: #20680
+#16655 := (~ #10846 #10846)
+#16653 := (~ #10843 #10843)
+#16654 := [refl]: #16653
+#16656 := [nnf-pos #16654]: #16655
+#3980 := (and #3972 #3979)
+#3981 := (and #3968 #3980)
+#3982 := (and #1002 #3981)
+#3987 := (or #3982 #3986)
+#3988 := (and #1178 #3987)
+#3989 := (iff #3959 #3988)
+#3990 := (forall (vars (?v0 S10) (?v1 S11)) (:pat #3958) #3989)
+#10847 := (iff #3990 #10846)
+#10844 := (iff #3989 #10843)
+#10841 := (iff #3988 #10840)
+#10838 := (iff #3987 #10837)
+#10835 := (iff #3982 #10834)
+#10836 := [rewrite]: #10835
+#10839 := [monotonicity #10836]: #10838
+#10842 := [monotonicity #10839]: #10841
+#10845 := [monotonicity #10842]: #10844
+#10848 := [quant-intro #10845]: #10847
+#10830 := [asserted]: #3990
+#10849 := [mp #10830 #10848]: #10846
+#16657 := [mp~ #10849 #16656]: #10846
+#20682 := [mp #16657 #20681]: #20679
+#24794 := (not #20679)
+#24803 := (or #24794 #24209)
+#24804 := [quant-inst #4649 #23991]: #24803
+#24792 := [unit-resolution #24804 #20682]: #24209
+#24544 := (not #24164)
+#25021 := (iff #17067 #24544)
+#25015 := (iff #4745 #24164)
+#24960 := (iff #24164 #4745)
+#24958 := (= #24163 #4744)
+#24959 := [monotonicity #25100]: #24958
+#25018 := [monotonicity #24959]: #24960
+#25016 := [symm #25018]: #25015
+#25022 := [monotonicity #25016]: #25021
+#24793 := [hypothesis]: #17067
+#25004 := [mp #24793 #25022]: #24544
+#24541 := (not #24209)
+#24542 := (or #24541 #24164 #24207)
+#24543 := [def-axiom]: #24542
+#25051 := [unit-resolution #24543 #25004 #24792]: #24207
+#24751 := (or #24208 #24170 #24206)
+#24540 := [def-axiom]: #24751
+#25052 := [unit-resolution #24540 #25051 #25104]: #24206
+#22792 := (f61 f35)
+#22793 := (= #22792 f3)
+#22800 := (iff #4477 #22793)
+#3856 := (pattern #1291)
+#4540 := (= #4530 f3)
+#4541 := (iff #1292 #4540)
+#4542 := (forall (vars (?v0 S4)) (:pat #3856) #4541)
+#17010 := (~ #4542 #4542)
+#17008 := (~ #4541 #4541)
+#17009 := [refl]: #17008
+#17011 := [nnf-pos #17009]: #17010
+#11188 := [asserted]: #4542
+#17012 := [mp~ #11188 #17011]: #4542
+#22524 := (not #4542)
+#22803 := (or #22524 #22800)
+#22804 := [quant-inst #356]: #22803
+#25017 := [unit-resolution #22804 #17012]: #22800
+#22805 := (not #22800)
+#24919 := (or #22805 #22793)
+#22810 := (or #22805 #22809 #22793)
+#22811 := [def-axiom]: #22810
+#24920 := [unit-resolution #22811 #11138]: #24919
+#24539 := [unit-resolution #24920 #25017]: #22793
+#25055 := (= #24172 #22792)
+#25063 := (= #24171 f35)
+#25049 := (or #24309 #24001)
+#24531 := (or #24309 #17058 #24001)
+#24532 := [def-axiom]: #24531
+#25050 := [unit-resolution #24532 #24483]: #25049
+#25053 := [unit-resolution #25050 #24368]: #24001
+#25054 := (= #24171 #24000)
+#24643 := [monotonicity #25100]: #25054
+#25064 := [trans #24643 #25053]: #25063
+#25056 := [monotonicity #25064]: #25055
+#25048 := [trans #25056 #24539]: #24173
+#24296 := (not #24094)
+#25070 := (iff #24296 #24178)
+#24731 := (iff #24094 #24177)
+#25057 := (iff #24177 #24094)
+#24648 := (= #24176 #24093)
+#25286 := (= #24175 #23963)
+#25287 := [monotonicity #25100]: #25286
+#25068 := [monotonicity #25287]: #24648
+#25047 := [monotonicity #25068]: #25057
+#24629 := [symm #25047]: #24731
+#25072 := [monotonicity #24629]: #25070
+#24297 := (or #24095 #24296)
+#24295 := [def-axiom]: #24297
+#24647 := [unit-resolution #24295 #24453]: #24296
+#25073 := [mp #24647 #25072]: #24178
+#24805 := (or #24183 #24177)
+#24806 := [def-axiom]: #24805
+#25108 := [unit-resolution #24806 #25073]: #24183
+#25117 := (or #24196 #24174 #24184)
+#24890 := (f55 f206 #23413)
+#25219 := (f87 #4654 #24890)
+#25193 := (f153 f154 #23413)
+#25194 := (f140 #25193 0::Int)
+#25201 := (f139 #25194 f35)
+#25222 := (= #25201 #25219)
+#25225 := (not #25222)
+#25202 := (f125 f243 #25201)
+#25203 := (f71 #25202 #23413)
+#25204 := (= #25203 f1)
+#25205 := (not #25204)
+#25228 := (or #25205 #25225)
+#25231 := (not #25228)
+#25337 := [hypothesis]: #25228
+#25234 := (or #24299 #25231)
+#25206 := (+ #24890 #24025)
+#25207 := (f87 #4654 #25206)
+#25208 := (= #25201 #25207)
+#25209 := (not #25208)
+#25210 := (or #25205 #25209)
+#25211 := (not #25210)
+#25235 := (or #24299 #25211)
+#25237 := (iff #25235 #25234)
+#25239 := (iff #25234 #25234)
+#25240 := [rewrite]: #25239
+#25232 := (iff #25211 #25231)
+#25229 := (iff #25210 #25228)
+#25226 := (iff #25209 #25225)
+#25223 := (iff #25208 #25222)
+#25220 := (= #25207 #25219)
+#25217 := (= #25206 #24890)
+#25212 := (+ #24890 0::Int)
+#25215 := (= #25212 #24890)
+#25216 := [rewrite]: #25215
+#25213 := (= #25206 #25212)
+#25214 := [monotonicity #24033]: #25213
+#25218 := [trans #25214 #25216]: #25217
+#25221 := [monotonicity #25218]: #25220
+#25224 := [monotonicity #25221]: #25223
+#25227 := [monotonicity #25224]: #25226
+#25230 := [monotonicity #25227]: #25229
+#25233 := [monotonicity #25230]: #25232
+#25238 := [monotonicity #25233]: #25237
+#25241 := [trans #25238 #25240]: #25237
+#25236 := [quant-inst #23413 #1138 #356]: #25235
+#25242 := [mp #25236 #25241]: #25234
+#25338 := [unit-resolution #25242 #19813 #25337]: false
+#25339 := [lemma #25338]: #25231
+#25245 := (or #25228 #25222)
+#25246 := [def-axiom]: #25245
+#25109 := [unit-resolution #25246 #25339]: #25222
+#25335 := (or #25225 #24189)
+#25331 := (= #24188 #4662)
+#25298 := (= #24179 #4658)
+#25296 := (= #24179 #24084)
+#25120 := (f153 f154 #24084)
+#25121 := (f140 #25120 0::Int)
+#25122 := (f139 #25121 f35)
+#25123 := (f134 #4883 #25122)
+#25124 := (f235 f236 #25123)
+#25125 := (= #25124 #24084)
+#25132 := (f71 #4667 #25122)
+#25133 := (= #25132 f1)
+#25134 := (not #25133)
+decl f156 :: S69
+#1366 := f156
+#25129 := (f155 f156 #25123)
+#25130 := (= #25129 f1)
+#25131 := (not #25130)
+#25127 := (f155 f237 #25123)
+#25128 := (= #25127 f1)
+#25126 := (not #25125)
+#25135 := (or #25126 #25128 #25131 #25134)
+#25136 := (not #25135)
+#25190 := [hypothesis]: #25135
+#25111 := (f71 #4667 #24084)
+#25112 := (= #25111 f1)
+#25182 := (= #25111 #4668)
+#25183 := [monotonicity #24984]: #25182
+#25184 := [trans #25183 #13468]: #25112
+#25119 := (not #25112)
+#25181 := [hypothesis]: #25119
+#25185 := [unit-resolution #25181 #25184]: false
+#25186 := [lemma #25185]: #25112
+#1351 := (:var 3 Int)
+#1398 := (:var 2 S4)
+#2758 := (f216 f217 #1398)
+#2759 := (f215 #2758 #1306)
+#2760 := (f113 f114 #2759)
+#2761 := (f87 #2760 #1351)
+#2603 := (f113 f114 #1398)
+#2753 := (f87 #2603 #1351)
+#2754 := (f153 f154 #2753)
+#2755 := (f140 #2754 #1197)
+#2756 := (f139 #2755 #1398)
+#1010 := (:var 4 S10)
+#2763 := (f137 f138 #1010)
+#2764 := (f135 f136 #2763)
+#2765 := (f134 #2764 #2756)
+#2766 := (pattern #2765 #2761)
+#2751 := (f110 f111 #1010)
+#2752 := (f108 f109 #2751)
+#2757 := (f107 #2752 #2756)
+#2762 := (pattern #2757 #2761)
+#2771 := (f153 f154 #2761)
+#2772 := (f140 #2771 #1197)
+#2773 := (f139 #2772 #1398)
+#2767 := (f80 f81 #1010)
+#2783 := (f71 #2767 #2773)
+#2784 := (= #2783 f1)
+#19581 := (not #2784)
+#2774 := (f134 #2764 #2773)
+#2781 := (f155 f156 #2774)
+#2782 := (= #2781 f1)
+#19580 := (not #2782)
+#2778 := (f155 f237 #2774)
+#2779 := (= #2778 f1)
+#2775 := (f235 f236 #2774)
+#2776 := (= #2775 #2761)
+#19579 := (not #2776)
+#19582 := (or #19579 #2779 #19580 #19581)
+#19583 := (not #19582)
+#7650 := (* -1::Int #1306)
+#8261 := (+ #1197 #7650)
+#8262 := (>= #8261 0::Int)
+#2768 := (f71 #2767 #2761)
+#2769 := (= #2768 f1)
+#9684 := (not #2769)
+#19589 := (or #9684 #7428 #8262 #19583)
+#19594 := (forall (vars (?v0 S10) (?v1 Int) (?v2 S4) (?v3 Int) (?v4 Int)) (:pat #2762 #2766) #19589)
+#2780 := (not #2779)
+#9693 := (and #2776 #2780 #2782 #2784)
+#9479 := (not #8262)
+#9482 := (and #6706 #9479)
+#9485 := (not #9482)
+#9702 := (or #9684 #9485 #9693)
+#9707 := (forall (vars (?v0 S10) (?v1 Int) (?v2 S4) (?v3 Int) (?v4 Int)) (:pat #2762 #2766) #9702)
+#19595 := (iff #9707 #19594)
+#19592 := (iff #9702 #19589)
+#19464 := (or #7428 #8262)
+#19586 := (or #9684 #19464 #19583)
+#19590 := (iff #19586 #19589)
+#19591 := [rewrite]: #19590
+#19587 := (iff #9702 #19586)
+#19584 := (iff #9693 #19583)
+#19585 := [rewrite]: #19584
+#19473 := (iff #9485 #19464)
+#19465 := (not #19464)
+#19468 := (not #19465)
+#19471 := (iff #19468 #19464)
+#19472 := [rewrite]: #19471
+#19469 := (iff #9485 #19468)
+#19466 := (iff #9482 #19465)
+#19467 := [rewrite]: #19466
+#19470 := [monotonicity #19467]: #19469
+#19474 := [trans #19470 #19472]: #19473
+#19588 := [monotonicity #19474 #19585]: #19587
+#19593 := [trans #19588 #19591]: #19592
+#19596 := [quant-intro #19593]: #19595
+#15871 := (~ #9707 #9707)
+#15869 := (~ #9702 #9702)
+#15870 := [refl]: #15869
+#15872 := [nnf-pos #15870]: #15871
+#2785 := (and #2782 #2784)
+#2786 := (and #2780 #2785)
+#2787 := (and #2776 #2786)
+#2612 := (< #1197 #1306)
+#2613 := (and #1363 #2612)
+#2788 := (implies #2613 #2787)
+#2789 := (implies #2769 #2788)
+#2790 := (forall (vars (?v0 S10) (?v1 Int) (?v2 S4) (?v3 Int) (?v4 Int)) (:pat #2762 #2766) #2789)
+#9710 := (iff #2790 #9707)
+#9451 := (not #2613)
+#9678 := (or #9451 #2787)
+#9685 := (or #9684 #9678)
+#9690 := (forall (vars (?v0 S10) (?v1 Int) (?v2 S4) (?v3 Int) (?v4 Int)) (:pat #2762 #2766) #9685)
+#9708 := (iff #9690 #9707)
+#9705 := (iff #9685 #9702)
+#9696 := (or #9485 #9693)
+#9699 := (or #9684 #9696)
+#9703 := (iff #9699 #9702)
+#9704 := [rewrite]: #9703
+#9700 := (iff #9685 #9699)
+#9697 := (iff #9678 #9696)
+#9694 := (iff #2787 #9693)
+#9695 := [rewrite]: #9694
+#9486 := (iff #9451 #9485)
+#9483 := (iff #2613 #9482)
+#9480 := (iff #2612 #9479)
+#9481 := [rewrite]: #9480
+#9484 := [monotonicity #6705 #9481]: #9483
+#9487 := [monotonicity #9484]: #9486
+#9698 := [monotonicity #9487 #9695]: #9697
+#9701 := [monotonicity #9698]: #9700
+#9706 := [trans #9701 #9704]: #9705
+#9709 := [quant-intro #9706]: #9708
+#9691 := (iff #2790 #9690)
+#9688 := (iff #2789 #9685)
+#9681 := (implies #2769 #9678)
+#9686 := (iff #9681 #9685)
+#9687 := [rewrite]: #9686
+#9682 := (iff #2789 #9681)
+#9679 := (iff #2788 #9678)
+#9680 := [rewrite]: #9679
+#9683 := [monotonicity #9680]: #9682
+#9689 := [trans #9683 #9687]: #9688
+#9692 := [quant-intro #9689]: #9691
+#9711 := [trans #9692 #9709]: #9710
+#9677 := [asserted]: #2790
+#9712 := [mp #9677 #9711]: #9707
+#15873 := [mp~ #9712 #15872]: #9707
+#19597 := [mp #15873 #19596]: #19594
+#25115 := (not #19594)
+#25113 := (or #25115 #25119 #12634 #25136)
+#25137 := (or #25119 #24090 #24092 #25136)
+#25147 := (or #25115 #25137)
+#25160 := (iff #25147 #25113)
+#25141 := (or #25119 #12634 #25136)
+#25154 := (or #25115 #25141)
+#25157 := (iff #25154 #25113)
+#25158 := [rewrite]: #25157
+#25155 := (iff #25147 #25154)
+#25144 := (iff #25137 #25141)
+#25138 := (or #25119 false #12634 #25136)
+#25142 := (iff #25138 #25141)
+#25143 := [rewrite]: #25142
+#25139 := (iff #25137 #25138)
+#25140 := [monotonicity #24103 #24112]: #25139
+#25145 := [trans #25140 #25143]: #25144
+#25156 := [monotonicity #25145]: #25155
+#25161 := [trans #25156 #25158]: #25160
+#25153 := [quant-inst #4649 #4655 #356 #4646 #1138]: #25147
+#25162 := [mp #25153 #25161]: #25113
+#25176 := [unit-resolution #25162 #19597 #13463 #25186 #25190]: false
+#25177 := [lemma #25176]: #25136
+#24735 := (or #25135 #25125)
+#24722 := [def-axiom]: #24735
+#25319 := [unit-resolution #24722 #25177]: #25125
+#25294 := (= #24179 #25124)
+#25292 := (= #24175 #25123)
+#25290 := (= #23963 #25123)
+#25288 := (= #25123 #23963)
+#25284 := (= #25122 #4736)
+#25276 := (= #24041 #4736)
+#25277 := [symm #24456]: #25276
+#25282 := (= #25122 #24041)
+#25274 := (= #4656 #24041)
+#25275 := [monotonicity #24982]: #25274
+#25280 := (= #25122 #4656)
+#25272 := (= #25219 #4656)
+#25256 := (= #24890 f445)
+#25254 := (= #24890 #4657)
+#23488 := (= #22490 #4657)
+#23497 := (or #23460 #23488)
+#23498 := [quant-inst #4652 #4657]: #23497
+#25251 := [unit-resolution #23498 #21829]: #23488
+#25252 := (= #24890 #22490)
+#25253 := [monotonicity #23932]: #25252
+#25255 := [trans #25253 #25251]: #25254
+#25257 := [trans #25255 #24980]: #25256
+#25273 := [monotonicity #25257]: #25272
+#25278 := (= #25122 #25219)
+#25320 := [hypothesis]: #25222
+#25270 := (= #25122 #25201)
+#25268 := (= #25121 #25194)
+#25266 := (= #25194 #25121)
+#25264 := (= #25193 #25120)
+#25262 := (= #23413 #24084)
+#25260 := (= #4658 #24084)
+#25261 := [symm #24984]: #25260
+#25263 := [trans #23932 #25261]: #25262
+#25265 := [monotonicity #25263]: #25264
+#25267 := [monotonicity #25265]: #25266
+#25269 := [symm #25267]: #25268
+#25271 := [monotonicity #25269]: #25270
+#25321 := [trans #25271 #25320]: #25278
+#25322 := [trans #25321 #25273]: #25280
+#25323 := [trans #25322 #25275]: #25282
+#25324 := [trans #25323 #25277]: #25284
+#25325 := [monotonicity #25324]: #25288
+#25326 := [symm #25325]: #25290
+#25327 := [trans #25287 #25326]: #25292
+#25328 := [monotonicity #25327]: #25294
+#25329 := [trans #25328 #25319]: #25296
+#25330 := [trans #25329 #24984]: #25298
+#25332 := [monotonicity #25330]: #25331
+#25333 := [trans #25332 #13466]: #24189
+#24789 := (not #24189)
+#25318 := [hypothesis]: #24789
+#25334 := [unit-resolution #25318 #25333]: false
+#25336 := [lemma #25334]: #25335
+#25114 := [unit-resolution #25336 #25109]: #24189
+#24940 := (or #24193 #24789)
+#24941 := [def-axiom]: #24940
+#25116 := [unit-resolution #24941 #25114]: #24193
+#24945 := (not #24187)
+#24307 := (f235 f236 #23963)
+#24308 := (f62 f63 #24307)
+#24311 := (f61 #24308)
+#24312 := (= #24311 f3)
+#25019 := [hypothesis]: #24187
+#25005 := (= #24311 #24186)
+#24955 := (= #24308 #24185)
+#24834 := (= #24307 #24179)
+#24835 := (= #23963 #24175)
+#25001 := [symm #25287]: #24835
+#24954 := [monotonicity #25001]: #24834
+#24972 := [monotonicity #24954]: #24955
+#25006 := [monotonicity #24972]: #25005
+#25023 := [trans #25006 #25019]: #24312
+#24939 := (not #24312)
+#24313 := (f45 f79 #24308)
+#24314 := (= #24313 f1)
+#24315 := (not #24314)
+#24316 := (or #24312 #24315)
+#24317 := (not #24316)
+#4275 := (:var 0 S56)
+#4276 := (f235 f236 #4275)
+#4277 := (pattern #4276)
+#4278 := (f62 f63 #4276)
+#4282 := (f45 f79 #4278)
+#4283 := (= #4282 f1)
+#20836 := (not #4283)
+#4279 := (f61 #4278)
+#4280 := (= #4279 f3)
+#20837 := (or #4280 #20836)
+#20838 := (not #20837)
+#20841 := (forall (vars (?v0 S56)) (:pat #4277) #20838)
+#4281 := (not #4280)
+#4284 := (and #4281 #4283)
+#4285 := (forall (vars (?v0 S56)) (:pat #4277) #4284)
+#20842 := (iff #4285 #20841)
+#20839 := (iff #4284 #20838)
+#20840 := [rewrite]: #20839
+#20843 := [quant-intro #20840]: #20842
+#16870 := (~ #4285 #4285)
+#16868 := (~ #4284 #4284)
+#16869 := [refl]: #16868
+#16871 := [nnf-pos #16869]: #16870
+#11056 := [asserted]: #4285
+#16872 := [mp~ #11056 #16871]: #4285
+#20844 := [mp #16872 #20843]: #20841
+#24538 := (not #20841)
+#24950 := (or #24538 #24317)
+#24938 := [quant-inst #23963]: #24950
+#24787 := [unit-resolution #24938 #20844]: #24317
+#25002 := (or #24316 #24939)
+#25003 := [def-axiom]: #25002
+#24788 := [unit-resolution #25003 #24787]: #24939
+#25024 := [unit-resolution #24788 #25023]: false
+#25020 := [lemma #25024]: #24945
+#24628 := (or #24196 #24174 #24184 #24187 #24194)
+#24644 := [def-axiom]: #24628
+#25071 := [unit-resolution #24644 #25020 #25116]: #25117
+#25069 := [unit-resolution #25071 #25108 #25048]: #24196
+#24082 := (or #24205 #24195)
+#24083 := [def-axiom]: #24082
+#24730 := [unit-resolution #24083 #25069 #25052]: false
+#24749 := [lemma #24730]: #4745
+#25458 := (or #17067 #22090)
+#22093 := (or #17058 #17067 #22090)
+#22096 := (not #22093)
+#22099 := (or #17058 #17061 #22096)
+#22102 := (not #22099)
+#22105 := (or #17058 #17061 #22102)
+#22108 := (not #22105)
+#22111 := (or #12412 #22108)
+#22114 := (not #22111)
+#22117 := (or #12412 #22114)
+#21170 := (forall (vars (?v0 Int)) #21159)
+#21177 := (not #21170)
+#21155 := (forall (vars (?v0 Int)) #21150)
+#21176 := (not #21155)
+#21178 := (or #21176 #21177)
+#21179 := (not #21178)
+#21184 := (or #21144 #21179)
+#21190 := (not #21184)
+#21191 := (or #12923 #12777 #21067 #11388 #11379 #11370 #11361 #21190)
+#21192 := (not #21191)
+#20982 := (forall (vars (?v0 Int)) #20977)
+#21000 := (not #20982)
+#21001 := (or #21000 #20987)
+#21002 := (not #21001)
+#21007 := (or #20971 #21002)
+#21013 := (not #21007)
+#21014 := (or #12681 #21013)
+#21015 := (not #21014)
+#21020 := (or #12681 #21015)
+#21028 := (not #21020)
+#21029 := (or #17209 #17212 #12743 #11705 #21026 #21027 #21028)
+#21030 := (not #21029)
+#21035 := (or #17209 #17212 #21030)
+#21041 := (not #21035)
+#21078 := (or #21067 #12828 #11892 #11883 #12777 #21027 #21041)
+#21079 := (not #21078)
+#21042 := (or #17171 #17180 #11803 #11794 #11785 #11760 #11751 #12777 #21027 #21041)
+#21043 := (not #21042)
+#21048 := (or #17171 #17180 #21043)
+#21054 := (not #21048)
+#21055 := (or #17171 #17174 #21054)
+#21056 := (not #21055)
+#21061 := (or #17171 #17174 #21056)
+#21068 := (not #21061)
+#21069 := (or #12777 #21067 #12829 #21068)
+#21070 := (not #21069)
+#21084 := (or #21070 #21079)
+#21090 := (not #21084)
+#21091 := (or #17171 #17180 #12777 #21067 #21090)
+#21092 := (not #21091)
+#21097 := (or #17171 #17180 #21092)
+#21103 := (not #21097)
+#21104 := (or #17171 #17174 #21103)
+#21105 := (not #21104)
+#21110 := (or #17171 #17174 #21105)
+#21116 := (not #21110)
+#21117 := (or #12777 #21067 #12922 #21116)
+#21118 := (not #21117)
+#21197 := (or #21118 #21192)
+#21212 := (not #21197)
+#20947 := (forall (vars (?v0 Int)) #20942)
+#21208 := (not #20947)
+#21213 := (or #12634 #17117 #21203 #21204 #21205 #21206 #21207 #12777 #21067 #13142 #21208 #13095 #21209 #13090 #15031 #12144 #12135 #12126 #12117 #21210 #21211 #21212)
+#21214 := (not #21213)
+#21219 := (or #12634 #17117 #21214)
+#21226 := (not #21219)
+#20936 := (forall (vars (?v0 Int)) #20931)
+#21225 := (not #20936)
+#21227 := (or #21225 #21226)
+#21228 := (not #21227)
+#21233 := (or #20925 #21228)
+#21239 := (not #21233)
+#21240 := (or #12598 #21239)
+#21241 := (not #21240)
+#21246 := (or #12598 #21241)
+#21252 := (not #21246)
+#21253 := (or #17058 #17067 #12379 #12370 #12361 #12352 #21252)
+#21254 := (not #21253)
+#21259 := (or #17058 #17067 #21254)
+#21265 := (not #21259)
+#21266 := (or #17058 #17061 #21265)
+#21267 := (not #21266)
+#21272 := (or #17058 #17061 #21267)
+#21278 := (not #21272)
+#21279 := (or #12412 #21278)
+#21280 := (not #21279)
+#21285 := (or #12412 #21280)
+#22118 := (iff #21285 #22117)
+#22115 := (iff #21280 #22114)
+#22112 := (iff #21279 #22111)
+#22109 := (iff #21278 #22108)
+#22106 := (iff #21272 #22105)
+#22103 := (iff #21267 #22102)
+#22100 := (iff #21266 #22099)
+#22097 := (iff #21265 #22096)
+#22094 := (iff #21259 #22093)
+#22091 := (iff #21254 #22090)
+#22088 := (iff #21253 #22087)
+#22085 := (iff #21252 #22084)
+#22082 := (iff #21246 #22081)
+#22079 := (iff #21241 #22078)
+#22076 := (iff #21240 #22075)
+#22073 := (iff #21239 #22072)
+#22070 := (iff #21233 #22069)
+#22067 := (iff #21228 #22066)
+#22064 := (iff #21227 #22063)
+#22061 := (iff #21226 #22060)
+#22058 := (iff #21219 #22057)
+#22055 := (iff #21214 #22054)
+#22052 := (iff #21213 #22051)
+#22049 := (iff #21212 #22048)
+#22046 := (iff #21197 #22045)
+#22043 := (iff #21192 #22042)
+#22040 := (iff #21191 #22039)
+#22037 := (iff #21190 #22036)
+#22034 := (iff #21184 #22033)
+#22031 := (iff #21179 #22030)
+#22028 := (iff #21178 #22027)
+#22025 := (iff #21177 #22024)
+#22022 := (iff #21170 #22019)
+#22020 := (iff #21159 #21159)
+#22021 := [refl]: #22020
+#22023 := [quant-intro #22021]: #22022
+#22026 := [monotonicity #22023]: #22025
+#22017 := (iff #21176 #22016)
+#22014 := (iff #21155 #22011)
+#22012 := (iff #21150 #21150)
+#22013 := [refl]: #22012
+#22015 := [quant-intro #22013]: #22014
+#22018 := [monotonicity #22015]: #22017
+#22029 := [monotonicity #22018 #22026]: #22028
+#22032 := [monotonicity #22029]: #22031
+#22035 := [monotonicity #22032]: #22034
+#22038 := [monotonicity #22035]: #22037
+#22041 := [monotonicity #22038]: #22040
+#22044 := [monotonicity #22041]: #22043
+#22009 := (iff #21118 #22008)
+#22006 := (iff #21117 #22005)
+#22003 := (iff #21116 #22002)
+#22000 := (iff #21110 #21999)
+#21997 := (iff #21105 #21996)
+#21994 := (iff #21104 #21993)
+#21991 := (iff #21103 #21990)
+#21988 := (iff #21097 #21987)
+#21985 := (iff #21092 #21984)
+#21982 := (iff #21091 #21981)
+#21979 := (iff #21090 #21978)
+#21976 := (iff #21084 #21975)
+#21973 := (iff #21079 #21972)
+#21970 := (iff #21078 #21969)
+#21937 := (iff #21041 #21936)
+#21934 := (iff #21035 #21933)
+#21931 := (iff #21030 #21930)
+#21928 := (iff #21029 #21927)
+#21925 := (iff #21028 #21924)
+#21922 := (iff #21020 #21921)
+#21919 := (iff #21015 #21918)
+#21916 := (iff #21014 #21915)
+#21913 := (iff #21013 #21912)
+#21910 := (iff #21007 #21909)
+#21907 := (iff #21002 #21906)
+#21904 := (iff #21001 #21903)
+#21901 := (iff #21000 #21900)
+#21898 := (iff #20982 #21895)
+#21896 := (iff #20977 #20977)
+#21897 := [refl]: #21896
+#21899 := [quant-intro #21897]: #21898
+#21902 := [monotonicity #21899]: #21901
+#21905 := [monotonicity #21902]: #21904
+#21908 := [monotonicity #21905]: #21907
+#21911 := [monotonicity #21908]: #21910
+#21914 := [monotonicity #21911]: #21913
+#21917 := [monotonicity #21914]: #21916
+#21920 := [monotonicity #21917]: #21919
+#21923 := [monotonicity #21920]: #21922
+#21926 := [monotonicity #21923]: #21925
+#21929 := [monotonicity #21926]: #21928
+#21932 := [monotonicity #21929]: #21931
+#21935 := [monotonicity #21932]: #21934
+#21938 := [monotonicity #21935]: #21937
+#21971 := [monotonicity #21938]: #21970
+#21974 := [monotonicity #21971]: #21973
+#21967 := (iff #21070 #21966)
+#21964 := (iff #21069 #21963)
+#21961 := (iff #21068 #21960)
+#21958 := (iff #21061 #21957)
+#21955 := (iff #21056 #21954)
+#21952 := (iff #21055 #21951)
+#21949 := (iff #21054 #21948)
+#21946 := (iff #21048 #21945)
+#21943 := (iff #21043 #21942)
+#21940 := (iff #21042 #21939)
+#21941 := [monotonicity #21938]: #21940
+#21944 := [monotonicity #21941]: #21943
+#21947 := [monotonicity #21944]: #21946
+#21950 := [monotonicity #21947]: #21949
+#21953 := [monotonicity #21950]: #21952
+#21956 := [monotonicity #21953]: #21955
+#21959 := [monotonicity #21956]: #21958
+#21962 := [monotonicity #21959]: #21961
+#21965 := [monotonicity #21962]: #21964
+#21968 := [monotonicity #21965]: #21967
+#21977 := [monotonicity #21968 #21974]: #21976
+#21980 := [monotonicity #21977]: #21979
+#21983 := [monotonicity #21980]: #21982
+#21986 := [monotonicity #21983]: #21985
+#21989 := [monotonicity #21986]: #21988
+#21992 := [monotonicity #21989]: #21991
+#21995 := [monotonicity #21992]: #21994
+#21998 := [monotonicity #21995]: #21997
+#22001 := [monotonicity #21998]: #22000
+#22004 := [monotonicity #22001]: #22003
+#22007 := [monotonicity #22004]: #22006
+#22010 := [monotonicity #22007]: #22009
+#22047 := [monotonicity #22010 #22044]: #22046
+#22050 := [monotonicity #22047]: #22049
+#21893 := (iff #21208 #21892)
+#21890 := (iff #20947 #21887)
+#21888 := (iff #20942 #20942)
+#21889 := [refl]: #21888
+#21891 := [quant-intro #21889]: #21890
+#21894 := [monotonicity #21891]: #21893
+#22053 := [monotonicity #21894 #22050]: #22052
+#22056 := [monotonicity #22053]: #22055
+#22059 := [monotonicity #22056]: #22058
+#22062 := [monotonicity #22059]: #22061
+#21885 := (iff #21225 #21884)
+#21882 := (iff #20936 #21879)
+#21880 := (iff #20931 #20931)
+#21881 := [refl]: #21880
+#21883 := [quant-intro #21881]: #21882
+#21886 := [monotonicity #21883]: #21885
+#22065 := [monotonicity #21886 #22062]: #22064
+#22068 := [monotonicity #22065]: #22067
+#22071 := [monotonicity #22068]: #22070
+#22074 := [monotonicity #22071]: #22073
+#22077 := [monotonicity #22074]: #22076
+#22080 := [monotonicity #22077]: #22079
+#22083 := [monotonicity #22080]: #22082
+#22086 := [monotonicity #22083]: #22085
+#22089 := [monotonicity #22086]: #22088
+#22092 := [monotonicity #22089]: #22091
+#22095 := [monotonicity #22092]: #22094
+#22098 := [monotonicity #22095]: #22097
+#22101 := [monotonicity #22098]: #22100
+#22104 := [monotonicity #22101]: #22103
+#22107 := [monotonicity #22104]: #22106
+#22110 := [monotonicity #22107]: #22109
+#22113 := [monotonicity #22110]: #22112
+#22116 := [monotonicity #22113]: #22115
+#22119 := [monotonicity #22116]: #22118
+#13642 := (and #6706 #14917 #12952 #4826)
+#17379 := (not #13642)
+#17382 := (forall (vars (?v0 Int)) #17379)
+#14273 := (and #6706 #14917)
+#14268 := (not #14273)
+#13653 := (or #14268 #12950 #12964)
+#13648 := (forall (vars (?v0 Int)) #13653)
+#17386 := (and #13648 #17382)
+#17357 := (and #17355 #17356)
+#17358 := (not #17357)
+#17685 := (or #17358 #17667 #17680)
+#17688 := (not #17685)
+#17691 := (or #17688 #17386)
+#17697 := (and #12922 #12639 #12642 #4814 #4816 #4818 #4820 #17691)
+#17257 := (not #12724)
+#13714 := (or #14268 #12684 #12698)
+#13713 := (forall (vars (?v0 Int)) #13714)
+#17260 := (and #13713 #17257)
+#17235 := (and #17233 #17234)
+#17236 := (not #17235)
+#17550 := (or #17236 #17532 #17545)
+#17553 := (not #17550)
+#17556 := (or #17553 #17260)
+#17559 := (and #12676 #17556)
+#17562 := (or #12681 #17559)
+#17568 := (and #12660 #13727 #12739 #4978 #12668 #12651 #17562)
+#17573 := (or #17209 #17212 #17568)
+#17613 := (and #12642 #12829 #5026 #5027 #12639 #12651 #17573)
+#17579 := (and #4940 #4945 #4950 #4955 #4960 #4963 #4965 #12639 #12651 #17573)
+#17584 := (or #17171 #17180 #17579)
+#17590 := (and #4940 #4942 #17584)
+#17595 := (or #17171 #17174 #17590)
+#17601 := (and #12639 #12642 #12828 #17595)
+#17618 := (or #17601 #17613)
+#17624 := (and #4940 #4945 #12639 #12642 #17618)
+#17629 := (or #17171 #17180 #17624)
+#17635 := (and #4940 #4942 #17629)
+#17640 := (or #17171 #17174 #17635)
+#17646 := (and #12639 #12642 #12923 #17640)
+#17702 := (or #17646 #17697)
+#13738 := (or #14268 #13105 #13119)
+#13737 := (forall (vars (?v0 Int)) #13738)
+#17708 := (and #12635 #4780 #13167 #13785 #13766 #13145 #13751 #12639 #12642 #13139 #13737 #13096 #4806 #4898 #4811 #4909 #4913 #4917 #4921 #4925 #4930 #17702)
+#17713 := (or #12634 #17117 #17708)
+#13798 := (or #14268 #12601 #12613)
+#13797 := (forall (vars (?v0 Int)) #13798)
+#17716 := (and #13797 #17713)
+#17093 := (and #17091 #17092)
+#17094 := (not #17093)
+#17102 := (or #17094 #17095 #17101)
+#17103 := (not #17102)
+#17719 := (or #17103 #17716)
+#17722 := (and #12595 #17719)
+#17725 := (or #12598 #17722)
+#17731 := (and #4739 #4745 #4750 #4755 #4760 #4765 #17725)
+#17736 := (or #17058 #17067 #17731)
+#17742 := (and #4739 #4741 #17736)
+#17747 := (or #17058 #17061 #17742)
+#17750 := (and #4733 #17747)
+#17753 := (or #12412 #17750)
+#21286 := (iff #17753 #21285)
+#21283 := (iff #17750 #21280)
+#21275 := (and #4733 #21272)
+#21281 := (iff #21275 #21280)
+#21282 := [rewrite]: #21281
+#21276 := (iff #17750 #21275)
+#21273 := (iff #17747 #21272)
+#21270 := (iff #17742 #21267)
+#21262 := (and #4739 #4741 #21259)
+#21268 := (iff #21262 #21267)
+#21269 := [rewrite]: #21268
+#21263 := (iff #17742 #21262)
+#21260 := (iff #17736 #21259)
+#21257 := (iff #17731 #21254)
+#21249 := (and #4739 #4745 #4750 #4755 #4760 #4765 #21246)
+#21255 := (iff #21249 #21254)
+#21256 := [rewrite]: #21255
+#21250 := (iff #17731 #21249)
+#21247 := (iff #17725 #21246)
+#21244 := (iff #17722 #21241)
+#21236 := (and #12595 #21233)
+#21242 := (iff #21236 #21241)
+#21243 := [rewrite]: #21242
+#21237 := (iff #17722 #21236)
+#21234 := (iff #17719 #21233)
+#21231 := (iff #17716 #21228)
+#21222 := (and #20936 #21219)
+#21229 := (iff #21222 #21228)
+#21230 := [rewrite]: #21229
+#21223 := (iff #17716 #21222)
+#21220 := (iff #17713 #21219)
+#21217 := (iff #17708 #21214)
+#21200 := (and #12635 #4780 #13167 #13785 #13766 #13145 #13751 #12639 #12642 #13139 #20947 #13096 #4806 #4898 #4811 #4909 #4913 #4917 #4921 #4925 #4930 #21197)
+#21215 := (iff #21200 #21214)
+#21216 := [rewrite]: #21215
+#21201 := (iff #17708 #21200)
+#21198 := (iff #17702 #21197)
+#21195 := (iff #17697 #21192)
+#21187 := (and #12922 #12639 #12642 #4814 #4816 #4818 #4820 #21184)
+#21193 := (iff #21187 #21192)
+#21194 := [rewrite]: #21193
+#21188 := (iff #17697 #21187)
+#21185 := (iff #17691 #21184)
+#21182 := (iff #17386 #21179)
+#21173 := (and #21155 #21170)
+#21180 := (iff #21173 #21179)
+#21181 := [rewrite]: #21180
+#21174 := (iff #17386 #21173)
+#21171 := (iff #17382 #21170)
+#21168 := (iff #17379 #21159)
+#21160 := (not #21159)
+#21163 := (not #21160)
+#21166 := (iff #21163 #21159)
+#21167 := [rewrite]: #21166
+#21164 := (iff #17379 #21163)
+#21161 := (iff #13642 #21160)
+#21162 := [rewrite]: #21161
+#21165 := [monotonicity #21162]: #21164
+#21169 := [trans #21165 #21167]: #21168
+#21172 := [quant-intro #21169]: #21171
+#21156 := (iff #13648 #21155)
+#21153 := (iff #13653 #21150)
+#18824 := (or #7428 #18181)
+#21147 := (or #18824 #12950 #12964)
+#21151 := (iff #21147 #21150)
+#21152 := [rewrite]: #21151
+#21148 := (iff #13653 #21147)
+#18833 := (iff #14268 #18824)
+#18825 := (not #18824)
+#18828 := (not #18825)
+#18831 := (iff #18828 #18824)
+#18832 := [rewrite]: #18831
+#18829 := (iff #14268 #18828)
+#18826 := (iff #14273 #18825)
+#18827 := [rewrite]: #18826
+#18830 := [monotonicity #18827]: #18829
+#18834 := [trans #18830 #18832]: #18833
+#21149 := [monotonicity #18834]: #21148
+#21154 := [trans #21149 #21152]: #21153
+#21157 := [quant-intro #21154]: #21156
+#21175 := [monotonicity #21157 #21172]: #21174
+#21183 := [trans #21175 #21181]: #21182
+#21145 := (iff #17688 #21144)
+#21142 := (iff #17685 #21139)
+#21125 := (or #21123 #21124)
+#21136 := (or #21125 #17667 #17680)
+#21140 := (iff #21136 #21139)
+#21141 := [rewrite]: #21140
+#21137 := (iff #17685 #21136)
+#21134 := (iff #17358 #21125)
+#21126 := (not #21125)
+#21129 := (not #21126)
+#21132 := (iff #21129 #21125)
+#21133 := [rewrite]: #21132
+#21130 := (iff #17358 #21129)
+#21127 := (iff #17357 #21126)
+#21128 := [rewrite]: #21127
+#21131 := [monotonicity #21128]: #21130
+#21135 := [trans #21131 #21133]: #21134
+#21138 := [monotonicity #21135]: #21137
+#21143 := [trans #21138 #21141]: #21142
+#21146 := [monotonicity #21143]: #21145
+#21186 := [monotonicity #21146 #21183]: #21185
+#21189 := [monotonicity #21186]: #21188
+#21196 := [trans #21189 #21194]: #21195
+#21121 := (iff #17646 #21118)
+#21113 := (and #12639 #12642 #12923 #21110)
+#21119 := (iff #21113 #21118)
+#21120 := [rewrite]: #21119
+#21114 := (iff #17646 #21113)
+#21111 := (iff #17640 #21110)
+#21108 := (iff #17635 #21105)
+#21100 := (and #4940 #4942 #21097)
+#21106 := (iff #21100 #21105)
+#21107 := [rewrite]: #21106
+#21101 := (iff #17635 #21100)
+#21098 := (iff #17629 #21097)
+#21095 := (iff #17624 #21092)
+#21087 := (and #4940 #4945 #12639 #12642 #21084)
+#21093 := (iff #21087 #21092)
+#21094 := [rewrite]: #21093
+#21088 := (iff #17624 #21087)
+#21085 := (iff #17618 #21084)
+#21082 := (iff #17613 #21079)
+#21075 := (and #12642 #12829 #5026 #5027 #12639 #12651 #21035)
+#21080 := (iff #21075 #21079)
+#21081 := [rewrite]: #21080
+#21076 := (iff #17613 #21075)
+#21036 := (iff #17573 #21035)
+#21033 := (iff #17568 #21030)
+#21023 := (and #12660 #13727 #12739 #4978 #12668 #12651 #21020)
+#21031 := (iff #21023 #21030)
+#21032 := [rewrite]: #21031
+#21024 := (iff #17568 #21023)
+#21021 := (iff #17562 #21020)
+#21018 := (iff #17559 #21015)
+#21010 := (and #12676 #21007)
+#21016 := (iff #21010 #21015)
+#21017 := [rewrite]: #21016
+#21011 := (iff #17559 #21010)
+#21008 := (iff #17556 #21007)
+#21005 := (iff #17260 #21002)
+#20997 := (and #20982 #20986)
+#21003 := (iff #20997 #21002)
+#21004 := [rewrite]: #21003
+#20998 := (iff #17260 #20997)
+#20995 := (iff #17257 #20986)
+#20990 := (not #20987)
+#20993 := (iff #20990 #20986)
+#20994 := [rewrite]: #20993
+#20991 := (iff #17257 #20990)
+#20988 := (iff #12724 #20987)
+#20989 := [rewrite]: #20988
+#20992 := [monotonicity #20989]: #20991
+#20996 := [trans #20992 #20994]: #20995
+#20983 := (iff #13713 #20982)
+#20980 := (iff #13714 #20977)
+#20974 := (or #18824 #12684 #12698)
+#20978 := (iff #20974 #20977)
+#20979 := [rewrite]: #20978
+#20975 := (iff #13714 #20974)
+#20976 := [monotonicity #18834]: #20975
+#20981 := [trans #20976 #20979]: #20980
+#20984 := [quant-intro #20981]: #20983
+#20999 := [monotonicity #20984 #20996]: #20998
+#21006 := [trans #20999 #21004]: #21005
+#20972 := (iff #17553 #20971)
+#20969 := (iff #17550 #20966)
+#20952 := (or #20950 #20951)
+#20963 := (or #20952 #17532 #17545)
+#20967 := (iff #20963 #20966)
+#20968 := [rewrite]: #20967
+#20964 := (iff #17550 #20963)
+#20961 := (iff #17236 #20952)
+#20953 := (not #20952)
+#20956 := (not #20953)
+#20959 := (iff #20956 #20952)
+#20960 := [rewrite]: #20959
+#20957 := (iff #17236 #20956)
+#20954 := (iff #17235 #20953)
+#20955 := [rewrite]: #20954
+#20958 := [monotonicity #20955]: #20957
+#20962 := [trans #20958 #20960]: #20961
+#20965 := [monotonicity #20962]: #20964
+#20970 := [trans #20965 #20968]: #20969
+#20973 := [monotonicity #20970]: #20972
+#21009 := [monotonicity #20973 #21006]: #21008
+#21012 := [monotonicity #21009]: #21011
+#21019 := [trans #21012 #21017]: #21018
+#21022 := [monotonicity #21019]: #21021
+#21025 := [monotonicity #21022]: #21024
+#21034 := [trans #21025 #21032]: #21033
+#21037 := [monotonicity #21034]: #21036
+#21077 := [monotonicity #21037]: #21076
+#21083 := [trans #21077 #21081]: #21082
+#21073 := (iff #17601 #21070)
+#21064 := (and #12639 #12642 #12828 #21061)
+#21071 := (iff #21064 #21070)
+#21072 := [rewrite]: #21071
+#21065 := (iff #17601 #21064)
+#21062 := (iff #17595 #21061)
+#21059 := (iff #17590 #21056)
+#21051 := (and #4940 #4942 #21048)
+#21057 := (iff #21051 #21056)
+#21058 := [rewrite]: #21057
+#21052 := (iff #17590 #21051)
+#21049 := (iff #17584 #21048)
+#21046 := (iff #17579 #21043)
+#21038 := (and #4940 #4945 #4950 #4955 #4960 #4963 #4965 #12639 #12651 #21035)
+#21044 := (iff #21038 #21043)
+#21045 := [rewrite]: #21044
+#21039 := (iff #17579 #21038)
+#21040 := [monotonicity #21037]: #21039
+#21047 := [trans #21040 #21045]: #21046
+#21050 := [monotonicity #21047]: #21049
+#21053 := [monotonicity #21050]: #21052
+#21060 := [trans #21053 #21058]: #21059
+#21063 := [monotonicity #21060]: #21062
+#21066 := [monotonicity #21063]: #21065
+#21074 := [trans #21066 #21072]: #21073
+#21086 := [monotonicity #21074 #21083]: #21085
+#21089 := [monotonicity #21086]: #21088
+#21096 := [trans #21089 #21094]: #21095
+#21099 := [monotonicity #21096]: #21098
+#21102 := [monotonicity #21099]: #21101
+#21109 := [trans #21102 #21107]: #21108
+#21112 := [monotonicity #21109]: #21111
+#21115 := [monotonicity #21112]: #21114
+#21122 := [trans #21115 #21120]: #21121
+#21199 := [monotonicity #21122 #21196]: #21198
+#20948 := (iff #13737 #20947)
+#20945 := (iff #13738 #20942)
+#20939 := (or #18824 #13105 #13119)
+#20943 := (iff #20939 #20942)
+#20944 := [rewrite]: #20943
+#20940 := (iff #13738 #20939)
+#20941 := [monotonicity #18834]: #20940
+#20946 := [trans #20941 #20944]: #20945
+#20949 := [quant-intro #20946]: #20948
+#21202 := [monotonicity #20949 #21199]: #21201
+#21218 := [trans #21202 #21216]: #21217
+#21221 := [monotonicity #21218]: #21220
+#20937 := (iff #13797 #20936)
+#20934 := (iff #13798 #20931)
+#20928 := (or #18824 #12601 #12613)
+#20932 := (iff #20928 #20931)
+#20933 := [rewrite]: #20932
+#20929 := (iff #13798 #20928)
+#20930 := [monotonicity #18834]: #20929
+#20935 := [trans #20930 #20933]: #20934
+#20938 := [quant-intro #20935]: #20937
+#21224 := [monotonicity #20938 #21221]: #21223
+#21232 := [trans #21224 #21230]: #21231
+#20926 := (iff #17103 #20925)
+#20923 := (iff #17102 #20920)
+#20906 := (or #20904 #20905)
+#20917 := (or #20906 #17095 #17101)
+#20921 := (iff #20917 #20920)
+#20922 := [rewrite]: #20921
+#20918 := (iff #17102 #20917)
+#20915 := (iff #17094 #20906)
+#20907 := (not #20906)
+#20910 := (not #20907)
+#20913 := (iff #20910 #20906)
+#20914 := [rewrite]: #20913
+#20911 := (iff #17094 #20910)
+#20908 := (iff #17093 #20907)
+#20909 := [rewrite]: #20908
+#20912 := [monotonicity #20909]: #20911
+#20916 := [trans #20912 #20914]: #20915
+#20919 := [monotonicity #20916]: #20918
+#20924 := [trans #20919 #20922]: #20923
+#20927 := [monotonicity #20924]: #20926
+#21235 := [monotonicity #20927 #21232]: #21234
+#21238 := [monotonicity #21235]: #21237
+#21245 := [trans #21238 #21243]: #21244
+#21248 := [monotonicity #21245]: #21247
+#21251 := [monotonicity #21248]: #21250
+#21258 := [trans #21251 #21256]: #21257
+#21261 := [monotonicity #21258]: #21260
+#21264 := [monotonicity #21261]: #21263
+#21271 := [trans #21264 #21269]: #21270
+#21274 := [monotonicity #21271]: #21273
+#21277 := [monotonicity #21274]: #21276
+#21284 := [trans #21277 #21282]: #21283
+#21287 := [monotonicity #21284]: #21286
+#17364 := (+ #17363 #12962)
+#17365 := (<= #17364 0::Int)
+#17359 := (+ ?v0!15 #12568)
+#17360 := (>= #17359 0::Int)
+#17366 := (or #17358 #17360 #17365)
+#17367 := (not #17366)
+#17390 := (or #17367 #17386)
+#17351 := (not #11361)
+#17348 := (not #11370)
+#17345 := (not #11379)
+#17342 := (not #11388)
+#17132 := (not #12647)
+#17394 := (and #12926 #17132 #17342 #17345 #17348 #17351 #17390)
+#17242 := (+ #17241 #12696)
+#17243 := (<= #17242 0::Int)
+#17237 := (+ ?v0!14 #12677)
+#17238 := (>= #17237 0::Int)
+#17244 := (or #17236 #17238 #17243)
+#17245 := (not #17244)
+#17264 := (or #17245 #17260)
+#17229 := (not #12681)
+#17268 := (and #17229 #17264)
+#17272 := (or #12681 #17268)
+#17224 := (not #12673)
+#17221 := (not #11705)
+#17218 := (not #12743)
+#13720 := (and #12660 #13727)
+#13719 := (not #13720)
+#17215 := (not #13719)
+#17276 := (and #17215 #17218 #17221 #17224 #17272)
+#17280 := (or #17209 #17212 #17276)
+#17206 := (not #12656)
+#17309 := (not #11883)
+#17306 := (not #11892)
+#17312 := (and #17132 #12834 #17306 #17309 #17206 #17280)
+#17203 := (not #11751)
+#17200 := (not #11760)
+#17197 := (not #12777)
+#17194 := (not #11785)
+#17191 := (not #11794)
+#17188 := (not #11803)
+#17183 := (not #11812)
+#17284 := (and #17183 #17188 #17191 #17194 #17197 #17200 #17203 #17206 #17280)
+#17288 := (or #17171 #17180 #17284)
+#17177 := (not #11824)
+#17292 := (and #17177 #17288)
+#17296 := (or #17171 #17174 #17292)
+#17300 := (and #17132 #12828 #17296)
+#17316 := (or #17300 #17312)
+#17320 := (and #17183 #17132 #17316)
+#17324 := (or #17171 #17180 #17320)
+#17328 := (and #17177 #17324)
+#17332 := (or #17171 #17174 #17328)
+#17336 := (and #17132 #12923 #17332)
+#17398 := (or #17336 #17394)
+#17166 := (not #12108)
+#17163 := (not #12117)
+#17160 := (not #12126)
+#17157 := (not #12135)
+#17154 := (not #12144)
+#17151 := (not #15031)
+#17148 := (not #13090)
+#17145 := (not #13102)
+#17135 := (not #13142)
+#13744 := (and #13145 #13751)
+#13743 := (not #13744)
+#17129 := (not #13743)
+#13763 := (and #12642 #13766)
+#13758 := (not #13763)
+#17126 := (not #13758)
+#13778 := (and #13167 #13785)
+#13777 := (not #13778)
+#17123 := (not #13777)
+#17120 := (not #13183)
+#17402 := (and #17120 #17123 #17126 #17129 #17132 #17135 #13737 #17145 #17148 #17151 #17154 #17157 #17160 #17163 #17166 #17398)
+#17406 := (or #13340 #17117 #17402)
+#17410 := (and #13797 #17406)
+#17414 := (or #17103 #17410)
+#17087 := (not #12598)
+#17418 := (and #17087 #17414)
+#17422 := (or #12598 #17418)
+#17082 := (not #12352)
+#17079 := (not #12361)
+#17076 := (not #12370)
+#17073 := (not #12379)
+#17070 := (not #12388)
+#17426 := (and #17070 #17073 #17076 #17079 #17082 #17422)
+#17430 := (or #17058 #17067 #17426)
+#17064 := (not #12400)
+#17434 := (and #17064 #17430)
+#17438 := (or #17058 #17061 #17434)
+#17055 := (not #12412)
+#17442 := (and #17055 #17438)
+#17446 := (or #12412 #17442)
+#17754 := (iff #17446 #17753)
+#17751 := (iff #17442 #17750)
+#17748 := (iff #17438 #17747)
+#17745 := (iff #17434 #17742)
+#17739 := (and #4742 #17736)
+#17743 := (iff #17739 #17742)
+#17744 := [rewrite]: #17743
+#17740 := (iff #17434 #17739)
+#17737 := (iff #17430 #17736)
+#17734 := (iff #17426 #17731)
+#17728 := (and #4746 #4750 #4755 #4760 #4765 #17725)
+#17732 := (iff #17728 #17731)
+#17733 := [rewrite]: #17732
+#17729 := (iff #17426 #17728)
+#17726 := (iff #17422 #17725)
+#17723 := (iff #17418 #17722)
+#17720 := (iff #17414 #17719)
+#17717 := (iff #17410 #17716)
+#17714 := (iff #17406 #17713)
+#17711 := (iff #17402 #17708)
+#17705 := (and #13180 #13778 #13763 #13744 #12644 #13139 #13737 #13099 #4898 #4811 #4909 #4913 #4917 #4921 #4931 #17702)
+#17709 := (iff #17705 #17708)
+#17710 := [rewrite]: #17709
+#17706 := (iff #17402 #17705)
+#17703 := (iff #17398 #17702)
+#17700 := (iff #17394 #17697)
+#17694 := (and #12922 #12644 #4814 #4816 #4818 #4820 #17691)
+#17698 := (iff #17694 #17697)
+#17699 := [rewrite]: #17698
+#17695 := (iff #17394 #17694)
+#17692 := (iff #17390 #17691)
+#17689 := (iff #17367 #17688)
+#17686 := (iff #17366 #17685)
+#17683 := (iff #17365 #17680)
+#17672 := (+ #12962 #17363)
+#17675 := (<= #17672 0::Int)
+#17681 := (iff #17675 #17680)
+#17682 := [rewrite]: #17681
+#17676 := (iff #17365 #17675)
+#17673 := (= #17364 #17672)
+#17674 := [rewrite]: #17673
+#17677 := [monotonicity #17674]: #17676
+#17684 := [trans #17677 #17682]: #17683
+#17670 := (iff #17360 #17667)
+#17659 := (+ #12568 ?v0!15)
+#17662 := (>= #17659 0::Int)
+#17668 := (iff #17662 #17667)
+#17669 := [rewrite]: #17668
+#17663 := (iff #17360 #17662)
+#17660 := (= #17359 #17659)
+#17661 := [rewrite]: #17660
+#17664 := [monotonicity #17661]: #17663
+#17671 := [trans #17664 #17669]: #17670
+#17687 := [monotonicity #17671 #17684]: #17686
+#17690 := [monotonicity #17687]: #17689
+#17693 := [monotonicity #17690]: #17692
+#17657 := (iff #17351 #4820)
+#17658 := [rewrite]: #17657
+#17655 := (iff #17348 #4818)
+#17656 := [rewrite]: #17655
+#17653 := (iff #17345 #4816)
+#17654 := [rewrite]: #17653
+#17651 := (iff #17342 #4814)
+#17652 := [rewrite]: #17651
+#17476 := (iff #17132 #12644)
+#17477 := [rewrite]: #17476
+#17696 := [monotonicity #12930 #17477 #17652 #17654 #17656 #17658 #17693]: #17695
+#17701 := [trans #17696 #17699]: #17700
+#17649 := (iff #17336 #17646)
+#17643 := (and #12644 #12923 #17640)
+#17647 := (iff #17643 #17646)
+#17648 := [rewrite]: #17647
+#17644 := (iff #17336 #17643)
+#17641 := (iff #17332 #17640)
+#17638 := (iff #17328 #17635)
+#17632 := (and #4943 #17629)
+#17636 := (iff #17632 #17635)
+#17637 := [rewrite]: #17636
+#17633 := (iff #17328 #17632)
+#17630 := (iff #17324 #17629)
+#17627 := (iff #17320 #17624)
+#17621 := (and #4946 #12644 #17618)
+#17625 := (iff #17621 #17624)
+#17626 := [rewrite]: #17625
+#17622 := (iff #17320 #17621)
+#17619 := (iff #17316 #17618)
+#17616 := (iff #17312 #17613)
+#17610 := (and #12644 #12829 #5026 #5027 #12653 #17573)
+#17614 := (iff #17610 #17613)
+#17615 := [rewrite]: #17614
+#17611 := (iff #17312 #17610)
+#17574 := (iff #17280 #17573)
+#17571 := (iff #17276 #17568)
+#17565 := (and #13720 #12739 #4978 #12670 #17562)
+#17569 := (iff #17565 #17568)
+#17570 := [rewrite]: #17569
+#17566 := (iff #17276 #17565)
+#17563 := (iff #17272 #17562)
+#17560 := (iff #17268 #17559)
+#17557 := (iff #17264 #17556)
+#17554 := (iff #17245 #17553)
+#17551 := (iff #17244 #17550)
+#17548 := (iff #17243 #17545)
+#17537 := (+ #12696 #17241)
+#17540 := (<= #17537 0::Int)
+#17546 := (iff #17540 #17545)
+#17547 := [rewrite]: #17546
+#17541 := (iff #17243 #17540)
+#17538 := (= #17242 #17537)
+#17539 := [rewrite]: #17538
+#17542 := [monotonicity #17539]: #17541
+#17549 := [trans #17542 #17547]: #17548
+#17535 := (iff #17238 #17532)
+#17524 := (+ #12677 ?v0!14)
+#17527 := (>= #17524 0::Int)
+#17533 := (iff #17527 #17532)
+#17534 := [rewrite]: #17533
+#17528 := (iff #17238 #17527)
+#17525 := (= #17237 #17524)
+#17526 := [rewrite]: #17525
+#17529 := [monotonicity #17526]: #17528
+#17536 := [trans #17529 #17534]: #17535
+#17552 := [monotonicity #17536 #17549]: #17551
+#17555 := [monotonicity #17552]: #17554
+#17558 := [monotonicity #17555]: #17557
+#17522 := (iff #17229 #12676)
+#17523 := [rewrite]: #17522
+#17561 := [monotonicity #17523 #17558]: #17560
+#17564 := [monotonicity #17561]: #17563
+#17520 := (iff #17224 #12670)
+#17521 := [rewrite]: #17520
+#17518 := (iff #17221 #4978)
+#17519 := [rewrite]: #17518
+#17516 := (iff #17218 #12739)
+#17517 := [rewrite]: #17516
+#17514 := (iff #17215 #13720)
+#17515 := [rewrite]: #17514
+#17567 := [monotonicity #17515 #17517 #17519 #17521 #17564]: #17566
+#17572 := [trans #17567 #17570]: #17571
+#17575 := [monotonicity #17572]: #17574
+#17512 := (iff #17206 #12653)
+#17513 := [rewrite]: #17512
+#17608 := (iff #17309 #5027)
+#17609 := [rewrite]: #17608
+#17606 := (iff #17306 #5026)
+#17607 := [rewrite]: #17606
+#17612 := [monotonicity #17477 #12838 #17607 #17609 #17513 #17575]: #17611
+#17617 := [trans #17612 #17615]: #17616
+#17604 := (iff #17300 #17601)
+#17598 := (and #12644 #12828 #17595)
+#17602 := (iff #17598 #17601)
+#17603 := [rewrite]: #17602
+#17599 := (iff #17300 #17598)
+#17596 := (iff #17296 #17595)
+#17593 := (iff #17292 #17590)
+#17587 := (and #4943 #17584)
+#17591 := (iff #17587 #17590)
+#17592 := [rewrite]: #17591
+#17588 := (iff #17292 #17587)
+#17585 := (iff #17288 #17584)
+#17582 := (iff #17284 #17579)
+#17576 := (and #4946 #4950 #4955 #4960 #12639 #4963 #4965 #12653 #17573)
+#17580 := (iff #17576 #17579)
+#17581 := [rewrite]: #17580
+#17577 := (iff #17284 #17576)
+#17510 := (iff #17203 #4965)
+#17511 := [rewrite]: #17510
+#17508 := (iff #17200 #4963)
+#17509 := [rewrite]: #17508
+#17506 := (iff #17197 #12639)
+#17507 := [rewrite]: #17506
+#17504 := (iff #17194 #4960)
+#17505 := [rewrite]: #17504
+#17502 := (iff #17191 #4955)
+#17503 := [rewrite]: #17502
+#17500 := (iff #17188 #4950)
+#17501 := [rewrite]: #17500
+#17498 := (iff #17183 #4946)
+#17499 := [rewrite]: #17498
+#17578 := [monotonicity #17499 #17501 #17503 #17505 #17507 #17509 #17511 #17513 #17575]: #17577
+#17583 := [trans #17578 #17581]: #17582
+#17586 := [monotonicity #17583]: #17585
+#17496 := (iff #17177 #4943)
+#17497 := [rewrite]: #17496
+#17589 := [monotonicity #17497 #17586]: #17588
+#17594 := [trans #17589 #17592]: #17593
+#17597 := [monotonicity #17594]: #17596
+#17600 := [monotonicity #17477 #17597]: #17599
+#17605 := [trans #17600 #17603]: #17604
+#17620 := [monotonicity #17605 #17617]: #17619
+#17623 := [monotonicity #17499 #17477 #17620]: #17622
+#17628 := [trans #17623 #17626]: #17627
+#17631 := [monotonicity #17628]: #17630
+#17634 := [monotonicity #17497 #17631]: #17633
+#17639 := [trans #17634 #17637]: #17638
+#17642 := [monotonicity #17639]: #17641
+#17645 := [monotonicity #17477 #17642]: #17644
+#17650 := [trans #17645 #17648]: #17649
+#17704 := [monotonicity #17650 #17701]: #17703
+#17494 := (iff #17166 #4931)
+#17495 := [rewrite]: #17494
+#17492 := (iff #17163 #4921)
+#17493 := [rewrite]: #17492
+#17490 := (iff #17160 #4917)
+#17491 := [rewrite]: #17490
+#17488 := (iff #17157 #4913)
+#17489 := [rewrite]: #17488
+#17486 := (iff #17154 #4909)
+#17487 := [rewrite]: #17486
+#17484 := (iff #17151 #4811)
+#17485 := [rewrite]: #17484
+#17482 := (iff #17148 #4898)
+#17483 := [rewrite]: #17482
+#17480 := (iff #17145 #13099)
+#17481 := [rewrite]: #17480
+#17478 := (iff #17135 #13139)
+#17479 := [rewrite]: #17478
+#17474 := (iff #17129 #13744)
+#17475 := [rewrite]: #17474
+#17472 := (iff #17126 #13763)
+#17473 := [rewrite]: #17472
+#17470 := (iff #17123 #13778)
+#17471 := [rewrite]: #17470
+#17468 := (iff #17120 #13180)
+#17469 := [rewrite]: #17468
+#17707 := [monotonicity #17469 #17471 #17473 #17475 #17477 #17479 #17481 #17483 #17485 #17487 #17489 #17491 #17493 #17495 #17704]: #17706
+#17712 := [trans #17707 #17710]: #17711
+#17715 := [monotonicity #13344 #17712]: #17714
+#17718 := [monotonicity #17715]: #17717
+#17721 := [monotonicity #17718]: #17720
+#17466 := (iff #17087 #12595)
+#17467 := [rewrite]: #17466
+#17724 := [monotonicity #17467 #17721]: #17723
+#17727 := [monotonicity #17724]: #17726
+#17464 := (iff #17082 #4765)
+#17465 := [rewrite]: #17464
+#17462 := (iff #17079 #4760)
+#17463 := [rewrite]: #17462
+#17460 := (iff #17076 #4755)
+#17461 := [rewrite]: #17460
+#17458 := (iff #17073 #4750)
+#17459 := [rewrite]: #17458
+#17456 := (iff #17070 #4746)
+#17457 := [rewrite]: #17456
+#17730 := [monotonicity #17457 #17459 #17461 #17463 #17465 #17727]: #17729
+#17735 := [trans #17730 #17733]: #17734
+#17738 := [monotonicity #17735]: #17737
+#17454 := (iff #17064 #4742)
+#17455 := [rewrite]: #17454
+#17741 := [monotonicity #17455 #17738]: #17740
+#17746 := [trans #17741 #17744]: #17745
+#17749 := [monotonicity #17746]: #17748
+#17452 := (iff #17055 #4733)
+#17453 := [rewrite]: #17452
+#17752 := [monotonicity #17453 #17749]: #17751
+#17755 := [monotonicity #17752]: #17754
+#13641 := (exists (vars (?v0 Int)) #13642)
+#13647 := (not #13648)
+#13636 := (or #13647 #13641)
+#13635 := (and #13648 #13636)
+#13630 := (or #12923 #12647 #11388 #11379 #11370 #11361 #13635)
+#13708 := (not #13713)
+#13707 := (or #13708 #12724)
+#13702 := (and #13713 #13707)
+#13701 := (or #12681 #13702)
+#13696 := (and #12676 #13701)
+#13695 := (or #13719 #12743 #11705 #12673 #13696)
+#13690 := (and #12660 #13727 #13695)
+#13672 := (or #12647 #12828 #11892 #11883 #12656 #13690)
+#13689 := (or #11812 #11803 #11794 #11785 #12777 #11760 #11751 #12656 #13690)
+#13684 := (and #4940 #4945 #13689)
+#13683 := (or #11824 #13684)
+#13678 := (and #4940 #4942 #13683)
+#13677 := (or #12647 #12829 #13678)
+#13671 := (and #13677 #13672)
+#13666 := (or #11812 #12647 #13671)
+#13665 := (and #4940 #4945 #13666)
+#13660 := (or #11824 #13665)
+#13659 := (and #4940 #4942 #13660)
+#13654 := (or #12647 #12922 #13659)
+#13629 := (and #13654 #13630)
+#13732 := (not #13737)
+#13624 := (or #13183 #13777 #13758 #13743 #12647 #13142 #13732 #13102 #13090 #15031 #12144 #12135 #12126 #12117 #12108 #13629)
+#13623 := (and #12635 #4780 #13624)
+#13792 := (not #13797)
+#13618 := (or #13792 #13623)
+#13617 := (and #13797 #13618)
+#13612 := (or #12598 #13617)
+#13611 := (and #12595 #13612)
+#13606 := (or #12388 #12379 #12370 #12361 #12352 #13611)
+#13605 := (and #4739 #4745 #13606)
+#13600 := (or #12400 #13605)
+#13599 := (and #4739 #4741 #13600)
+#13594 := (or #12412 #13599)
+#13593 := (and #4733 #13594)
+#13588 := (not #13593)
+#17447 := (~ #13588 #17446)
+#17443 := (not #13594)
+#17444 := (~ #17443 #17442)
+#17439 := (not #13599)
+#17440 := (~ #17439 #17438)
+#17435 := (not #13600)
+#17436 := (~ #17435 #17434)
+#17431 := (not #13605)
+#17432 := (~ #17431 #17430)
+#17427 := (not #13606)
+#17428 := (~ #17427 #17426)
+#17423 := (not #13611)
+#17424 := (~ #17423 #17422)
+#17419 := (not #13612)
+#17420 := (~ #17419 #17418)
+#17415 := (not #13617)
+#17416 := (~ #17415 #17414)
+#17411 := (not #13618)
+#17412 := (~ #17411 #17410)
+#17407 := (not #13623)
+#17408 := (~ #17407 #17406)
+#17403 := (not #13624)
+#17404 := (~ #17403 #17402)
+#17399 := (not #13629)
+#17400 := (~ #17399 #17398)
+#17395 := (not #13630)
+#17396 := (~ #17395 #17394)
+#17391 := (not #13635)
+#17392 := (~ #17391 #17390)
+#17387 := (not #13636)
+#17388 := (~ #17387 #17386)
+#17383 := (not #13641)
+#17384 := (~ #17383 #17382)
+#17380 := (~ #17379 #17379)
+#17381 := [refl]: #17380
+#17385 := [nnf-neg #17381]: #17384
+#17376 := (not #13647)
+#17377 := (~ #17376 #13648)
+#17374 := (~ #13648 #13648)
+#17372 := (~ #13653 #13653)
+#17373 := [refl]: #17372
+#17375 := [nnf-pos #17373]: #17374
+#17378 := [nnf-neg #17375]: #17377
+#17389 := [nnf-neg #17378 #17385]: #17388
+#17368 := (~ #13647 #17367)
+#17369 := [sk]: #17368
+#17393 := [nnf-neg #17369 #17389]: #17392
+#17352 := (~ #17351 #17351)
+#17353 := [refl]: #17352
+#17349 := (~ #17348 #17348)
+#17350 := [refl]: #17349
+#17346 := (~ #17345 #17345)
+#17347 := [refl]: #17346
+#17343 := (~ #17342 #17342)
+#17344 := [refl]: #17343
+#17133 := (~ #17132 #17132)
+#17134 := [refl]: #17133
+#17340 := (~ #12926 #12926)
+#17341 := [refl]: #17340
+#17397 := [nnf-neg #17341 #17134 #17344 #17347 #17350 #17353 #17393]: #17396
+#17337 := (not #13654)
+#17338 := (~ #17337 #17336)
+#17333 := (not #13659)
+#17334 := (~ #17333 #17332)
+#17329 := (not #13660)
+#17330 := (~ #17329 #17328)
+#17325 := (not #13665)
+#17326 := (~ #17325 #17324)
+#17321 := (not #13666)
+#17322 := (~ #17321 #17320)
+#17317 := (not #13671)
+#17318 := (~ #17317 #17316)
+#17313 := (not #13672)
+#17314 := (~ #17313 #17312)
+#17281 := (not #13690)
+#17282 := (~ #17281 #17280)
+#17277 := (not #13695)
+#17278 := (~ #17277 #17276)
+#17273 := (not #13696)
+#17274 := (~ #17273 #17272)
+#17269 := (not #13701)
+#17270 := (~ #17269 #17268)
+#17265 := (not #13702)
+#17266 := (~ #17265 #17264)
+#17261 := (not #13707)
+#17262 := (~ #17261 #17260)
+#17258 := (~ #17257 #17257)
+#17259 := [refl]: #17258
+#17254 := (not #13708)
+#17255 := (~ #17254 #13713)
+#17252 := (~ #13713 #13713)
+#17250 := (~ #13714 #13714)
+#17251 := [refl]: #17250
+#17253 := [nnf-pos #17251]: #17252
+#17256 := [nnf-neg #17253]: #17255
+#17263 := [nnf-neg #17256 #17259]: #17262
+#17246 := (~ #13708 #17245)
+#17247 := [sk]: #17246
+#17267 := [nnf-neg #17247 #17263]: #17266
+#17230 := (~ #17229 #17229)
+#17231 := [refl]: #17230
+#17271 := [nnf-neg #17231 #17267]: #17270
+#17227 := (~ #12681 #12681)
+#17228 := [refl]: #17227
+#17275 := [nnf-neg #17228 #17271]: #17274
+#17225 := (~ #17224 #17224)
+#17226 := [refl]: #17225
+#17222 := (~ #17221 #17221)
+#17223 := [refl]: #17222
+#17219 := (~ #17218 #17218)
+#17220 := [refl]: #17219
+#17216 := (~ #17215 #17215)
+#17217 := [refl]: #17216
+#17279 := [nnf-neg #17217 #17220 #17223 #17226 #17275]: #17278
+#17213 := (~ #17212 #17212)
+#17214 := [refl]: #17213
+#17210 := (~ #17209 #17209)
+#17211 := [refl]: #17210
+#17283 := [nnf-neg #17211 #17214 #17279]: #17282
+#17207 := (~ #17206 #17206)
+#17208 := [refl]: #17207
+#17310 := (~ #17309 #17309)
+#17311 := [refl]: #17310
+#17307 := (~ #17306 #17306)
+#17308 := [refl]: #17307
+#17304 := (~ #12834 #12834)
+#17305 := [refl]: #17304
+#17315 := [nnf-neg #17134 #17305 #17308 #17311 #17208 #17283]: #17314
+#17301 := (not #13677)
+#17302 := (~ #17301 #17300)
+#17297 := (not #13678)
+#17298 := (~ #17297 #17296)
+#17293 := (not #13683)
+#17294 := (~ #17293 #17292)
+#17289 := (not #13684)
+#17290 := (~ #17289 #17288)
+#17285 := (not #13689)
+#17286 := (~ #17285 #17284)
+#17204 := (~ #17203 #17203)
+#17205 := [refl]: #17204
+#17201 := (~ #17200 #17200)
+#17202 := [refl]: #17201
+#17198 := (~ #17197 #17197)
+#17199 := [refl]: #17198
+#17195 := (~ #17194 #17194)
+#17196 := [refl]: #17195
+#17192 := (~ #17191 #17191)
+#17193 := [refl]: #17192
+#17189 := (~ #17188 #17188)
+#17190 := [refl]: #17189
+#17184 := (~ #17183 #17183)
+#17185 := [refl]: #17184
+#17287 := [nnf-neg #17185 #17190 #17193 #17196 #17199 #17202 #17205 #17208 #17283]: #17286
+#17181 := (~ #17180 #17180)
+#17182 := [refl]: #17181
+#17172 := (~ #17171 #17171)
+#17173 := [refl]: #17172
+#17291 := [nnf-neg #17173 #17182 #17287]: #17290
+#17178 := (~ #17177 #17177)
+#17179 := [refl]: #17178
+#17295 := [nnf-neg #17179 #17291]: #17294
+#17175 := (~ #17174 #17174)
+#17176 := [refl]: #17175
+#17299 := [nnf-neg #17173 #17176 #17295]: #17298
+#17186 := (~ #12828 #12828)
+#17187 := [refl]: #17186
+#17303 := [nnf-neg #17134 #17187 #17299]: #17302
+#17319 := [nnf-neg #17303 #17315]: #17318
+#17323 := [nnf-neg #17185 #17134 #17319]: #17322
+#17327 := [nnf-neg #17173 #17182 #17323]: #17326
+#17331 := [nnf-neg #17179 #17327]: #17330
+#17335 := [nnf-neg #17173 #17176 #17331]: #17334
+#17169 := (~ #12923 #12923)
+#17170 := [refl]: #17169
+#17339 := [nnf-neg #17134 #17170 #17335]: #17338
+#17401 := [nnf-neg #17339 #17397]: #17400
+#17167 := (~ #17166 #17166)
+#17168 := [refl]: #17167
+#17164 := (~ #17163 #17163)
+#17165 := [refl]: #17164
+#17161 := (~ #17160 #17160)
+#17162 := [refl]: #17161
+#17158 := (~ #17157 #17157)
+#17159 := [refl]: #17158
+#17155 := (~ #17154 #17154)
+#17156 := [refl]: #17155
+#17152 := (~ #17151 #17151)
+#17153 := [refl]: #17152
+#17149 := (~ #17148 #17148)
+#17150 := [refl]: #17149
+#17146 := (~ #17145 #17145)
+#17147 := [refl]: #17146
+#17142 := (not #13732)
+#17143 := (~ #17142 #13737)
+#17140 := (~ #13737 #13737)
+#17138 := (~ #13738 #13738)
+#17139 := [refl]: #17138
+#17141 := [nnf-pos #17139]: #17140
+#17144 := [nnf-neg #17141]: #17143
+#17136 := (~ #17135 #17135)
+#17137 := [refl]: #17136
+#17130 := (~ #17129 #17129)
+#17131 := [refl]: #17130
+#17127 := (~ #17126 #17126)
+#17128 := [refl]: #17127
+#17124 := (~ #17123 #17123)
+#17125 := [refl]: #17124
+#17121 := (~ #17120 #17120)
+#17122 := [refl]: #17121
+#17405 := [nnf-neg #17122 #17125 #17128 #17131 #17134 #17137 #17144 #17147 #17150 #17153 #17156 #17159 #17162 #17165 #17168 #17401]: #17404
+#17118 := (~ #17117 #17117)
+#17119 := [refl]: #17118
+#17115 := (~ #13340 #13340)
+#17116 := [refl]: #17115
+#17409 := [nnf-neg #17116 #17119 #17405]: #17408
+#17112 := (not #13792)
+#17113 := (~ #17112 #13797)
+#17110 := (~ #13797 #13797)
+#17108 := (~ #13798 #13798)
+#17109 := [refl]: #17108
+#17111 := [nnf-pos #17109]: #17110
+#17114 := [nnf-neg #17111]: #17113
+#17413 := [nnf-neg #17114 #17409]: #17412
+#17104 := (~ #13792 #17103)
+#17105 := [sk]: #17104
+#17417 := [nnf-neg #17105 #17413]: #17416
+#17088 := (~ #17087 #17087)
+#17089 := [refl]: #17088
+#17421 := [nnf-neg #17089 #17417]: #17420
+#17085 := (~ #12598 #12598)
+#17086 := [refl]: #17085
+#17425 := [nnf-neg #17086 #17421]: #17424
+#17083 := (~ #17082 #17082)
+#17084 := [refl]: #17083
+#17080 := (~ #17079 #17079)
+#17081 := [refl]: #17080
+#17077 := (~ #17076 #17076)
+#17078 := [refl]: #17077
+#17074 := (~ #17073 #17073)
+#17075 := [refl]: #17074
+#17071 := (~ #17070 #17070)
+#17072 := [refl]: #17071
+#17429 := [nnf-neg #17072 #17075 #17078 #17081 #17084 #17425]: #17428
+#17068 := (~ #17067 #17067)
+#17069 := [refl]: #17068
+#17059 := (~ #17058 #17058)
+#17060 := [refl]: #17059
+#17433 := [nnf-neg #17060 #17069 #17429]: #17432
+#17065 := (~ #17064 #17064)
+#17066 := [refl]: #17065
+#17437 := [nnf-neg #17066 #17433]: #17436
+#17062 := (~ #17061 #17061)
+#17063 := [refl]: #17062
+#17441 := [nnf-neg #17060 #17063 #17437]: #17440
+#17056 := (~ #17055 #17055)
+#17057 := [refl]: #17056
+#17445 := [nnf-neg #17057 #17441]: #17444
+#17053 := (~ #12412 #12412)
+#17054 := [refl]: #17053
+#17448 := [nnf-neg #17054 #17445]: #17447
+#15056 := (or #12923 #12647 #11388 #11379 #11370 #11361 #13004)
+#15061 := (and #12945 #15056)
+#15064 := (or #13183 #13177 #13164 #13154 #12647 #13142 #13136 #13102 #13090 #15031 #12144 #12135 #12126 #12117 #12108 #15061)
+#15067 := (and #12635 #4780 #15064)
+#15070 := (or #12631 #15067)
+#15073 := (and #12628 #15070)
+#15076 := (or #12598 #15073)
+#15079 := (and #12595 #15076)
+#15082 := (or #12388 #12379 #12370 #12361 #12352 #15079)
+#15085 := (and #4739 #4745 #15082)
+#15088 := (or #12400 #15085)
+#15091 := (and #4739 #4741 #15088)
+#15094 := (or #12412 #15091)
+#15097 := (and #4733 #15094)
+#15100 := (not #15097)
+#13589 := (iff #15100 #13588)
+#13590 := (iff #15097 #13593)
+#13595 := (iff #15094 #13594)
+#13596 := (iff #15091 #13599)
+#13601 := (iff #15088 #13600)
+#13602 := (iff #15085 #13605)
+#13607 := (iff #15082 #13606)
+#13608 := (iff #15079 #13611)
+#13613 := (iff #15076 #13612)
+#13614 := (iff #15073 #13617)
+#13619 := (iff #15070 #13618)
+#13620 := (iff #15067 #13623)
+#13625 := (iff #15064 #13624)
+#13626 := (iff #15061 #13629)
+#13631 := (iff #15056 #13630)
+#13632 := (iff #13004 #13635)
+#13637 := (iff #13001 #13636)
+#13638 := (iff #12998 #13641)
+#13643 := (iff #12993 #13642)
+#14912 := (iff #6890 #14917)
+#14949 := -4294967295::Int
+#14925 := (+ -4294967295::Int #1197)
+#14918 := (<= #14925 0::Int)
+#14914 := (iff #14918 #14917)
+#14915 := [rewrite]: #14914
+#14919 := (iff #6890 #14918)
+#14920 := (= #6889 #14925)
+#14926 := (+ #1197 -4294967295::Int)
+#14922 := (= #14926 #14925)
+#14923 := [rewrite]: #14922
+#14927 := (= #6889 #14926)
+#14944 := (= #6888 -4294967295::Int)
+#14950 := (* -1::Int 4294967295::Int)
+#14946 := (= #14950 -4294967295::Int)
+#14947 := [rewrite]: #14946
+#14951 := (= #6888 #14950)
+#9364 := (= f168 4294967295::Int)
+#2153 := 65536::Int
+#2552 := (* 65536::Int 65536::Int)
+#2557 := (- #2552 1::Int)
+#2558 := (= f168 #2557)
+#9365 := (iff #2558 #9364)
+#9362 := (= #2557 4294967295::Int)
+#2216 := 4294967296::Int
+#9355 := (- 4294967296::Int 1::Int)
+#9360 := (= #9355 4294967295::Int)
+#9361 := [rewrite]: #9360
+#9357 := (= #2557 #9355)
+#9326 := (= #2552 4294967296::Int)
+#9327 := [rewrite]: #9326
+#9358 := [monotonicity #9327]: #9357
+#9363 := [trans #9358 #9361]: #9362
+#9366 := [monotonicity #9363]: #9365
+#9354 := [asserted]: #2558
+#9369 := [mp #9354 #9366]: #9364
+#14948 := [monotonicity #9369]: #14951
+#14945 := [trans #14948 #14947]: #14944
+#14924 := [monotonicity #14945]: #14927
+#14921 := [trans #14924 #14923]: #14920
+#14916 := [monotonicity #14921]: #14919
+#14913 := [trans #14916 #14915]: #14912
+#13640 := [monotonicity #14913]: #13643
+#13639 := [quant-intro #13640]: #13638
+#13644 := (iff #12981 #13647)
+#13649 := (iff #12978 #13648)
+#13650 := (iff #12973 #13653)
+#14269 := (iff #7910 #14268)
+#14270 := (iff #6897 #14273)
+#14271 := [monotonicity #14913]: #14270
+#14266 := [monotonicity #14271]: #14269
+#13651 := [monotonicity #14266]: #13650
+#13646 := [quant-intro #13651]: #13649
+#13645 := [monotonicity #13646]: #13644
+#13634 := [monotonicity #13645 #13639]: #13637
+#13633 := [monotonicity #13646 #13634]: #13632
+#13628 := [monotonicity #13633]: #13631
+#13655 := (iff #12945 #13654)
+#13656 := (iff #12916 #13659)
+#13661 := (iff #12910 #13660)
+#13662 := (iff #12905 #13665)
+#13667 := (iff #12897 #13666)
+#13668 := (iff #12888 #13671)
+#13673 := (iff #12883 #13672)
+#13691 := (iff #12772 #13690)
+#13692 := (iff #12764 #13695)
+#13697 := (iff #12736 #13696)
+#13698 := (iff #12733 #13701)
+#13703 := (iff #12730 #13702)
+#13704 := (iff #12727 #13707)
+#13709 := (iff #12715 #13708)
+#13710 := (iff #12712 #13713)
+#13715 := (iff #12707 #13714)
+#13712 := [monotonicity #14266]: #13715
+#13711 := [quant-intro #13712]: #13710
+#13706 := [monotonicity #13711]: #13709
+#13705 := [monotonicity #13706]: #13704
+#13700 := [monotonicity #13711 #13705]: #13703
+#13699 := [monotonicity #13700]: #13698
+#13694 := [monotonicity #13699]: #13697
+#13716 := (iff #12749 #13719)
+#13721 := (iff #12746 #13720)
+#13722 := (iff #12662 #13727)
+#13757 := (+ 4294967295::Int #12663)
+#13731 := (>= #13757 1::Int)
+#13724 := (iff #13731 #13727)
+#13725 := [rewrite]: #13724
+#13728 := (iff #12662 #13731)
+#13754 := (= #12664 #13757)
+#13755 := [monotonicity #9369]: #13754
+#13729 := [monotonicity #13755]: #13728
+#13723 := [trans #13729 #13725]: #13722
+#13718 := [monotonicity #13723]: #13721
+#13717 := [monotonicity #13718]: #13716
+#13693 := [monotonicity #13717 #13694]: #13692
+#13688 := [monotonicity #13723 #13693]: #13691
+#13670 := [monotonicity #13688]: #13673
+#13674 := (iff #12853 #13677)
+#13679 := (iff #12823 #13678)
+#13680 := (iff #12817 #13683)
+#13685 := (iff #12812 #13684)
+#13686 := (iff #12804 #13689)
+#13687 := [monotonicity #13688]: #13686
+#13682 := [monotonicity #13687]: #13685
+#13681 := [monotonicity #13682]: #13680
+#13676 := [monotonicity #13681]: #13679
+#13675 := [monotonicity #13676]: #13674
+#13669 := [monotonicity #13675 #13670]: #13668
+#13664 := [monotonicity #13669]: #13667
+#13663 := [monotonicity #13664]: #13662
+#13658 := [monotonicity #13663]: #13661
+#13657 := [monotonicity #13658]: #13656
+#13652 := [monotonicity #13657]: #13655
+#13627 := [monotonicity #13652 #13628]: #13626
+#13733 := (iff #13136 #13732)
+#13734 := (iff #13133 #13737)
+#13739 := (iff #13128 #13738)
+#13736 := [monotonicity #14266]: #13739
+#13735 := [quant-intro #13736]: #13734
+#13730 := [monotonicity #13735]: #13733
+#13740 := (iff #13154 #13743)
+#13745 := (iff #13151 #13744)
+#13746 := (iff #13148 #13751)
+#13752 := (>= #13757 0::Int)
+#13748 := (iff #13752 #13751)
+#13749 := [rewrite]: #13748
+#13753 := (iff #13148 #13752)
+#13750 := [monotonicity #13755]: #13753
+#13747 := [trans #13750 #13749]: #13746
+#13742 := [monotonicity #13747]: #13745
+#13741 := [monotonicity #13742]: #13740
+#13759 := (iff #13164 #13758)
+#13760 := (iff #13161 #13763)
+#13765 := (iff #13157 #13766)
+#13772 := (+ 4294967295::Int #13093)
+#13771 := (>= #13772 0::Int)
+#13767 := (iff #13771 #13766)
+#13764 := [rewrite]: #13767
+#13768 := (iff #13157 #13771)
+#13773 := (= #13158 #13772)
+#13770 := [monotonicity #9369]: #13773
+#13769 := [monotonicity #13770]: #13768
+#13762 := [trans #13769 #13764]: #13765
+#13761 := [monotonicity #13762]: #13760
+#13756 := [monotonicity #13761]: #13759
+#13774 := (iff #13177 #13777)
+#13779 := (iff #13174 #13778)
+#13780 := (iff #13170 #13785)
+#13791 := (+ 255::Int #13117)
+#13786 := (>= #13791 0::Int)
+#13782 := (iff #13786 #13785)
+#13783 := [rewrite]: #13782
+#13787 := (iff #13170 #13786)
+#13788 := (= #13171 #13791)
+#2562 := (= f170 255::Int)
+#9368 := [asserted]: #2562
+#13789 := [monotonicity #9368]: #13788
+#13784 := [monotonicity #13789]: #13787
+#13781 := [trans #13784 #13783]: #13780
+#13776 := [monotonicity #13781]: #13779
+#13775 := [monotonicity #13776]: #13774
+#13622 := [monotonicity #13775 #13756 #13741 #13730 #13627]: #13625
+#13621 := [monotonicity #13622]: #13620
+#13793 := (iff #12631 #13792)
+#13794 := (iff #12628 #13797)
+#13799 := (iff #12623 #13798)
+#13796 := [monotonicity #14266]: #13799
+#13795 := [quant-intro #13796]: #13794
+#13790 := [monotonicity #13795]: #13793
+#13616 := [monotonicity #13790 #13621]: #13619
+#13615 := [monotonicity #13795 #13616]: #13614
+#13610 := [monotonicity #13615]: #13613
+#13609 := [monotonicity #13610]: #13608
+#13604 := [monotonicity #13609]: #13607
+#13603 := [monotonicity #13604]: #13602
+#13598 := [monotonicity #13603]: #13601
+#13597 := [monotonicity #13598]: #13596
+#13592 := [monotonicity #13597]: #13595
+#13591 := [monotonicity #13592]: #13590
+#13586 := [monotonicity #13591]: #13589
+#13485 := (not #13319)
+#15101 := (iff #13485 #15100)
+#15098 := (iff #13319 #15097)
+#15095 := (iff #13316 #15094)
+#15092 := (iff #13311 #15091)
+#15089 := (iff #13305 #15088)
+#15086 := (iff #13300 #15085)
+#15083 := (iff #13292 #15082)
+#15080 := (iff #13271 #15079)
+#15077 := (iff #13268 #15076)
+#15074 := (iff #13265 #15073)
+#15071 := (iff #13262 #15070)
+#15068 := (iff #13257 #15067)
+#15065 := (iff #13249 #15064)
+#15062 := (iff #13066 #15061)
+#15059 := (iff #13061 #15056)
+#15041 := (or #12647 #11388 #11379 #11370 #11361 #13004)
+#15053 := (or #12647 #12923 #15041)
+#15057 := (iff #15053 #15056)
+#15058 := [rewrite]: #15057
+#15054 := (iff #13061 #15053)
+#15051 := (iff #13036 #15041)
+#15046 := (and true #15041)
+#15049 := (iff #15046 #15041)
+#15050 := [rewrite]: #15049
+#15047 := (iff #13036 #15046)
+#15044 := (iff #13031 #15041)
+#15038 := (or false #12647 #11388 #11379 #11370 #11361 #13004)
+#15042 := (iff #15038 #15041)
+#15043 := [rewrite]: #15042
+#15039 := (iff #13031 #15038)
+#15036 := (iff #11436 false)
+#15034 := (iff #11436 #4808)
+#14750 := (iff #4072 true)
+#10920 := [asserted]: #4072
+#14751 := [iff-true #10920]: #14750
+#15035 := [monotonicity #14751]: #15034
+#15037 := [trans #15035 #11314]: #15036
+#15040 := [monotonicity #15037]: #15039
+#15045 := [trans #15040 #15043]: #15044
+#15048 := [monotonicity #14751 #15045]: #15047
+#15052 := [trans #15048 #15050]: #15051
+#15055 := [monotonicity #15052]: #15054
+#15060 := [trans #15055 #15058]: #15059
+#15063 := [monotonicity #15060]: #15062
+#15032 := (iff #11471 #15031)
+#15029 := (iff #4812 #4811)
+#15024 := (and #4811 true)
+#15027 := (iff #15024 #4811)
+#15028 := [rewrite]: #15027
+#15025 := (iff #4812 #15024)
+#15006 := (iff #4686 true)
+#15007 := [iff-true #13474]: #15006
+#15026 := [monotonicity #15007]: #15025
+#15030 := [trans #15026 #15028]: #15029
+#15033 := [monotonicity #15030]: #15032
+#15066 := [monotonicity #15033 #15063]: #15065
+#15069 := [monotonicity #15066]: #15068
+#15072 := [monotonicity #15069]: #15071
+#15075 := [monotonicity #15072]: #15074
+#15078 := [monotonicity #15075]: #15077
+#15081 := [monotonicity #15078]: #15080
+#15084 := [monotonicity #15081]: #15083
+#15087 := [monotonicity #15084]: #15086
+#15090 := [monotonicity #15087]: #15089
+#15093 := [monotonicity #15090]: #15092
+#15096 := [monotonicity #15093]: #15095
+#15099 := [monotonicity #15096]: #15098
+#15102 := [monotonicity #15099]: #15101
+#13486 := [not-or-elim #13452]: #13485
+#15103 := [mp #13486 #15102]: #15100
+#13587 := [mp #15103 #13586]: #13588
+#17449 := [mp~ #13587 #17448]: #17446
+#17450 := [mp #17449 #17755]: #17753
+#21288 := [mp #17450 #21287]: #21285
+#22120 := [mp #21288 #22119]: #22117
+#25454 := [unit-resolution #22120 #24007]: #22114
+#22345 := (or #22111 #22105)
+#22346 := [def-axiom]: #22345
+#25455 := [unit-resolution #22346 #25454]: #22105
+#22341 := (or #22108 #17058 #17061 #22102)
+#22342 := [def-axiom]: #22341
+#25456 := [unit-resolution #22342 #24483 #25074 #25455]: #22102
+#22331 := (or #22099 #22093)
+#22332 := [def-axiom]: #22331
+#25457 := [unit-resolution #22332 #25456]: #22093
+#22325 := (or #22096 #17058 #17067 #22090)
+#22326 := [def-axiom]: #22325
+#25459 := [unit-resolution #22326 #24483 #25457]: #25458
+#25460 := [unit-resolution #25459 #24749]: #22090
+#22307 := (or #22087 #4750)
+#22308 := [def-axiom]: #22307
+#25461 := [unit-resolution #22308 #25460]: #4750
+#25643 := [mp #25461 #25642]: #4780
+#22315 := (or #22087 #22081)
+#22316 := [def-axiom]: #22315
+#25644 := [unit-resolution #22316 #25460]: #22081
+#25645 := (or #22084 #22078)
+#24646 := [hypothesis]: #12598
+#24679 := [th-lemma arith farkas 1 1 #13463 #24646]: false
+#24680 := [lemma #24679]: #12595
+#22301 := (or #22084 #12598 #22078)
+#22302 := [def-axiom]: #22301
+#25646 := [unit-resolution #22302 #24680]: #25645
+#25647 := [unit-resolution #25646 #25644]: #22078
+#22293 := (or #22075 #22069)
+#22294 := [def-axiom]: #22293
+#25648 := [unit-resolution #22294 #25647]: #22069
+#25443 := (= f461 #17098)
+#25464 := (= #4749 #17098)
+#25462 := (= #17098 #4749)
+#25452 := (= #17097 #4736)
+#25450 := (= #17096 #4735)
+#25448 := (= ?v0!13 0::Int)
+#21540 := (not #17095)
+#25445 := [hypothesis]: #20925
+#21571 := (or #20920 #21540)
+#21574 := [def-axiom]: #21571
+#25446 := [unit-resolution #21574 #25445]: #21540
+#21618 := (or #20920 #17091)
+#21598 := [def-axiom]: #21618
+#25447 := [unit-resolution #21598 #25445]: #17091
+#25449 := [th-lemma arith eq-propagate 0 0 #25447 #25446]: #25448
+#25451 := [monotonicity #25449]: #25450
+#25453 := [monotonicity #25451]: #25452
+#25463 := [monotonicity #25453]: #25462
+#25465 := [symm #25463]: #25464
+#25466 := [trans #25461 #25465]: #25443
+#21550 := (not #17101)
+#21533 := (or #20920 #21550)
+#21551 := [def-axiom]: #21533
+#25467 := [unit-resolution #21551 #25445]: #21550
+#25468 := (not #25443)
+#25469 := (or #25468 #17101)
+#25470 := [th-lemma arith triangle-eq]: #25469
+#25471 := [unit-resolution #25470 #25467 #25466]: false
+#25472 := [lemma #25471]: #20920
+#22289 := (or #22072 #20925 #22066)
+#22290 := [def-axiom]: #22289
+#25649 := [unit-resolution #22290 #25472 #25648]: #22066
+#22281 := (or #22063 #22057)
+#22282 := [def-axiom]: #22281
+#25650 := [unit-resolution #22282 #25649]: #22057
+#25651 := (or #22060 #17117 #22054)
+#22277 := (or #22060 #12634 #17117 #22054)
+#22278 := [def-axiom]: #22277
+#25652 := [unit-resolution #22278 #13463]: #25651
+#25653 := [unit-resolution #25652 #25650 #25643]: #22054
+#22267 := (or #22051 #22045)
+#22268 := [def-axiom]: #22267
+#25917 := [unit-resolution #22268 #25653]: #22045
+#24775 := (+ f462 #17678)
+#24776 := (>= #24775 0::Int)
+#24763 := (+ f464 #17665)
+#24764 := (<= #24763 0::Int)
+#25683 := (not #24764)
+#22180 := (not #17667)
+#25686 := [hypothesis]: #22042
+#22215 := (or #22039 #22033)
+#22216 := [def-axiom]: #22215
+#25687 := [unit-resolution #22216 #25686]: #22033
+#22233 := (or #22051 #13766)
+#22234 := [def-axiom]: #22233
+#25688 := [unit-resolution #22234 #25653]: #13766
+#22249 := (or #22051 #4806)
+#22250 := [def-axiom]: #22249
+#25689 := [unit-resolution #22250 #25653]: #4806
+#22247 := (or #22051 #13096)
+#22248 := [def-axiom]: #22247
+#25690 := [unit-resolution #22248 #25653]: #13096
+#22241 := (or #22051 #12642)
+#22242 := [def-axiom]: #22241
+#25691 := [unit-resolution #22242 #25653]: #12642
+#22213 := (or #22039 #4820)
+#22214 := [def-axiom]: #22213
+#25692 := [unit-resolution #22214 #25686]: #4820
+#24697 := (or #22024 #21205 #21067 #13095 #21209 #11361)
+#24653 := (= #4805 f468)
+#24602 := (= f462 f468)
+#24688 := [hypothesis]: #4820
+#24690 := [symm #24688]: #24602
+#24689 := [hypothesis]: #4806
+#24691 := [trans #24689 #24690]: #24653
+#24692 := [hypothesis]: #22019
+#24693 := [hypothesis]: #13096
+#24694 := [hypothesis]: #12642
+#24695 := [hypothesis]: #13766
+#24654 := (not #24653)
+#24659 := (or #22024 #21067 #21205 #13095 #24654)
+#24546 := (+ f463 #12568)
+#24547 := (>= #24546 0::Int)
+#24655 := (or #21067 #21205 #24547 #24654)
+#24660 := (or #22024 #24655)
+#24667 := (iff #24660 #24659)
+#24656 := (or #21067 #21205 #13095 #24654)
+#24662 := (or #22024 #24656)
+#24665 := (iff #24662 #24659)
+#24666 := [rewrite]: #24665
+#24663 := (iff #24660 #24662)
+#24657 := (iff #24655 #24656)
+#24559 := (iff #24547 #13095)
+#24551 := (+ #12568 f463)
+#24554 := (>= #24551 0::Int)
+#24557 := (iff #24554 #13095)
+#24558 := [rewrite]: #24557
+#24555 := (iff #24547 #24554)
+#24552 := (= #24546 #24551)
+#24553 := [rewrite]: #24552
+#24556 := [monotonicity #24553]: #24555
+#24560 := [trans #24556 #24558]: #24559
+#24658 := [monotonicity #24560]: #24657
+#24664 := [monotonicity #24658]: #24663
+#24668 := [trans #24664 #24666]: #24667
+#24661 := [quant-inst #4786]: #24660
+#24669 := [mp #24661 #24668]: #24659
+#24696 := [unit-resolution #24669 #24695 #24694 #24693 #24692 #24691]: false
+#24698 := [lemma #24696]: #24697
+#25693 := [unit-resolution #24698 #25692 #25691 #25690 #25689 #25688]: #22024
+#22191 := (or #22027 #22019)
+#22192 := [def-axiom]: #22191
+#25694 := [unit-resolution #22192 #25693]: #22027
+#22199 := (or #22036 #21144 #22030)
+#22200 := [def-axiom]: #22199
+#25695 := [unit-resolution #22200 #25694 #25687]: #21144
+#22181 := (or #21139 #22180)
+#22182 := [def-axiom]: #22181
+#25696 := [unit-resolution #22182 #25695]: #22180
+#22201 := (or #22039 #12922)
+#22202 := [def-axiom]: #22201
+#25697 := [unit-resolution #22202 #25686]: #12922
+#25684 := (or #25683 #12923 #17667)
+#25679 := [hypothesis]: #22180
+#25680 := [hypothesis]: #12922
+#25681 := [hypothesis]: #24764
+#25682 := [th-lemma arith farkas -1 -1 1 #25681 #25680 #25679]: false
+#25685 := [lemma #25682]: #25684
+#25698 := [unit-resolution #25685 #25697 #25696]: #25683
+#25701 := (or #24764 #24776)
+#22178 := (or #21139 #17356)
+#22179 := [def-axiom]: #22178
+#25699 := [unit-resolution #22179 #25695]: #17356
+#22176 := (or #21139 #17355)
+#22177 := [def-axiom]: #22176
+#25700 := [unit-resolution #22177 #25695]: #17355
+#22245 := (or #22051 #21887)
+#22246 := [def-axiom]: #22245
+#25659 := [unit-resolution #22246 #25653]: #21887
+#25593 := (or #21892 #21123 #21124 #24764 #24776)
+#24754 := (+ #17363 #13117)
+#24755 := (<= #24754 0::Int)
+#24746 := (+ ?v0!15 #12663)
+#24747 := (>= #24746 0::Int)
+#24756 := (or #21123 #21124 #24747 #24755)
+#25594 := (or #21892 #24756)
+#25609 := (iff #25594 #25593)
+#24781 := (or #21123 #21124 #24764 #24776)
+#25604 := (or #21892 #24781)
+#25607 := (iff #25604 #25593)
+#25608 := [rewrite]: #25607
+#25605 := (iff #25594 #25604)
+#24782 := (iff #24756 #24781)
+#24779 := (iff #24755 #24776)
+#24769 := (+ #13117 #17363)
+#24772 := (<= #24769 0::Int)
+#24777 := (iff #24772 #24776)
+#24778 := [rewrite]: #24777
+#24773 := (iff #24755 #24772)
+#24770 := (= #24754 #24769)
+#24771 := [rewrite]: #24770
+#24774 := [monotonicity #24771]: #24773
+#24780 := [trans #24774 #24778]: #24779
+#24767 := (iff #24747 #24764)
+#24757 := (+ #12663 ?v0!15)
+#24760 := (>= #24757 0::Int)
+#24765 := (iff #24760 #24764)
+#24766 := [rewrite]: #24765
+#24761 := (iff #24747 #24760)
+#24758 := (= #24746 #24757)
+#24759 := [rewrite]: #24758
+#24762 := [monotonicity #24759]: #24761
+#24768 := [trans #24762 #24766]: #24767
+#24783 := [monotonicity #24768 #24780]: #24782
+#25606 := [monotonicity #24783]: #25605
+#25610 := [trans #25606 #25608]: #25609
+#25603 := [quant-inst #17354]: #25594
+#25611 := [mp #25603 #25610]: #25593
+#25702 := [unit-resolution #25611 #25659 #25700 #25699]: #25701
+#25703 := [unit-resolution #25702 #25698]: #24776
+#22183 := (not #17680)
+#22184 := (or #21139 #22183)
+#22185 := [def-axiom]: #22184
+#25704 := [unit-resolution #22185 #25695]: #22183
+#25559 := (+ f462 #12962)
+#25562 := (<= #25559 0::Int)
+#25705 := [symm #25692]: #24602
+#25706 := (not #24602)
+#25707 := (or #25706 #25562)
+#25708 := [th-lemma arith triangle-eq]: #25707
+#25709 := [unit-resolution #25708 #25705]: #25562
+#25710 := [th-lemma arith farkas -1 -1 1 #25709 #25704 #25703]: false
+#25711 := [lemma #25710]: #22039
+#22223 := (or #22048 #22008 #22042)
+#22224 := [def-axiom]: #22223
+#25918 := [unit-resolution #22224 #25711 #25917]: #22008
+#22170 := (or #22005 #12923)
+#22171 := [def-axiom]: #22170
+#25919 := [unit-resolution #22171 #25918]: #12923
+#22235 := (or #22051 #13145)
+#22236 := [def-axiom]: #22235
+#25920 := [unit-resolution #22236 #25653]: #13145
+#25721 := (or #24348 #22428 #22809 #24088 #21206 #12922 #25785)
+#25780 := (+ f464 #12568)
+#25781 := (>= #25780 0::Int)
+#25786 := (or #22428 #22809 #24088 #21206 #25781 #25785)
+#25726 := (or #24348 #25786)
+#25894 := (iff #25726 #25721)
+#25797 := (or #22428 #22809 #24088 #21206 #12922 #25785)
+#25833 := (or #24348 #25797)
+#25843 := (iff #25833 #25721)
+#25893 := [rewrite]: #25843
+#25834 := (iff #25726 #25833)
+#25798 := (iff #25786 #25797)
+#25795 := (iff #25781 #12922)
+#25787 := (+ #12568 f464)
+#25790 := (>= #25787 0::Int)
+#25793 := (iff #25790 #12922)
+#25794 := [rewrite]: #25793
+#25791 := (iff #25781 #25790)
+#25788 := (= #25780 #25787)
+#25789 := [rewrite]: #25788
+#25792 := [monotonicity #25789]: #25791
+#25796 := [trans #25792 #25794]: #25795
+#25799 := [monotonicity #25796]: #25798
+#25842 := [monotonicity #25799]: #25834
+#25895 := [trans #25842 #25893]: #25894
+#25747 := [quant-inst #4649 #4655 #23413 #4646 #4790 #356]: #25726
+#25896 := [mp #25747 #25895]: #25721
+#25921 := [unit-resolution #25896 #20277 #11138 #13474 #25920 #25919 #24429 #25916]: false
+#25922 := [lemma #25921]: #25785
+#25615 := (or #25784 #4942)
+#25616 := [def-axiom]: #25615
+#25947 := [unit-resolution #25616 #25922]: #4942
+#26236 := (= #25848 #4941)
+#26234 := (= #25639 #4937)
+#24370 := (f153 f154 #23991)
+#25586 := (f140 #24370 f464)
+#25595 := (f139 #25586 f35)
+#26232 := (= #25595 #4937)
+#25809 := (= #4937 #25595)
+#25807 := (= #4936 #25586)
+#25800 := (= #25586 #4936)
+#25804 := (= #24370 #4734)
+#25802 := (= #23991 #4656)
+#25816 := [symm #25275]: #24457
+#25817 := (= #23991 #24041)
+#25801 := [trans #25100 #24456]: #25817
+#25803 := [trans #25801 #25816]: #25802
+#25805 := [monotonicity #25803]: #25804
+#25806 := [monotonicity #25805]: #25800
+#25808 := [symm #25806]: #25807
+#25810 := [monotonicity #25808]: #25809
+#26233 := [symm #25810]: #26232
+#26216 := (= #25639 #25595)
+#25612 := (= #25595 #25639)
+#25712 := (not #25612)
+#25613 := (f125 f243 #25595)
+#25619 := (f71 #25613 #23991)
+#25620 := (= #25619 f1)
+#25621 := (not #25620)
+#25715 := (or #25621 #25712)
+#25718 := (not #25715)
+#25568 := (or #24299 #25718)
+#25622 := (* f464 #4624)
+#25623 := (+ #24379 #25622)
+#25626 := (f87 #4654 #25623)
+#25627 := (= #25595 #25626)
+#25625 := (not #25627)
+#25628 := (or #25621 #25625)
+#25624 := (not #25628)
+#24275 := (or #24299 #25624)
+#25580 := (iff #24275 #25568)
+#25583 := (iff #25568 #25568)
+#25584 := [rewrite]: #25583
+#25719 := (iff #25624 #25718)
+#25716 := (iff #25628 #25715)
+#25713 := (iff #25625 #25712)
+#25635 := (iff #25627 #25612)
+#25640 := (= #25626 #25639)
+#25633 := (= #25623 #25632)
+#25630 := (= #25622 #25629)
+#25631 := [rewrite]: #25630
+#25634 := [monotonicity #25631]: #25633
+#25563 := [monotonicity #25634]: #25640
+#25636 := [monotonicity #25563]: #25635
+#25714 := [monotonicity #25636]: #25713
+#25717 := [monotonicity #25714]: #25716
+#25720 := [monotonicity #25717]: #25719
+#25582 := [monotonicity #25720]: #25580
+#25599 := [trans #25582 #25584]: #25580
+#25581 := [quant-inst #23991 #4790 #356]: #24275
+#25597 := [mp #25581 #25599]: #25568
+#25814 := [unit-resolution #25597 #19813]: #25718
+#25602 := (or #25715 #25612)
+#25614 := [def-axiom]: #25602
+#25815 := [unit-resolution #25614 #25814]: #25612
+#26217 := [symm #25815]: #26216
+#26235 := [trans #26217 #26233]: #26234
+#26237 := [monotonicity #26235]: #26236
+#26238 := [trans #26237 #25947]: #25849
+#25850 := (not #25849)
+#25886 := (or #25850 #25885)
+#25887 := (not #25886)
+#25846 := (f71 #4743 #25639)
+#25847 := (= #25846 f1)
+#25888 := (iff #25847 #25887)
+#26057 := (or #24794 #25888)
+#26049 := [quant-inst #4649 #25639]: #26057
+#26055 := [unit-resolution #26049 #20682]: #25888
+#26188 := (not #25847)
+#26198 := (iff #17180 #26188)
+#26212 := (iff #4945 #25847)
+#25844 := (iff #25847 #4945)
+#26209 := (= #25846 #4944)
+#26210 := [monotonicity #26235]: #26209
+#26211 := [monotonicity #26210]: #25844
+#26213 := [symm #26211]: #26212
+#26199 := [monotonicity #26213]: #26198
+#26056 := [hypothesis]: #17180
+#26197 := [mp #26056 #26199]: #26188
+#26037 := (not #25888)
+#26038 := (or #26037 #25847 #25886)
+#26187 := [def-axiom]: #26038
+#25845 := [unit-resolution #26187 #26197 #26055]: #25886
+#26179 := (or #25887 #25850 #25885)
+#25779 := [def-axiom]: #26179
+#26207 := [unit-resolution #25779 #25845 #26238]: #25885
+#26280 := (= #25852 #22792)
+#25617 := (= #25851 f35)
+#25592 := (f62 f63 #4937)
+#25565 := (= #25592 f35)
+#25585 := (iff #4940 #25565)
+#24274 := (or #23440 #25585)
+#24269 := [quant-inst #4937 #356]: #24274
+#25750 := [unit-resolution #24269 #21822]: #25585
+#24273 := (not #25585)
+#26208 := (or #24273 #25565)
+#25724 := (or #23455 #25617)
+#25725 := [quant-inst #356 #25632]: #25724
+#25813 := [unit-resolution #25725 #21835]: #25617
+#25826 := (= #25592 #25851)
+#25827 := (= #4937 #25639)
+#25828 := [trans #25810 #25815]: #25827
+#25829 := [monotonicity #25828]: #25826
+#25830 := [trans #25829 #25813]: #25565
+#24270 := (not #25565)
+#25752 := (or #24273 #24270)
+#25751 := [hypothesis]: #17171
+#24212 := (or #24273 #4940 #24270)
+#25579 := [def-axiom]: #24212
+#25811 := [unit-resolution #25579 #25751]: #25752
+#25812 := [unit-resolution #25811 #25750]: #24270
+#25831 := [unit-resolution #25812 #25830]: false
+#25832 := [lemma #25831]: #4940
+#25905 := (or #24273 #17171 #25565)
+#25908 := [def-axiom]: #25905
+#26219 := [unit-resolution #25908 #25832]: #26208
+#26220 := [unit-resolution #26219 #25750]: #25565
+#26218 := (= #25851 #25592)
+#26221 := [monotonicity #26235]: #26218
+#26279 := [trans #26221 #26220]: #25617
+#26281 := [monotonicity #26279]: #26280
+#26175 := [trans #26281 #24539]: #25853
+#25596 := (not #25783)
+#26276 := (iff #25596 #25858)
+#26282 := (iff #25783 #25857)
+#26166 := (iff #25857 #25783)
+#26164 := (= #25856 #25782)
+#26271 := (= #25855 #24580)
+#26174 := (= #24580 #25855)
+#26284 := [monotonicity #25828]: #26174
+#26272 := [symm #26284]: #26271
+#26165 := [monotonicity #26272]: #26164
+#26274 := [monotonicity #26165]: #26166
+#26283 := [symm #26274]: #26282
+#26277 := [monotonicity #26283]: #26276
+#25598 := (or #25784 #25596)
+#25731 := [def-axiom]: #25598
+#26176 := [unit-resolution #25731 #25922]: #25596
+#26275 := [mp #26176 #26277]: #25858
+#26050 := (or #25863 #25857)
+#26066 := [def-axiom]: #26050
+#26270 := [unit-resolution #26066 #26275]: #25863
+#26278 := (or #25875 #25854 #25864)
+#26553 := (+ #24890 #25629)
+#26751 := (= #25632 #26553)
+#26752 := (* -1::Int #26553)
+#26753 := (+ #25632 #26752)
+#26754 := (<= #26753 0::Int)
+#24569 := (* -1::Int #23971)
+#24572 := (+ #22490 #24569)
+#24574 := (>= #24572 0::Int)
+#24568 := (= #22490 #23971)
+#26764 := (= #4657 #23971)
+#26762 := (= #23971 #4657)
+#26761 := [trans #24456 #25816]: #24459
+#26763 := [monotonicity #26761]: #26762
+#26765 := [symm #26763]: #26764
+#26766 := [trans #25251 #26765]: #24568
+#26767 := (not #24568)
+#26789 := (or #26767 #24574)
+#26790 := [th-lemma arith triangle-eq]: #26789
+#26791 := [unit-resolution #26790 #26766]: #24574
+#25530 := (* -1::Int #24379)
+#25531 := (+ #23971 #25530)
+#25533 := (>= #25531 0::Int)
+#25529 := (= #23971 #24379)
+#26771 := (= #24379 #23971)
+#26772 := [monotonicity #25100]: #26771
+#26773 := [symm #26772]: #25529
+#26774 := (not #25529)
+#26792 := (or #26774 #25533)
+#26793 := [th-lemma arith triangle-eq]: #26792
+#26794 := [unit-resolution #26793 #26773]: #25533
+#26658 := (* -1::Int #24890)
+#26659 := (+ #22490 #26658)
+#26660 := (<= #26659 0::Int)
+#26653 := (= #22490 #24890)
+#26778 := [symm #25253]: #26653
+#26779 := (not #26653)
+#26795 := (or #26779 #26660)
+#26796 := [th-lemma arith triangle-eq]: #26795
+#26797 := [unit-resolution #26796 #26778]: #26660
+#26800 := (not #24574)
+#26799 := (not #26660)
+#26798 := (not #25533)
+#26801 := (or #26754 #26798 #26799 #26800)
+#26802 := [th-lemma arith assign-bounds 1 -1 1]: #26801
+#26803 := [unit-resolution #26802 #26797 #26794 #26791]: #26754
+#26755 := (>= #26753 0::Int)
+#24573 := (<= #24572 0::Int)
+#26768 := (or #26767 #24573)
+#26769 := [th-lemma arith triangle-eq]: #26768
+#26770 := [unit-resolution #26769 #26766]: #24573
+#25532 := (<= #25531 0::Int)
+#26775 := (or #26774 #25532)
+#26776 := [th-lemma arith triangle-eq]: #26775
+#26777 := [unit-resolution #26776 #26773]: #25532
+#26673 := (>= #26659 0::Int)
+#26780 := (or #26779 #26673)
+#26781 := [th-lemma arith triangle-eq]: #26780
+#26782 := [unit-resolution #26781 #26778]: #26673
+#26785 := (not #24573)
+#26784 := (not #26673)
+#26783 := (not #25532)
+#26786 := (or #26755 #26783 #26784 #26785)
+#26787 := [th-lemma arith assign-bounds 1 -1 1]: #26786
+#26788 := [unit-resolution #26787 #26782 #26777 #26770]: #26755
+#26805 := (not #26755)
+#26804 := (not #26754)
+#26806 := (or #26751 #26804 #26805)
+#26807 := [th-lemma arith triangle-eq]: #26806
+#26388 := [unit-resolution #26807 #26788 #26803]: #26751
+#26908 := (not #26751)
+#26909 := (or #26908 #25869)
+#26904 := (= #25868 #4662)
+#26845 := (= #25859 #4658)
+#26843 := (= #25859 #24084)
+#26447 := (f140 #25120 f464)
+#26448 := (f139 #26447 f35)
+#26449 := (f134 #4883 #26448)
+#26450 := (f235 f236 #26449)
+#26451 := (= #26450 #24084)
+#26458 := (f71 #4667 #26448)
+#26459 := (= #26458 f1)
+#26460 := (not #26459)
+#26455 := (f155 f156 #26449)
+#26456 := (= #26455 f1)
+#26457 := (not #26456)
+#26453 := (f155 f237 #26449)
+#26454 := (= #26453 f1)
+#26452 := (not #26451)
+#26461 := (or #26452 #26454 #26457 #26460)
+#26462 := (not #26461)
+#26380 := (or #25115 #25119 #21206 #12922 #26462)
+#26463 := (or #25119 #21206 #25781 #26462)
+#26381 := (or #25115 #26463)
+#25840 := (iff #26381 #26380)
+#26464 := (or #25119 #21206 #12922 #26462)
+#26020 := (or #25115 #26464)
+#25835 := (iff #26020 #26380)
+#25836 := [rewrite]: #25835
+#26177 := (iff #26381 #26020)
+#26465 := (iff #26463 #26464)
+#26466 := [monotonicity #25796]: #26465
+#26178 := [monotonicity #26466]: #26177
+#26352 := [trans #26178 #25836]: #25840
+#26413 := [quant-inst #4649 #4655 #356 #4646 #4790]: #26381
+#26405 := [mp #26413 #26352]: #26380
+#26757 := [unit-resolution #26405 #19597 #25920 #25919 #25186]: #26462
+#26406 := (or #26461 #26451)
+#26417 := [def-axiom]: #26406
+#26758 := [unit-resolution #26417 #26757]: #26451
+#26841 := (= #25859 #26450)
+#26839 := (= #25855 #26449)
+#26837 := (= #26449 #25855)
+#26835 := (= #26448 #25639)
+#26833 := (= #26448 #25595)
+#26831 := (= #26448 #4937)
+#24810 := (f55 f206 #4937)
+#25570 := (f87 #4654 #24810)
+#26825 := (= #25570 #4937)
+#25571 := (= #4937 #25570)
+#25900 := (or #23430 #17171 #25571)
+#25591 := (or #17171 #25571)
+#25901 := (or #23430 #25591)
+#25904 := (iff #25901 #25900)
+#25906 := [rewrite]: #25904
+#25903 := [quant-inst #4937 #356]: #25901
+#25907 := [mp #25903 #25906]: #25900
+#26759 := [unit-resolution #25907 #16892 #25832]: #25571
+#26826 := [symm #26759]: #26825
+#26829 := (= #26448 #25570)
+#26556 := (f87 #4654 #26553)
+#26823 := (= #26556 #25570)
+#26813 := (= #26553 #24810)
+#26811 := (= #25632 #24810)
+#26746 := (= #24810 #25632)
+#26747 := (* -1::Int #25632)
+#26748 := (+ #24810 #26747)
+#26749 := (<= #26748 0::Int)
+#25818 := (f55 f206 #25639)
+#25823 := (* -1::Int #25818)
+#25824 := (+ #25629 #25823)
+#25825 := (+ #24379 #25824)
+#25940 := (>= #25825 0::Int)
+#25821 := (= #25825 0::Int)
+#25915 := (or #23460 #25821)
+#25819 := (= #25818 #25632)
+#25923 := (or #23460 #25819)
+#25933 := (iff #25923 #25915)
+#25935 := (iff #25915 #25915)
+#25936 := [rewrite]: #25935
+#25820 := (iff #25819 #25821)
+#25822 := [rewrite]: #25820
+#25934 := [monotonicity #25822]: #25933
+#25937 := [trans #25934 #25936]: #25933
+#25932 := [quant-inst #356 #25632]: #25923
+#25938 := [mp #25932 #25937]: #25915
+#26537 := [unit-resolution #25938 #21829]: #25821
+#26517 := (not #25821)
+#26520 := (or #26517 #25940)
+#26519 := [th-lemma arith triangle-eq]: #26520
+#26521 := [unit-resolution #26519 #26537]: #25940
+#25942 := (+ #24810 #25823)
+#25945 := (<= #25942 0::Int)
+#25941 := (= #24810 #25818)
+#26522 := (= #25818 #24810)
+#26538 := [monotonicity #26235]: #26522
+#26580 := [symm #26538]: #25941
+#26581 := (not #25941)
+#26571 := (or #26581 #25945)
+#26640 := [th-lemma arith triangle-eq]: #26571
+#26492 := [unit-resolution #26640 #26580]: #25945
+#26506 := (not #25940)
+#26577 := (not #25945)
+#26578 := (or #26749 #26577 #26506)
+#26579 := [th-lemma arith assign-bounds -1 1]: #26578
+#26498 := [unit-resolution #26579 #26492 #26521]: #26749
+#26750 := (>= #26748 0::Int)
+#25939 := (<= #25825 0::Int)
+#26582 := (or #26517 #25939)
+#26497 := [th-lemma arith triangle-eq]: #26582
+#26576 := [unit-resolution #26497 #26537]: #25939
+#25946 := (>= #25942 0::Int)
+#26584 := (or #26581 #25946)
+#26575 := [th-lemma arith triangle-eq]: #26584
+#26585 := [unit-resolution #26575 #26580]: #25946
+#26573 := (not #25939)
+#26572 := (not #25946)
+#26439 := (or #26750 #26572 #26573)
+#26587 := [th-lemma arith assign-bounds -1 1]: #26439
+#26583 := [unit-resolution #26587 #26585 #26576]: #26750
+#26586 := (not #26750)
+#26588 := (not #26749)
+#26637 := (or #26746 #26588 #26586)
+#26594 := [th-lemma arith triangle-eq]: #26637
+#26595 := [unit-resolution #26594 #26583 #26498]: #26746
+#26892 := [symm #26595]: #26811
+#26809 := (= #26553 #25632)
+#26890 := [hypothesis]: #26751
+#26891 := [symm #26890]: #26809
+#26893 := [trans #26891 #26892]: #26813
+#26894 := [monotonicity #26893]: #26823
+#26827 := (= #26448 #26556)
+#26535 := (f140 #25193 f464)
+#26536 := (f139 #26535 f35)
+#26559 := (= #26536 #26556)
+#26562 := (not #26559)
+#26543 := (f125 f243 #26536)
+#26544 := (f71 #26543 #23413)
+#26545 := (= #26544 f1)
+#26546 := (not #26545)
+#26565 := (or #26546 #26562)
+#26568 := (not #26565)
+#26542 := (or #24299 #26568)
+#26547 := (+ #24890 #25622)
+#26548 := (f87 #4654 #26547)
+#26549 := (= #26536 #26548)
+#26550 := (not #26549)
+#26551 := (or #26546 #26550)
+#26552 := (not #26551)
+#26574 := (or #24299 #26552)
+#26629 := (iff #26574 #26542)
+#26665 := (iff #26542 #26542)
+#26671 := [rewrite]: #26665
+#26569 := (iff #26552 #26568)
+#26566 := (iff #26551 #26565)
+#26563 := (iff #26550 #26562)
+#26560 := (iff #26549 #26559)
+#26557 := (= #26548 #26556)
+#26554 := (= #26547 #26553)
+#26555 := [monotonicity #25631]: #26554
+#26558 := [monotonicity #26555]: #26557
+#26561 := [monotonicity #26558]: #26560
+#26564 := [monotonicity #26561]: #26563
+#26567 := [monotonicity #26564]: #26566
+#26570 := [monotonicity #26567]: #26569
+#26630 := [monotonicity #26570]: #26629
+#26642 := [trans #26630 #26671]: #26629
+#26628 := [quant-inst #23413 #4790 #356]: #26574
+#26645 := [mp #26628 #26642]: #26542
+#26815 := [unit-resolution #26645 #19813]: #26568
+#26654 := (or #26565 #26559)
+#26655 := [def-axiom]: #26654
+#26816 := [unit-resolution #26655 #26815]: #26559
+#26821 := (= #26448 #26536)
+#26819 := (= #26447 #26535)
+#26817 := (= #26535 #26447)
+#26818 := [monotonicity #25265]: #26817
+#26820 := [symm #26818]: #26819
+#26822 := [monotonicity #26820]: #26821
+#26828 := [trans #26822 #26816]: #26827
+#26895 := [trans #26828 #26894]: #26829
+#26896 := [trans #26895 #26826]: #26831
+#26897 := [trans #26896 #25810]: #26833
+#26898 := [trans #26897 #25815]: #26835
+#26899 := [monotonicity #26898]: #26837
+#26900 := [symm #26899]: #26839
+#26901 := [monotonicity #26900]: #26841
+#26902 := [trans #26901 #26758]: #26843
+#26903 := [trans #26902 #24984]: #26845
+#26905 := [monotonicity #26903]: #26904
+#26906 := [trans #26905 #13466]: #25869
+#26000 := (not #25869)
+#26889 := [hypothesis]: #26000
+#26907 := [unit-resolution #26889 #26906]: false
+#26910 := [lemma #26907]: #26909
+#26273 := [unit-resolution #26910 #26388]: #25869
+#26070 := (or #25872 #26000)
+#26074 := [def-axiom]: #26070
+#26291 := [unit-resolution #26074 #26273]: #25872
+#26039 := (not #25867)
+#26047 := (f45 f79 #25865)
+#26048 := (= #26047 f1)
+#26058 := (not #26048)
+#26009 := (or #25867 #26058)
+#26010 := (not #26009)
+#26035 := [hypothesis]: #26009
+#26181 := (or #24538 #26010)
+#26182 := [quant-inst #25855]: #26181
+#26036 := [unit-resolution #26182 #20844 #26035]: false
+#26194 := [lemma #26036]: #26010
+#25999 := (or #26009 #26039)
+#26001 := [def-axiom]: #25999
+#26292 := [unit-resolution #26001 #26194]: #26039
+#26118 := (or #25875 #25854 #25864 #25867 #25873)
+#26119 := [def-axiom]: #26118
+#26295 := [unit-resolution #26119 #26292 #26291]: #26278
+#26296 := [unit-resolution #26295 #26270 #26175]: #25875
+#26167 := (or #25884 #25874)
+#26168 := [def-axiom]: #26167
+#26300 := [unit-resolution #26168 #26296 #26207]: false
+#26293 := [lemma #26300]: #4945
+#26678 := (= f464 ?v0!14)
+#26712 := (not #26678)
+#26680 := (= #4947 #17241)
+#26686 := (not #26680)
+#26685 := (+ #4947 #17543)
+#26687 := (>= #26685 0::Int)
+#26696 := (not #26687)
+#26018 := (+ #4947 #12696)
+#26019 := (<= #26018 0::Int)
+#26867 := [hypothesis]: #12829
+#21362 := (+ f462 #12696)
+#21363 := (<= #21362 0::Int)
+#21359 := (= f462 f470)
+#26738 := (iff #5026 #21359)
+#26736 := (iff #21359 #5026)
+#26737 := [commutativity]: #26736
+#26739 := [symm #26737]: #26738
+#26868 := (or #17180 #21984)
+#22172 := (or #22005 #21999)
+#22173 := [def-axiom]: #22172
+#25948 := [unit-resolution #22173 #25918]: #21999
+#22164 := (or #22002 #17171 #17174 #21996)
+#22165 := [def-axiom]: #22164
+#25949 := [unit-resolution #22165 #25948]: #21999
+#25950 := [unit-resolution #25949 #25947 #25832]: #21996
+#22154 := (or #21993 #21987)
+#22155 := [def-axiom]: #22154
+#25951 := [unit-resolution #22155 #25950]: #21987
+#22148 := (or #21990 #17171 #17180 #21984)
+#22149 := [def-axiom]: #22148
+#26875 := [unit-resolution #22149 #25832 #25951]: #26868
+#26876 := [unit-resolution #26875 #26293]: #21984
+#22138 := (or #21981 #21975)
+#22139 := [def-axiom]: #22138
+#26885 := [unit-resolution #22139 #26876]: #21975
+#21370 := (or #21963 #12828)
+#21372 := [def-axiom]: #21370
+#26886 := [unit-resolution #21372 #26867]: #21963
+#22128 := (or #21978 #21966 #21972)
+#22129 := [def-axiom]: #22128
+#26734 := [unit-resolution #22129 #26886 #26885]: #21972
+#21356 := (or #21969 #5026)
+#21357 := [def-axiom]: #21356
+#26735 := [unit-resolution #21357 #26734]: #5026
+#26740 := [mp #26735 #26739]: #21359
+#26741 := (not #21359)
+#26744 := (or #26741 #21363)
+#26877 := [th-lemma arith triangle-eq]: #26744
+#26878 := [unit-resolution #26877 #26740]: #21363
+#26879 := (not #21363)
+#26880 := (or #26019 #12828 #26879)
+#26881 := [th-lemma arith assign-bounds 1 -1]: #26880
+#26882 := [unit-resolution #26881 #26878 #26867]: #26019
+#21516 := (not #17545)
+#26067 := [hypothesis]: #21936
+#21350 := (or #21969 #21933)
+#22121 := [def-axiom]: #21350
+#26379 := [unit-resolution #22121 #26067]: #21969
+#26007 := (or #21957 #21972)
+#25952 := [hypothesis]: #21960
+#21379 := (or #21957 #21951)
+#21380 := [def-axiom]: #21379
+#25953 := [unit-resolution #21380 #25952]: #21951
+#21385 := (or #21954 #17171 #17174 #21948)
+#21387 := [def-axiom]: #21385
+#25891 := [unit-resolution #21387 #25953 #25832 #25947]: #21948
+#21411 := (or #21945 #4945)
+#21412 := [def-axiom]: #21411
+#25892 := [unit-resolution #21412 #25891]: #4945
+#26002 := [hypothesis]: #21969
+#21373 := (or #21963 #21957)
+#21374 := [def-axiom]: #21373
+#26003 := [unit-resolution #21374 #25952]: #21963
+#26004 := [unit-resolution #22129 #26003 #26002]: #21978
+#26005 := [unit-resolution #22139 #26004]: #21981
+#26006 := [unit-resolution #22149 #26005 #25892 #25832 #25951]: false
+#26008 := [lemma #26006]: #26007
+#26025 := [unit-resolution #26008 #26379]: #21957
+#21423 := (or #21939 #21933)
+#21424 := [def-axiom]: #21423
+#26666 := [unit-resolution #21424 #26067]: #21939
+#26614 := (or #21948 #17180 #21942)
+#21398 := (or #21948 #17171 #17180 #21942)
+#21399 := [def-axiom]: #21398
+#26613 := [unit-resolution #21399 #25832]: #26614
+#26616 := [unit-resolution #26613 #26666 #26293]: #21948
+#21392 := (or #21951 #21945)
+#21404 := [def-axiom]: #21392
+#26617 := [unit-resolution #21404 #26616]: #21951
+#26639 := (or #21960 #21954)
+#21383 := (or #21960 #17171 #17174 #21954)
+#21378 := [def-axiom]: #21383
+#26591 := [unit-resolution #21378 #25832 #25947]: #26639
+#26589 := [unit-resolution #26591 #26617 #26025]: false
+#26625 := [lemma #26589]: #21933
+#26427 := (or #21936 #21930)
+#13804 := (<= f443 4294967295::Int)
+#13803 := (iff #12567 #13804)
+#13810 := (+ 4294967295::Int #12568)
+#13809 := (>= #13810 0::Int)
+#13805 := (iff #13809 #13804)
+#13802 := [rewrite]: #13805
+#13806 := (iff #12567 #13809)
+#13811 := (= #12569 #13810)
+#13808 := [monotonicity #9369]: #13811
+#13807 := [monotonicity #13808]: #13806
+#13800 := [trans #13807 #13802]: #13803
+#13482 := [not-or-elim #13452]: #12572
+#13484 := [and-elim #13482]: #12567
+#13801 := [mp #13484 #13800]: #13804
+#26422 := (not #13804)
+#26423 := (or #13727 #26422 #12922)
+#26424 := [th-lemma arith assign-bounds -1 1]: #26423
+#26411 := [unit-resolution #26424 #25919 #13801]: #13727
+#26425 := (or #21206 #12660)
+#26416 := [th-lemma arith farkas 1 1]: #26425
+#26426 := [unit-resolution #26416 #25920]: #12660
+#21458 := (or #21936 #17209 #17212 #21930)
+#21450 := [def-axiom]: #21458
+#26414 := [unit-resolution #21450 #26426 #26411]: #26427
+#26883 := [unit-resolution #26414 #26625]: #21930
+#21469 := (or #21927 #21921)
+#21477 := [def-axiom]: #21469
+#26884 := [unit-resolution #21477 #26883]: #21921
+#21524 := (>= #12740 -1::Int)
+#21468 := (or #21927 #12739)
+#21470 := [def-axiom]: #21468
+#26887 := [unit-resolution #21470 #26883]: #12739
+#26434 := (or #12743 #21524)
+#26435 := [th-lemma arith triangle-eq]: #26434
+#26888 := [unit-resolution #26435 #26887]: #21524
+#26442 := (not #21524)
+#26911 := (or #12676 #26442)
+#26436 := (or #12676 #26442 #12922)
+#26443 := [th-lemma arith assign-bounds -1 -1]: #26436
+#26912 := [unit-resolution #26443 #25919]: #26911
+#26913 := [unit-resolution #26912 #26888]: #12676
+#21487 := (or #21924 #12681 #21918)
+#21488 := [def-axiom]: #21487
+#26914 := [unit-resolution #21488 #26913 #26884]: #21918
+#21478 := (or #21915 #21909)
+#21480 := [def-axiom]: #21478
+#26915 := [unit-resolution #21480 #26914]: #21909
+#26923 := [symm #26735]: #21359
+#26924 := (= #4990 f462)
+#26921 := (= #4990 #4805)
+#26919 := (= #4989 #4804)
+#26917 := (= #4988 #4803)
+#21353 := (or #21969 #5027)
+#21358 := [def-axiom]: #21353
+#26916 := [unit-resolution #21358 #26734]: #5027
+#26918 := [monotonicity #26916]: #26917
+#26920 := [monotonicity #26918]: #26919
+#26922 := [monotonicity #26920]: #26921
+#26925 := [trans #26922 #25689]: #26924
+#26926 := [trans #26925 #26923]: #4991
+#21368 := (+ f463 #12718)
+#21369 := (>= #21368 0::Int)
+#21367 := (= f463 f471)
+#26929 := (iff #5027 #21367)
+#26927 := (iff #21367 #5027)
+#26928 := [commutativity]: #26927
+#26930 := [symm #26928]: #26929
+#26931 := [mp #26916 #26930]: #21367
+#26932 := (not #21367)
+#26933 := (or #26932 #21369)
+#26934 := [th-lemma arith triangle-eq]: #26933
+#26935 := [unit-resolution #26934 #26931]: #21369
+#26936 := (not #21369)
+#26937 := (or #12721 #13095 #26936)
+#26938 := [th-lemma arith assign-bounds -1 -1]: #26937
+#26939 := [unit-resolution #26938 #26935 #25690]: #12721
+#21505 := (or #20987 #12720 #20985)
+#21497 := [def-axiom]: #21505
+#26940 := [unit-resolution #21497 #26939 #26926]: #20987
+#21502 := (or #21903 #20986)
+#21506 := [def-axiom]: #21502
+#26941 := [unit-resolution #21506 #26940]: #21903
+#21494 := (or #21912 #20971 #21906)
+#21495 := [def-axiom]: #21494
+#26942 := [unit-resolution #21495 #26941 #26915]: #20971
+#21519 := (or #20966 #21516)
+#21517 := [def-axiom]: #21519
+#26943 := [unit-resolution #21517 #26942]: #21516
+#26697 := (not #26019)
+#26698 := (or #26696 #17545 #26697)
+#26692 := [hypothesis]: #26687
+#26693 := [hypothesis]: #26019
+#26694 := [hypothesis]: #21516
+#26695 := [th-lemma arith farkas -1 -1 1 #26694 #26693 #26692]: false
+#26699 := [lemma #26695]: #26698
+#26944 := [unit-resolution #26699 #26943 #26882]: #26696
+#26688 := (or #26686 #26687)
+#26689 := [th-lemma arith triangle-eq]: #26688
+#26945 := [unit-resolution #26689 #26944]: #26686
+#26713 := (or #26712 #26680)
+#26708 := (= #17241 #4947)
+#26706 := (= #17240 #4937)
+#26704 := (= #17239 #4936)
+#26702 := (= ?v0!14 f464)
+#26701 := [hypothesis]: #26678
+#26703 := [symm #26701]: #26702
+#26705 := [monotonicity #26703]: #26704
+#26707 := [monotonicity #26705]: #26706
+#26709 := [monotonicity #26707]: #26708
+#26710 := [symm #26709]: #26680
+#26700 := [hypothesis]: #26686
+#26711 := [unit-resolution #26700 #26710]: false
+#26714 := [lemma #26711]: #26713
+#26946 := [unit-resolution #26714 #26945]: #26712
+#26320 := (+ f464 #17530)
+#26420 := (>= #26320 0::Int)
+#21530 := (not #17532)
+#21509 := (or #20966 #21530)
+#21512 := [def-axiom]: #21509
+#26947 := [unit-resolution #21512 #26942]: #21530
+#26948 := (or #26420 #26442 #17532)
+#26949 := [th-lemma arith assign-bounds -1 -1]: #26948
+#26950 := [unit-resolution #26949 #26947 #26888]: #26420
+#26321 := (<= #26320 0::Int)
+#26332 := (+ f462 #17543)
+#26333 := (>= #26332 0::Int)
+#26493 := (not #26333)
+#26951 := (or #26493 #17545 #26879)
+#26952 := [th-lemma arith assign-bounds -1 -1]: #26951
+#26953 := [unit-resolution #26952 #26878 #26943]: #26493
+#21525 := (or #20966 #17234)
+#21527 := [def-axiom]: #21525
+#26954 := [unit-resolution #21527 #26942]: #17234
+#21528 := (or #20966 #17233)
+#21529 := [def-axiom]: #21528
+#26955 := [unit-resolution #21529 #26942]: #17233
+#26341 := (or #21892 #20950 #20951 #26321 #26333)
+#26311 := (+ #17241 #13117)
+#26312 := (<= #26311 0::Int)
+#26303 := (+ ?v0!14 #12663)
+#26304 := (>= #26303 0::Int)
+#26313 := (or #20950 #20951 #26304 #26312)
+#26342 := (or #21892 #26313)
+#26349 := (iff #26342 #26341)
+#26338 := (or #20950 #20951 #26321 #26333)
+#26344 := (or #21892 #26338)
+#26347 := (iff #26344 #26341)
+#26348 := [rewrite]: #26347
+#26345 := (iff #26342 #26344)
+#26339 := (iff #26313 #26338)
+#26336 := (iff #26312 #26333)
+#26326 := (+ #13117 #17241)
+#26329 := (<= #26326 0::Int)
+#26334 := (iff #26329 #26333)
+#26335 := [rewrite]: #26334
+#26330 := (iff #26312 #26329)
+#26327 := (= #26311 #26326)
+#26328 := [rewrite]: #26327
+#26331 := [monotonicity #26328]: #26330
+#26337 := [trans #26331 #26335]: #26336
+#26324 := (iff #26304 #26321)
+#26314 := (+ #12663 ?v0!14)
+#26317 := (>= #26314 0::Int)
+#26322 := (iff #26317 #26321)
+#26323 := [rewrite]: #26322
+#26318 := (iff #26304 #26317)
+#26315 := (= #26303 #26314)
+#26316 := [rewrite]: #26315
+#26319 := [monotonicity #26316]: #26318
+#26325 := [trans #26319 #26323]: #26324
+#26340 := [monotonicity #26325 #26337]: #26339
+#26346 := [monotonicity #26340]: #26345
+#26350 := [trans #26346 #26348]: #26349
+#26343 := [quant-inst #17232]: #26342
+#26351 := [mp #26343 #26350]: #26341
+#26956 := [unit-resolution #26351 #25659 #26955 #26954 #26953]: #26321
+#26539 := (not #26420)
+#26518 := (not #26321)
+#26526 := (or #26678 #26518 #26539)
+#26527 := [th-lemma arith triangle-eq]: #26526
+#26957 := [unit-resolution #26527 #26956 #26950 #26946]: false
+#26958 := [lemma #26957]: #12828
+#26540 := (or #21939 #12829)
+#26421 := [hypothesis]: #21942
+#26419 := [unit-resolution #21424 #26421]: #21933
+#26428 := [unit-resolution #26414 #26419]: #21930
+#26429 := [unit-resolution #21477 #26428]: #21921
+#26432 := [unit-resolution #21470 #26428]: #12739
+#26441 := [unit-resolution #26435 #26432]: #21524
+#26444 := [unit-resolution #26443 #26441 #25919]: #12676
+#26440 := [unit-resolution #21488 #26444 #26429]: #21918
+#26478 := [unit-resolution #21480 #26440]: #21909
+#26485 := (= f469 f470)
+#21437 := (or #21939 #4963)
+#21447 := [def-axiom]: #21437
+#26479 := [unit-resolution #21447 #26421]: #4963
+#26486 := [symm #26479]: #26485
+#26487 := (= #4990 f469)
+#26415 := (= #4947 f469)
+#21442 := (or #21939 #4950)
+#21443 := [def-axiom]: #21442
+#26445 := [unit-resolution #21443 #26421]: #4950
+#26484 := [symm #26445]: #26415
+#26482 := (= #4990 #4947)
+#26480 := (= #4989 #4937)
+#26469 := (= #4988 #4936)
+#21414 := (or #21939 #4965)
+#21416 := [def-axiom]: #21414
+#26468 := [unit-resolution #21416 #26421]: #4965
+#26470 := [monotonicity #26468]: #26469
+#26481 := [monotonicity #26470]: #26480
+#26483 := [monotonicity #26481]: #26482
+#26488 := [trans #26483 #26484]: #26487
+#26505 := [trans #26488 #26486]: #4991
+#26014 := (+ f464 #12718)
+#26016 := (>= #26014 0::Int)
+#26013 := (= f464 f471)
+#26431 := [symm #26468]: #26013
+#26471 := (not #26013)
+#26472 := (or #26471 #26016)
+#26467 := [th-lemma arith triangle-eq]: #26472
+#26473 := [unit-resolution #26467 #26431]: #26016
+#26418 := (not #26016)
+#26474 := (or #12721 #26418 #12922)
+#26475 := [th-lemma arith assign-bounds -1 -1]: #26474
+#26476 := [unit-resolution #26475 #26473 #25919]: #12721
+#26477 := [unit-resolution #21497 #26476 #26505]: #20987
+#26433 := [unit-resolution #21506 #26477]: #21903
+#26509 := [unit-resolution #21495 #26433 #26478]: #20971
+#26500 := [unit-resolution #21512 #26509]: #21530
+#26017 := (= #4947 f470)
+#26501 := [trans #26484 #26486]: #26017
+#26499 := (not #26017)
+#26502 := (or #26499 #26019)
+#26503 := [th-lemma arith triangle-eq]: #26502
+#26504 := [unit-resolution #26503 #26501]: #26019
+#26510 := [unit-resolution #21517 #26509]: #21516
+#26511 := [unit-resolution #26699 #26510 #26504]: #26696
+#26507 := [unit-resolution #26689 #26511]: #26686
+#26491 := [unit-resolution #26714 #26507]: #26712
+#26525 := (or #26678 #26539)
+#26494 := [hypothesis]: #12828
+#26495 := (or #26493 #17545 #26697 #12829)
+#26496 := [th-lemma arith assign-bounds 1 1 1]: #26495
+#26508 := [unit-resolution #26496 #26510 #26504 #26494]: #26493
+#26512 := (or #26321 #26333)
+#26523 := [unit-resolution #21527 #26509]: #17234
+#26524 := [unit-resolution #21529 #26509]: #17233
+#26531 := [unit-resolution #26351 #25659 #26524 #26523]: #26512
+#26532 := [unit-resolution #26531 #26508]: #26321
+#26528 := [unit-resolution #26527 #26532]: #26525
+#26529 := [unit-resolution #26528 #26491]: #26539
+#26530 := [th-lemma arith farkas 1 -1 1 #26441 #26529 #26500]: false
+#26541 := [lemma #26530]: #26540
+#26716 := [unit-resolution #26541 #26958]: #21939
+#26618 := [unit-resolution #26613 #26716 #26293]: #21948
+#21352 := (or #21969 #12829)
+#21355 := [def-axiom]: #21352
+#26661 := [unit-resolution #21355 #26958]: #21969
+#26593 := [unit-resolution #26008 #26661]: #21957
+#26438 := [unit-resolution #26591 #26593]: #21954
+[unit-resolution #21404 #26438 #26618]: false
+unsat
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/SMT_Examples/boogie.ML	Tue Jul 23 18:36:23 2013 +0200
@@ -0,0 +1,319 @@
+(*  Title:      HOL/SMT_Examples/boogie.ML
+    Author:     Sascha Boehme, TU Muenchen
+
+Proving Boogie-generated verficiation conditions.
+*)
+
+signature BOOGIE =
+sig
+  val boogie_prove: string -> theory -> theory
+end
+
+structure Boogie: BOOGIE =
+struct
+
+(* utility functions *)
+
+val as_int = fst o read_int o raw_explode
+
+
+val isabelle_name =
+  let 
+    fun purge s = if Symbol.is_letter s orelse Symbol.is_digit s then s else
+      (case s of
+        "." => "_o_"
+      | "_" => "_n_"
+      | "$" => "_S_"
+      | "@" => "_G_"
+      | "#" => "_H_"
+      | "^" => "_T_"
+      | _   => ("_" ^ string_of_int (ord s) ^ "_"))
+  in prefix "b_" o translate_string purge end
+
+
+
+(* context *)
+
+type context =
+  typ Symtab.table * (term * bool) Symtab.table * term list * term list
+
+val empty_context: context = (Symtab.empty, Symtab.empty, [], [])
+
+
+fun add_type name (tds, fds, axs, vcs) =
+  let
+    val T = TFree (isabelle_name name, @{sort type})
+    val tds' = Symtab.update (name, T) tds
+  in (tds', fds, axs, vcs) end
+
+
+fun add_func name Ts T unique (tds, fds, axs, vcs) =
+  let
+    val t = Free (isabelle_name name, Ts ---> T)
+    val fds' = Symtab.update (name, (t, unique)) fds
+  in (tds, fds', axs, vcs) end
+
+
+fun add_axiom t (tds, fds, axs, vcs) = (tds, fds, t :: axs, vcs)
+
+fun add_vc t (tds, fds, axs, vcs) = (tds, fds, axs, t :: vcs)
+
+
+fun lookup_type (tds, _, _, _) name =
+  (case Symtab.lookup tds name of
+    SOME T => T
+  | NONE => error "Undeclared type")
+
+
+fun lookup_func (_, fds, _, _) name =
+  (case Symtab.lookup fds name of
+    SOME t_unique => t_unique
+  | NONE => error "Undeclared function")
+
+
+
+(* constructors *)
+
+fun mk_var name T = Free ("V_" ^ isabelle_name name, T)
+
+
+fun mk_arrayT (Ts, T) = Type (@{type_name "fun"}, [HOLogic.mk_tupleT Ts, T])
+
+
+fun mk_binary t (t1, t2) = t $ t1 $ t2
+
+
+fun mk_nary _ t [] = t
+  | mk_nary f _ ts = uncurry (fold_rev f) (split_last ts)
+
+
+fun mk_distinct [] = @{const HOL.True}
+  | mk_distinct [_] = @{const HOL.True}
+  | mk_distinct (t :: ts) =
+      let
+        fun mk_noteq u u' =
+          HOLogic.mk_conj (HOLogic.mk_not (HOLogic.mk_eq (t, u)), u')
+      in fold_rev mk_noteq ts (mk_distinct ts) end
+
+
+fun mk_store m k v =
+  let
+    val mT = Term.fastype_of m and kT = Term.fastype_of k
+    val vT = Term.fastype_of v
+  in Const (@{const_name fun_upd}, mT --> kT --> vT --> mT) $ m $ k $ v end
+
+
+fun mk_quant q (Free (x, T)) t = q T $ absfree (x, T) t
+  | mk_quant _ t _ = raise TERM ("bad variable", [t])
+
+
+fun mk_list T = HOLogic.mk_list T
+
+
+val patternT = @{typ "SMT.pattern"}
+
+fun mk_pat t =
+  Const (@{const_name "SMT.pat"}, Term.fastype_of t --> patternT) $ t
+
+fun mk_pattern [] = raise TERM ("mk_pattern", [])
+  | mk_pattern ts = mk_list patternT (map mk_pat ts)
+
+
+fun mk_trigger [] t = t
+  | mk_trigger pss t =
+      @{term "SMT.trigger"} $
+        mk_list @{typ "SMT.pattern list"} (map mk_pattern pss) $ t
+
+
+(* parser *)
+
+fun repeat f n ls =
+  let fun apply (xs, ls) = f ls |>> (fn x => x :: xs)
+  in funpow (as_int n) apply ([], ls) |>> rev end
+
+
+fun parse_type _ (["bool"] :: ls) = (@{typ bool}, ls)
+  | parse_type _ (["int"] :: ls) = (@{typ int}, ls)
+  | parse_type cx (["array", arity] :: ls) =
+      repeat (parse_type cx) arity ls |>> mk_arrayT o split_last
+  | parse_type cx (("type-con" :: name :: _) :: ls) = (lookup_type cx name, ls)
+  | parse_type _ _ = error "Bad type"
+
+
+fun parse_expr _ (["true"] :: ls) = (@{term True}, ls)
+  | parse_expr _ (["false"] :: ls) = (@{term False}, ls)
+  | parse_expr cx (["not"] :: ls) = parse_expr cx ls |>> HOLogic.mk_not
+  | parse_expr cx (["and", n] :: ls) =
+      parse_nary_expr cx n HOLogic.mk_conj @{term True} ls
+  | parse_expr cx (["or", n] :: ls) =
+      parse_nary_expr cx n HOLogic.mk_disj @{term False} ls
+  | parse_expr cx (["implies"] :: ls) =
+      parse_bin_expr cx (mk_binary @{term HOL.implies}) ls
+  | parse_expr cx (["="] :: ls) = parse_bin_expr cx HOLogic.mk_eq ls
+  | parse_expr cx (["var", name] :: ls) = parse_type cx ls |>> mk_var name
+  | parse_expr cx (["fun", name, n] :: ls) =
+      let val (t, _) = lookup_func cx name
+      in repeat (parse_expr cx) n ls |>> curry Term.list_comb t end
+  | parse_expr cx (("label" :: _) :: ls) = parse_expr cx ls
+  | parse_expr _ (["int-num", n] :: ls) =
+      (HOLogic.mk_number @{typ int} (as_int n), ls)
+  | parse_expr cx (["<"] :: ls) =
+      parse_bin_expr cx (mk_binary @{term "op < :: int => _"}) ls
+  | parse_expr cx (["<="] :: ls) =
+      parse_bin_expr cx (mk_binary @{term "op <= :: int => _"}) ls
+  | parse_expr cx ([">"] :: ls) =
+      parse_bin_expr cx (mk_binary @{term "op < :: int => _"}o swap) ls
+  | parse_expr cx ([">="] :: ls) =
+      parse_bin_expr cx (mk_binary @{term "op <= :: int => _"} o swap) ls
+  | parse_expr cx (["+"] :: ls) =
+      parse_bin_expr cx (mk_binary @{term "op + :: int => _"}) ls
+  | parse_expr cx (["-"] :: ls) =
+      parse_bin_expr cx (mk_binary @{term "op - :: int => _"}) ls
+  | parse_expr cx (["*"] :: ls) =
+      parse_bin_expr cx (mk_binary @{term "op * :: int => _"}) ls
+  | parse_expr cx (["/"] :: ls) =
+      parse_bin_expr cx (mk_binary @{term boogie_div}) ls
+  | parse_expr cx (["%"] :: ls) =
+      parse_bin_expr cx (mk_binary @{term boogie_mod}) ls
+  | parse_expr cx (["select", n] :: ls) =
+      repeat (parse_expr cx) n ls
+      |>> (fn ts => hd ts $ HOLogic.mk_tuple (tl ts))
+  | parse_expr cx (["store", n] :: ls) =
+      repeat (parse_expr cx) n ls
+      |>> split_last
+      |>> (fn (ts, t) => mk_store (hd ts) (HOLogic.mk_tuple (tl ts)) t)
+  | parse_expr cx (["forall", vars, pats, atts] :: ls) =
+      parse_quant cx HOLogic.all_const vars pats atts ls
+  | parse_expr cx (["exists", vars, pats, atts] :: ls) =
+      parse_quant cx HOLogic.exists_const vars pats atts ls
+  | parse_expr _ _ = error "Bad expression"
+
+
+and parse_bin_expr cx f ls = ls |> parse_expr cx ||>> parse_expr cx |>> f
+
+
+and parse_nary_expr cx n f c ls =
+  repeat (parse_expr cx) n ls |>> mk_nary (curry f) c
+
+
+and parse_quant cx q vars pats atts ls =
+  let
+    val ((((vs, pss), _), t), ls') =
+      ls
+      |> repeat (parse_var cx) vars
+      ||>> repeat (parse_pat cx) pats
+      ||>> repeat (parse_attr cx) atts
+      ||>> parse_expr cx
+  in (fold_rev (mk_quant q) vs (mk_trigger pss t), ls') end
+
+
+and parse_var cx (["var", name] :: ls) = parse_type cx ls |>> mk_var name
+  | parse_var _ _ = error "Bad variable"
+
+
+and parse_pat cx (["pat", n] :: ls) = repeat (parse_expr cx) n ls
+  | parse_pat _ _ = error "Bad pattern"
+
+
+and parse_attr cx (["attribute", name, n] :: ls) =
+      let
+        fun attr (["expr-attr"] :: ls) = parse_expr cx ls |>> K ()
+          | attr (("string-attr" :: _) :: ls) = ((), ls)
+          | attr _ = error "Bad attribute value"
+      in repeat attr n ls |>> K name end
+  | parse_attr _ _ = error "Bad attribute"
+
+
+fun parse_func cx arity n ls =
+  let
+    val ((Ts, atts), ls') =
+      ls |> repeat (parse_type cx) arity ||>> repeat (parse_attr cx) n
+    val unique = member (op =) atts "unique"
+  in ((split_last Ts, unique), ls') end
+
+
+fun parse_decl (("type-decl" :: name :: _) :: ls) cx = (ls, add_type name cx)
+  | parse_decl (["fun-decl", name, arity, n] :: ls) cx =
+      let val (((Ts, T), unique), ls') = parse_func cx arity n ls
+      in (ls', add_func name Ts T unique cx) end
+  | parse_decl (("axiom" :: _) :: ls) cx =
+      let val (t, ls') = parse_expr cx ls
+      in (ls', add_axiom t cx) end
+  | parse_decl (("var-decl" :: _) :: ls) cx =
+      parse_type cx ls |> snd |> rpair cx
+  | parse_decl (("vc" :: _) :: ls) cx =
+      let val (t, ls') = parse_expr cx ls
+      in (ls', add_vc t cx) end
+  | parse_decl _ _ = error "Bad declaration"
+
+
+fun parse_lines [] cx = cx
+  | parse_lines ls cx = parse_decl ls cx |-> parse_lines
+
+
+
+(* splitting of text into a lists of lists of tokens *)
+
+fun is_blank c = (c = " " orelse c = "\t" orelse c = "\r" orelse c = "\n")
+
+fun explode_lines text =
+  text
+  |> split_lines
+  |> map (String.tokens (is_blank o str))
+  |> filter (fn [] => false | _ => true)
+
+
+
+(* proving verification conditions *)
+
+fun add_unique_axioms (tds, fds, axs, vcs) =
+  Symtab.fold (fn (_, (t, true)) => cons t | _ => I) fds []
+  |> map (swap o Term.dest_Free)
+  |> AList.group (op =)
+  |> map (fn (T, ns) => mk_distinct (map (Free o rpair T) ns))
+  |> (fn axs' => (tds, fds, axs' @ axs, vcs))
+
+
+fun build_proof_context thy (tds, fds, axs, vcs) =
+  let
+    val vc =
+      (case vcs of
+        [vc] => vc
+      | _ => error "Bad number of verification conditions")
+  in
+    Proof_Context.init_global thy
+    |> Symtab.fold (fn (_, T) => Variable.declare_typ T) tds
+    |> Symtab.fold (fn (_, (t, _)) => Variable.declare_term t) fds
+    |> fold Variable.declare_term axs
+    |> fold Variable.declare_term vcs
+    |> pair (map HOLogic.mk_Trueprop axs, HOLogic.mk_Trueprop vc)
+  end
+
+
+val boogie_rules =
+  [@{thm fst_conv}, @{thm snd_conv}, @{thm pair_collapse}] @
+  [@{thm fun_upd_same}, @{thm fun_upd_apply}]
+
+
+fun boogie_tac ctxt axioms =
+  ALLGOALS (SMT_Solver.smt_tac ctxt (boogie_rules @ axioms))
+
+
+fun boogie_prove file_name thy =
+  let
+    val (text, thy') = Thy_Load.use_file (Path.explode file_name) thy
+    val lines = explode_lines text
+    
+    val ((axioms, vc), ctxt) =
+      empty_context
+      |> parse_lines lines
+      |> add_unique_axioms
+      |> build_proof_context thy'
+
+    val _ = Goal.prove ctxt [] axioms vc (fn {prems, context} =>
+      boogie_tac context prems)
+    val _ = Output.writeln "Verification condition proved successfully"
+
+  in thy' end
+
+end