# HG changeset patch # User nipkow # Date 1257331259 -3600 # Node ID 1464ddca182b4f685feb089b73015080f399c1c0 # Parent 6f825ec18b4999c2577184a557efe1ba3520343e# Parent 2b5b0f9e271c7db2d72bd64aa236f328f099201f merged diff -r 2b5b0f9e271c -r 1464ddca182b Admin/isatest/isatest-settings --- a/Admin/isatest/isatest-settings Thu Oct 29 16:23:57 2009 +0100 +++ b/Admin/isatest/isatest-settings Wed Nov 04 11:40:59 2009 +0100 @@ -11,7 +11,7 @@ HOME=/home/isatest ## send email on failure to -MAILTO="kleing@cse.unsw.edu.au nipkow@in.tum.de berghofe@in.tum.de schirmer@in.tum.de lp15@cam.ac.uk makarius@sketis.net haftmann@in.tum.de krauss@in.tum.de blanchet@in.tum.de" +MAILTO="kleing@cse.unsw.edu.au nipkow@in.tum.de berghofe@in.tum.de schirmer@in.tum.de lp15@cam.ac.uk makarius@sketis.net haftmann@in.tum.de krauss@in.tum.de blanchet@in.tum.de bulwahn@in.tum.de" LOGPREFIX=$HOME/log MASTERLOG=$LOGPREFIX/isatest.log diff -r 2b5b0f9e271c -r 1464ddca182b Admin/isatest/settings/mac-poly-M4 --- a/Admin/isatest/settings/mac-poly-M4 Thu Oct 29 16:23:57 2009 +0100 +++ b/Admin/isatest/settings/mac-poly-M4 Wed Nov 04 11:40:59 2009 +0100 @@ -1,7 +1,7 @@ # -*- shell-script -*- :mode=shellscript: - POLYML_HOME="/home/polyml/polyml-5.2.1" - ML_SYSTEM="polyml-5.2.1" + POLYML_HOME="/home/polyml/polyml-svn" + ML_SYSTEM="polyml-experimental" ML_PLATFORM="x86-darwin" ML_HOME="$POLYML_HOME/$ML_PLATFORM" ML_OPTIONS="--mutable 800 --immutable 2000" diff -r 2b5b0f9e271c -r 1464ddca182b Admin/isatest/settings/mac-poly-M8 --- a/Admin/isatest/settings/mac-poly-M8 Thu Oct 29 16:23:57 2009 +0100 +++ b/Admin/isatest/settings/mac-poly-M8 Wed Nov 04 11:40:59 2009 +0100 @@ -1,7 +1,7 @@ # -*- shell-script -*- :mode=shellscript: - POLYML_HOME="/home/polyml/polyml-5.2.1" - ML_SYSTEM="polyml-5.2.1" + POLYML_HOME="/home/polyml/polyml-svn" + ML_SYSTEM="polyml-experimental" ML_PLATFORM="x86-darwin" ML_HOME="$POLYML_HOME/$ML_PLATFORM" ML_OPTIONS="--mutable 800 --immutable 2000" diff -r 2b5b0f9e271c -r 1464ddca182b Admin/spass/README --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Admin/spass/README Wed Nov 04 11:40:59 2009 +0100 @@ -0,0 +1,33 @@ + +This distribution of SPASS 3.0 has been compiled according to the +INSTALL instructions of the official source download from +http://www.spass-prover.org + +The following example is for x86-linux: + + $ tar xvzf spass30.tgz + $ cd SPASS-3.0 + $ ./configure --prefix="$HOME/tmp/spass" --exec-prefix="$HOME/tmp/spass/x86-linux" + $ make + $ make install + + +The x86-darwin platform works the same; use "make distclean" before +reconfiguration. + +Native x86_64 is avoided here, due to numerous compiler warnings about +int/pointer casts of different size. The x86 binaries are used +instead. + +For x86-cygwin we have removed the "__inline__" configuration to +workaround a problem with "___sgetc_r": + +diff configure-orig configure +3522c3522 +< CFLAGS="-O3 -D__inline__=" +--- +> CFLAGS="-O3" + + + Makarius + 30-Oct-2009 diff -r 2b5b0f9e271c -r 1464ddca182b CONTRIBUTORS --- a/CONTRIBUTORS Thu Oct 29 16:23:57 2009 +0100 +++ b/CONTRIBUTORS Wed Nov 04 11:40:59 2009 +0100 @@ -7,6 +7,9 @@ Contributions to this Isabelle version -------------------------------------- +* November 2009: Sascha Boehme, TUM + HOL-Boogie: an interactive prover back-end for Boogie and VCC + * October 2009: Jasmin Blanchette, TUM Nitpick: yet another counterexample generator for Isabelle/HOL diff -r 2b5b0f9e271c -r 1464ddca182b NEWS --- a/NEWS Thu Oct 29 16:23:57 2009 +0100 +++ b/NEWS Wed Nov 04 11:40:59 2009 +0100 @@ -37,6 +37,11 @@ *** HOL *** +* Combined former theories Divides and IntDiv to one theory Divides +in the spirit of other number theory theories in HOL; some constants +(and to a lesser extent also facts) have been suffixed with _nat und _int +respectively. INCOMPATIBILITY. + * Most rules produced by inductive and datatype package have mandatory prefixes. INCOMPATIBILITY. @@ -50,6 +55,10 @@ this method is proof-producing. Certificates are provided to avoid calling the external solvers solely for re-checking proofs. +* New commands to load and prove verification conditions +generated by the Boogie program verifier or derived systems +(e.g. the Verifying C Compiler (VCC) or Spec#). + * New counterexample generator tool "nitpick" based on the Kodkod relational model finder. diff -r 2b5b0f9e271c -r 1464ddca182b src/CCL/Wfd.thy --- a/src/CCL/Wfd.thy Thu Oct 29 16:23:57 2009 +0100 +++ b/src/CCL/Wfd.thy Wed Nov 04 11:40:59 2009 +0100 @@ -440,7 +440,7 @@ fun IHinst tac rls = SUBGOAL (fn (Bi,i) => let val bvs = bvars Bi [] - val ihs = List.filter could_IH (Logic.strip_assums_hyp Bi) + val ihs = filter could_IH (Logic.strip_assums_hyp Bi) val rnames = map (fn x=> let val (a,b) = get_bno [] 0 x in (List.nth(bvs,a),b) end) ihs diff -r 2b5b0f9e271c -r 1464ddca182b src/FOL/IFOL.thy --- a/src/FOL/IFOL.thy Thu Oct 29 16:23:57 2009 +0100 +++ b/src/FOL/IFOL.thy Wed Nov 04 11:40:59 2009 +0100 @@ -646,7 +646,7 @@ and [Pure.elim 2] = allE notE' impE' and [Pure.intro] = exI disjI2 disjI1 -setup {* ContextRules.addSWrapper (fn tac => hyp_subst_tac ORELSE' tac) *} +setup {* Context_Rules.addSWrapper (fn tac => hyp_subst_tac ORELSE' tac) *} lemma iff_not_sym: "~ (Q <-> P) ==> ~ (P <-> Q)" diff -r 2b5b0f9e271c -r 1464ddca182b src/FOLP/simp.ML --- a/src/FOLP/simp.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/FOLP/simp.ML Wed Nov 04 11:40:59 2009 +0100 @@ -66,7 +66,7 @@ fun eq_brl ((b1 : bool, th1), (b2, th2)) = b1 = b2 andalso Thm.eq_thm_prop (th1, th2); (*insert a thm in a discrimination net by its lhs*) -fun lhs_insert_thm (th,net) = +fun lhs_insert_thm th net = Net.insert_term eq_brl (lhs_of (concl_of th), (false,th)) net handle Net.INSERT => net; @@ -215,7 +215,7 @@ fun add_norm_tags congs = let val ccs = map cong_const congs - val new_asms = List.filter (exists not o #2) + val new_asms = filter (exists not o #2) (ccs ~~ (map (map atomic o prems_of) congs)); in add_norms(congs,ccs,new_asms) end; @@ -434,7 +434,7 @@ val thms = map (trivial o cterm_of(Thm.theory_of_thm thm)) As; val new_rws = maps mk_rew_rules thms; val rwrls = map mk_trans (maps mk_rew_rules thms); - val anet' = List.foldr lhs_insert_thm anet rwrls + val anet' = fold_rev lhs_insert_thm rwrls anet; in if !tracing andalso not(null new_rws) then writeln (cat_lines ("Adding rewrites:" :: map Display.string_of_thm_without_context new_rws)) diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/ATP_Linkup.thy --- a/src/HOL/ATP_Linkup.thy Thu Oct 29 16:23:57 2009 +0100 +++ b/src/HOL/ATP_Linkup.thy Wed Nov 04 11:40:59 2009 +0100 @@ -91,9 +91,9 @@ subsection {* Setup of external ATPs *} -use "Tools/res_axioms.ML" setup ResAxioms.setup +use "Tools/res_axioms.ML" setup Res_Axioms.setup use "Tools/res_hol_clause.ML" -use "Tools/res_reconstruct.ML" setup ResReconstruct.setup +use "Tools/res_reconstruct.ML" setup Res_Reconstruct.setup use "Tools/res_atp.ML" use "Tools/ATP_Manager/atp_wrapper.ML" setup ATP_Wrapper.setup diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/Auth/KerberosIV.thy --- a/src/HOL/Auth/KerberosIV.thy Thu Oct 29 16:23:57 2009 +0100 +++ b/src/HOL/Auth/KerberosIV.thy Wed Nov 04 11:40:59 2009 +0100 @@ -780,7 +780,7 @@ lemma u_NotexpiredSK_NotexpiredAK: "\ \ expiredSK Ts evs; servKlife + Ts <= authKlife + Ta \ \ \ expiredAK Ta evs" - by (metis nat_add_commute le_less_trans) + by (metis le_less_trans) subsection{* Reliability: friendly agents send something if something else happened*} diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/Boogie/Boogie.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Boogie/Boogie.thy Wed Nov 04 11:40:59 2009 +0100 @@ -0,0 +1,208 @@ +(* Title: HOL/Boogie/Boogie.thy + Author: Sascha Boehme, TU Muenchen +*) + +header {* Integration of the Boogie program verifier *} + +theory Boogie +imports SMT +uses + ("Tools/boogie_vcs.ML") + ("Tools/boogie_loader.ML") + ("Tools/boogie_commands.ML") + ("Tools/boogie_split.ML") +begin + +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 \ bool \ bool" where "assert_at l P = P" +definition block_at :: "bool \ bool \ bool" where "block_at l P = P" + +lemmas labels = assert_at_def block_at_def + + +subsection {* Arrays *} + +abbreviation (input) boogie_select :: "('i \ 'v) \ 'i \ 'v" +where "boogie_select \ (\f x. f x)" + +abbreviation (input) boogie_store :: + "('i \ 'v) \ 'i \ 'v \ 'i \ 'v" +where "boogie_store \ fun_upd" + + +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 \ int \ int" (infixl "div'_b" 70) and + boogie_mod :: "int \ int \ int" (infixl "mod'_b" 70) + + +subsection {* Bitvectors *} + +text {* +Boogie's and Z3's built-in bitvector functions are modelled with +functions of the HOL-Word library and the SMT theory of bitvectors. +Every of the following bitvector functions is supported by the SMT +binding. +*} + +abbreviation (input) boogie_bv_concat :: + "'a::len0 word \ 'b::len0 word \ 'c::len0 word" +where "boogie_bv_concat \ (\w1 w2. word_cat w1 w2)" + +abbreviation (input) boogie_bv_extract :: + "nat \ nat \ 'a::len0 word \ 'b::len0 word" +where "boogie_bv_extract \ (\mb lb w. slice lb w)" + +abbreviation (input) z3_bvnot :: "'a::len0 word \ 'a word" +where "z3_bvnot \ (\w. NOT w)" + +abbreviation (input) z3_bvand :: "'a::len0 word \ 'a word \ 'a word" +where "z3_bvand \ (\w1 w2. w1 AND w2)" + +abbreviation (input) z3_bvor :: "'a::len0 word \ 'a word \ 'a word" +where "z3_bvor \ (\w1 w2. w1 OR w2)" + +abbreviation (input) z3_bvxor :: "'a::len0 word \ 'a word \ 'a word" +where "z3_bvxor \ (\w1 w2. w1 XOR w2)" + +abbreviation (input) z3_bvneg :: "'a::len0 word \ 'a word" +where "z3_bvneg \ (\w. - w)" + +abbreviation (input) z3_bvadd :: "'a::len0 word \ 'a word \ 'a word" +where "z3_bvadd \ (\w1 w2. w1 + w2)" + +abbreviation (input) z3_bvsub :: "'a::len0 word \ 'a word \ 'a word" +where "z3_bvsub \ (\w1 w2. w1 - w2)" + +abbreviation (input) z3_bvmul :: "'a::len0 word \ 'a word \ 'a word" +where "z3_bvmul \ (\w1 w2. w1 * w2)" + +abbreviation (input) z3_bvudiv :: "'a::len0 word \ 'a word \ 'a word" +where "z3_bvudiv \ (\w1 w2. w1 div w2)" + +abbreviation (input) z3_bvurem :: "'a::len0 word \ 'a word \ 'a word" +where "z3_bvurem \ (\w1 w2. w1 mod w2)" + +abbreviation (input) z3_bvsdiv :: "'a::len word \ 'a word \ 'a word" +where "z3_bvsdiv \ (\w1 w2. w1 sdiv w2)" + +abbreviation (input) z3_bvsrem :: "'a::len word \ 'a word \ 'a word" +where "z3_bvsrem \ (\w1 w2. w1 srem w2)" + +abbreviation (input) z3_bvshl :: "'a::len0 word \ 'a word \ 'a word" +where "z3_bvshl \ (\w1 w2. bv_shl w1 w2)" + +abbreviation (input) z3_bvlshr :: "'a::len0 word \ 'a word \ 'a word" +where "z3_bvlshr \ (\w1 w2. bv_lshr w1 w2)" + +abbreviation (input) z3_bvashr :: "'a::len word \ 'a word \ 'a word" +where "z3_bvashr \ (\w1 w2. bv_ashr w1 w2)" + +abbreviation (input) z3_sign_extend :: "'a::len word \ 'b::len word" +where "z3_sign_extend \ (\w. scast w)" + +abbreviation (input) z3_zero_extend :: "'a::len0 word \ 'b::len0 word" +where "z3_zero_extend \ (\w. ucast w)" + +abbreviation (input) z3_rotate_left :: "nat \ 'a::len0 word \ 'a word" +where "z3_rotate_left \ (\n w. word_rotl n w)" + +abbreviation (input) z3_rotate_right :: "nat \ 'a::len0 word \ 'a word" +where "z3_rotate_right \ (\n w. word_rotr n w)" + +abbreviation (input) z3_bvult :: "'a::len0 word \ 'a word \ bool" +where "z3_bvult \ (\w1 w2. w1 < w2)" + +abbreviation (input) z3_bvule :: "'a::len0 word \ 'a word \ bool" +where "z3_bvule \ (\w1 w2. w1 \ w2)" + +abbreviation (input) z3_bvugt :: "'a::len0 word \ 'a word \ bool" +where "z3_bvugt \ (\w1 w2. w1 > w2)" + +abbreviation (input) z3_bvuge :: "'a::len0 word \ 'a word \ bool" +where "z3_bvuge \ (\w1 w2. w1 \ w2)" + +abbreviation (input) z3_bvslt :: "'a::len word \ 'a word \ bool" +where "z3_bvslt \ (\w1 w2. w1 'a word \ bool" +where "z3_bvsle \ (\w1 w2. w1 <=s w2)" + +abbreviation (input) z3_bvsgt :: "'a::len word \ 'a word \ bool" +where "z3_bvsgt \ (\w1 w2. w2 'a word \ bool" +where "z3_bvsge \ (\w1 w2. w2 <=s w1)" + + +section {* Boogie environment *} + +text {* +Proving Boogie-generated verification conditions happens inside +a Boogie environment: + + boogie_open "b2i file without 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 = "boogie" + val description = ("Boogie background axioms" ^ + " loaded along with Boogie verification conditions") +) +*} +setup Boogie_Axioms.setup + +text {* +Opening a Boogie environment causes the following list of theorems to be +enhanced by all theorems found in Boogie_Axioms. +*} +ML {* +structure Split_VC_SMT_Rules = Named_Thms +( + val name = "split_vc_smt" + val description = ("Theorems given to the SMT sub-tactic" ^ + " of the split_vc method") +) +*} +setup Split_VC_SMT_Rules.setup + +use "Tools/boogie_vcs.ML" +use "Tools/boogie_loader.ML" +use "Tools/boogie_commands.ML" +setup Boogie_Commands.setup + +use "Tools/boogie_split.ML" +setup Boogie_Split.setup + +end diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/Boogie/Examples/Boogie_Dijkstra.b2i --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Boogie/Examples/Boogie_Dijkstra.b2i Wed Nov 04 11:40:59 2009 +0100 @@ -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 diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/Boogie/Examples/Boogie_Dijkstra.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Boogie/Examples/Boogie_Dijkstra.thy Wed Nov 04 11:40:59 2009 +0100 @@ -0,0 +1,92 @@ +(* 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 "~~/src/HOL/Boogie/Examples/Boogie_Dijkstra" + +boogie_vc b_Dijkstra + unfolding labels + using [[smt_cert="~~/src/HOL/Boogie/Examples/cert/Boogie_b_Dijkstra"]] + using [[z3_proofs=true]] + by (smt boogie) + +boogie_end + +end diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/Boogie/Examples/Boogie_Max.b2i --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Boogie/Examples/Boogie_Max.b2i Wed Nov 04 11:40:59 2009 +0100 @@ -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 diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/Boogie/Examples/Boogie_Max.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Boogie/Examples/Boogie_Max.thy Wed Nov 04 11:40:59 2009 +0100 @@ -0,0 +1,72 @@ +(* Title: HOL/Boogie/Examples/Boogie_Dijkstra.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 "~~/src/HOL/Boogie/Examples/Boogie_Max" + +text {* +Approach 1: Discharge the verification condition fully automatically by SMT: +*} +boogie_vc b_max + unfolding labels + using [[smt_cert="~~/src/HOL/Boogie/Examples/cert/Boogie_b_max"]] + using [[z3_proofs=true]] + by (smt boogie) + +text {* +Approach 2: Split the verification condition, try to solve the splinters by +a selection of automated proof tools, and show the remaining subgoals by an +explicit proof. This approach is most useful in an interactive debug-and-fix +mode. +*} +boogie_vc b_max +proof (split_vc (verbose) try: fast simp) + print_cases + case L_14_5c + thus ?case by (metis abs_of_nonneg zabs_less_one_iff zle_linear) +next + case L_14_5b + thus ?case by (metis less_le_not_le linorder_not_le xt1(10) zle_linear + zless_add1_eq) +next + case L_14_5a + thus ?case by (metis less_le_not_le zle_add1_eq_le zless_linear) +qed + +boogie_end + +end diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/Boogie/Examples/ROOT.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Boogie/Examples/ROOT.ML Wed Nov 04 11:40:59 2009 +0100 @@ -0,0 +1,3 @@ +use_thy "Boogie_Max"; +use_thy "Boogie_Dijkstra"; +use_thy "VCC_Max"; diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/Boogie/Examples/VCC_Max.b2i --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Boogie/Examples/VCC_Max.b2i Wed Nov 04 11:40:59 2009 +0100 @@ -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 diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/Boogie/Examples/VCC_Max.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Boogie/Examples/VCC_Max.thy Wed Nov 04 11:40:59 2009 +0100 @@ -0,0 +1,59 @@ +(* 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 arr[i] <= result)) + ensures (exists(ulong i; i 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) "~~/src/HOL/Boogie/Examples/VCC_Max" + +boogie_status + +boogie_vc b_maximum + unfolding labels + using [[smt_cert="~~/src/HOL/Boogie/Examples/cert/VCC_b_maximum"]] + using [[z3_proofs=true]] + by (smt boogie) + +boogie_end + +end \ No newline at end of file diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/Boogie/Examples/cert/Boogie_b_Dijkstra --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Boogie/Examples/cert/Boogie_b_Dijkstra Wed Nov 04 11:40:59 2009 +0100 @@ -0,0 +1,43 @@ +(benchmark Isabelle +:extrasorts ( T1 T4 T3 T5 T2) +:extrafuns ( + (uf_2 T1 T2) + (uf_3 T1 T2) + (uf_1 T2 T2 T1) + (uf_8 T5) + (uf_7 T4 T2 T5 T4) + (uf_5 T3 T2 Int T3) + (uf_6 T4 T2 T5) + (uf_4 T3 T2 Int) + (uf_10 T1 Int) + (uf_11 T2) + (uf_9 Int) + (uf_21 T2) + (uf_16 T2) + (uf_20 T2) + (uf_12 T2 Int) + (uf_14 T3) + (uf_18 T2 Int) + (uf_22 T3) + (uf_24 T3) + (uf_23 T3) + (uf_15 T4) + (uf_17 T4) + (uf_19 T4) + ) +:extrapreds ( + (up_13 T2) + ) +:assumption (forall (?x1 T1) (= (uf_1 (uf_2 ?x1) (uf_3 ?x1)) ?x1)) +:assumption (forall (?x2 T2) (?x3 T2) (= (uf_3 (uf_1 ?x2 ?x3)) ?x3)) +:assumption (forall (?x4 T2) (?x5 T2) (= (uf_2 (uf_1 ?x4 ?x5)) ?x4)) +:assumption (forall (?x6 T3) (?x7 T2) (?x8 Int) (?x9 T2) (= (uf_4 (uf_5 ?x6 ?x7 ?x8) ?x9) (ite (= ?x9 ?x7) ?x8 (uf_4 ?x6 ?x9)))) +:assumption (forall (?x10 T4) (?x11 T2) (?x12 T5) (?x13 T2) (iff (= (uf_6 (uf_7 ?x10 ?x11 ?x12) ?x13) uf_8) (if_then_else (= ?x13 ?x11) (= ?x12 uf_8) (= (uf_6 ?x10 ?x13) uf_8)))) +:assumption (forall (?x14 T3) (?x15 T2) (?x16 Int) (= (uf_4 (uf_5 ?x14 ?x15 ?x16) ?x15) ?x16)) +:assumption (forall (?x17 T4) (?x18 T2) (?x19 T5) (iff (= (uf_6 (uf_7 ?x17 ?x18 ?x19) ?x18) uf_8) (= ?x19 uf_8))) +:assumption (< 0 uf_9) +:assumption (forall (?x20 T2) (?x21 T2) (implies (= ?x20 ?x21) (= (uf_10 (uf_1 ?x20 ?x21)) 0))) +:assumption (forall (?x22 T2) (?x23 T2) (implies (not (= ?x22 ?x23)) (< 0 (uf_10 (uf_1 ?x22 ?x23))))) +:assumption (not (implies true (implies true (implies (forall (?x24 T2) (implies (= ?x24 uf_11) (= (uf_12 ?x24) 0))) (implies (forall (?x25 T2) (implies (not (= ?x25 uf_11)) (= (uf_12 ?x25) uf_9))) (implies (forall (?x26 T2) (not (up_13 ?x26))) (implies true (and (implies (= (uf_12 uf_11) 0) (and (implies (forall (?x27 T2) (<= 0 (uf_12 ?x27))) (and (implies (forall (?x28 T2) (?x29 T2) (implies (and (up_13 ?x28) (not (up_13 ?x29))) (<= (uf_12 ?x28) (uf_12 ?x29)))) (and (implies (forall (?x30 T2) (?x31 T2) (implies (and (< (uf_10 (uf_1 ?x31 ?x30)) uf_9) (up_13 ?x31)) (<= (uf_12 ?x30) (+ (uf_12 ?x31) (uf_10 (uf_1 ?x31 ?x30)))))) (and (implies (forall (?x32 T2) (implies (and (< (uf_12 ?x32) uf_9) (not (= ?x32 uf_11))) (exists (?x33 T2) (and (= (uf_12 ?x32) (+ (uf_12 ?x33) (uf_10 (uf_1 ?x33 ?x32)))) (and (up_13 ?x33) (< (uf_12 ?x33) (uf_12 ?x32))))))) (implies true (implies true (implies (= (uf_4 uf_14 uf_11) 0) (implies (forall (?x34 T2) (<= 0 (uf_4 uf_14 ?x34))) (implies (forall (?x35 T2) (?x36 T2) (implies (and (= (uf_6 uf_15 ?x35) uf_8) (not (= (uf_6 uf_15 ?x36) uf_8))) (<= (uf_4 uf_14 ?x35) (uf_4 uf_14 ?x36)))) (implies (forall (?x37 T2) (?x38 T2) (implies (and (< (uf_10 (uf_1 ?x38 ?x37)) uf_9) (= (uf_6 uf_15 ?x38) uf_8)) (<= (uf_4 uf_14 ?x37) (+ (uf_4 uf_14 ?x38) (uf_10 (uf_1 ?x38 ?x37)))))) (implies (forall (?x39 T2) (implies (and (< (uf_4 uf_14 ?x39) uf_9) (not (= ?x39 uf_11))) (exists (?x40 T2) (and (= (uf_4 uf_14 ?x39) (+ (uf_4 uf_14 ?x40) (uf_10 (uf_1 ?x40 ?x39)))) (and (= (uf_6 uf_15 ?x40) uf_8) (< (uf_4 uf_14 ?x40) (uf_4 uf_14 ?x39))))))) (implies true (and (implies true (implies true (implies (exists (?x41 T2) (and (< (uf_4 uf_14 ?x41) uf_9) (not (= (uf_6 uf_15 ?x41) uf_8)))) (implies (not (= (uf_6 uf_15 uf_16) uf_8)) (implies (< (uf_4 uf_14 uf_16) uf_9) (implies (forall (?x42 T2) (implies (not (= (uf_6 uf_15 ?x42) uf_8)) (<= (uf_4 uf_14 uf_16) (uf_4 uf_14 ?x42)))) (implies (= uf_17 (uf_7 uf_15 uf_16 uf_8)) (implies (forall (?x43 T2) (implies (and (< (+ (uf_4 uf_14 uf_16) (uf_10 (uf_1 uf_16 ?x43))) (uf_4 uf_14 ?x43)) (< (uf_10 (uf_1 uf_16 ?x43)) uf_9)) (= (uf_18 ?x43) (+ (uf_4 uf_14 uf_16) (uf_10 (uf_1 uf_16 ?x43)))))) (implies (forall (?x44 T2) (implies (not (and (< (+ (uf_4 uf_14 uf_16) (uf_10 (uf_1 uf_16 ?x44))) (uf_4 uf_14 ?x44)) (< (uf_10 (uf_1 uf_16 ?x44)) uf_9))) (= (uf_18 ?x44) (uf_4 uf_14 ?x44)))) (and (implies (forall (?x45 T2) (<= (uf_18 ?x45) (uf_4 uf_14 ?x45))) (and (implies (forall (?x46 T2) (implies (= (uf_6 uf_17 ?x46) uf_8) (= (uf_18 ?x46) (uf_4 uf_14 ?x46)))) (implies true (implies true (and (implies (= (uf_18 uf_11) 0) (and (implies (forall (?x47 T2) (<= 0 (uf_18 ?x47))) (and (implies (forall (?x48 T2) (?x49 T2) (implies (and (= (uf_6 uf_17 ?x48) uf_8) (not (= (uf_6 uf_17 ?x49) uf_8))) (<= (uf_18 ?x48) (uf_18 ?x49)))) (and (implies (forall (?x50 T2) (?x51 T2) (implies (and (< (uf_10 (uf_1 ?x51 ?x50)) uf_9) (= (uf_6 uf_17 ?x51) uf_8)) (<= (uf_18 ?x50) (+ (uf_18 ?x51) (uf_10 (uf_1 ?x51 ?x50)))))) (and (implies (forall (?x52 T2) (implies (and (< (uf_18 ?x52) uf_9) (not (= ?x52 uf_11))) (exists (?x53 T2) (and (= (uf_18 ?x52) (+ (uf_18 ?x53) (uf_10 (uf_1 ?x53 ?x52)))) (and (= (uf_6 uf_17 ?x53) uf_8) (< (uf_18 ?x53) (uf_18 ?x52))))))) (implies false true)) (forall (?x54 T2) (implies (and (< (uf_18 ?x54) uf_9) (not (= ?x54 uf_11))) (exists (?x55 T2) (and (= (uf_18 ?x54) (+ (uf_18 ?x55) (uf_10 (uf_1 ?x55 ?x54)))) (and (= (uf_6 uf_17 ?x55) uf_8) (< (uf_18 ?x55) (uf_18 ?x54))))))))) (forall (?x56 T2) (?x57 T2) (implies (and (< (uf_10 (uf_1 ?x57 ?x56)) uf_9) (= (uf_6 uf_17 ?x57) uf_8)) (<= (uf_18 ?x56) (+ (uf_18 ?x57) (uf_10 (uf_1 ?x57 ?x56)))))))) (forall (?x58 T2) (?x59 T2) (implies (and (= (uf_6 uf_17 ?x58) uf_8) (not (= (uf_6 uf_17 ?x59) uf_8))) (<= (uf_18 ?x58) (uf_18 ?x59)))))) (forall (?x60 T2) (<= 0 (uf_18 ?x60))))) (= (uf_18 uf_11) 0))))) (forall (?x61 T2) (implies (= (uf_6 uf_17 ?x61) uf_8) (= (uf_18 ?x61) (uf_4 uf_14 ?x61)))))) (forall (?x62 T2) (<= (uf_18 ?x62) (uf_4 uf_14 ?x62))))))))))))) (implies true (implies true (implies (not (exists (?x63 T2) (and (< (uf_4 uf_14 ?x63) uf_9) (not (= (uf_6 uf_15 ?x63) uf_8))))) (implies true (implies true (implies (= uf_19 uf_15) (implies (= uf_20 uf_21) (implies (= uf_22 uf_14) (implies (= uf_23 uf_24) (implies true (and (implies (forall (?x64 T2) (implies (and (< (uf_4 uf_22 ?x64) uf_9) (not (= ?x64 uf_11))) (exists (?x65 T2) (and (= (uf_4 uf_22 ?x64) (+ (uf_4 uf_22 ?x65) (uf_10 (uf_1 ?x65 ?x64)))) (< (uf_4 uf_22 ?x65) (uf_4 uf_22 ?x64)))))) (and (implies (forall (?x66 T2) (?x67 T2) (implies (and (< (uf_10 (uf_1 ?x67 ?x66)) uf_9) (< (uf_4 uf_22 ?x67) uf_9)) (<= (uf_4 uf_22 ?x66) (+ (uf_4 uf_22 ?x67) (uf_10 (uf_1 ?x67 ?x66)))))) (and (implies (= (uf_4 uf_22 uf_11) 0) true) (= (uf_4 uf_22 uf_11) 0))) (forall (?x68 T2) (?x69 T2) (implies (and (< (uf_10 (uf_1 ?x69 ?x68)) uf_9) (< (uf_4 uf_22 ?x69) uf_9)) (<= (uf_4 uf_22 ?x68) (+ (uf_4 uf_22 ?x69) (uf_10 (uf_1 ?x69 ?x68)))))))) (forall (?x70 T2) (implies (and (< (uf_4 uf_22 ?x70) uf_9) (not (= ?x70 uf_11))) (exists (?x71 T2) (and (= (uf_4 uf_22 ?x70) (+ (uf_4 uf_22 ?x71) (uf_10 (uf_1 ?x71 ?x70)))) (< (uf_4 uf_22 ?x71) (uf_4 uf_22 ?x70))))))))))))))))))))))))))) (forall (?x72 T2) (implies (and (< (uf_12 ?x72) uf_9) (not (= ?x72 uf_11))) (exists (?x73 T2) (and (= (uf_12 ?x72) (+ (uf_12 ?x73) (uf_10 (uf_1 ?x73 ?x72)))) (and (up_13 ?x73) (< (uf_12 ?x73) (uf_12 ?x72))))))))) (forall (?x74 T2) (?x75 T2) (implies (and (< (uf_10 (uf_1 ?x75 ?x74)) uf_9) (up_13 ?x75)) (<= (uf_12 ?x74) (+ (uf_12 ?x75) (uf_10 (uf_1 ?x75 ?x74)))))))) (forall (?x76 T2) (?x77 T2) (implies (and (up_13 ?x76) (not (up_13 ?x77))) (<= (uf_12 ?x76) (uf_12 ?x77)))))) (forall (?x78 T2) (<= 0 (uf_12 ?x78))))) (= (uf_12 uf_11) 0))))))))) +:formula true +) diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/Boogie/Examples/cert/Boogie_b_Dijkstra.proof --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Boogie/Examples/cert/Boogie_b_Dijkstra.proof Wed Nov 04 11:40:59 2009 +0100 @@ -0,0 +1,7081 @@ +#2 := false +#55 := 0::int +decl uf_4 :: (-> T3 T2 int) +decl ?x40!7 :: (-> T2 T2) +decl ?x52!15 :: T2 +#2305 := ?x52!15 +#15992 := (?x40!7 ?x52!15) +decl uf_14 :: T3 +#107 := uf_14 +#15996 := (uf_4 uf_14 #15992) +#20405 := (>= #15996 0::int) +#11 := (:var 0 T2) +#110 := (uf_4 uf_14 #11) +#4403 := (pattern #110) +#1843 := (>= #110 0::int) +#4404 := (forall (vars (?x34 T2)) (:pat #4403) #1843) +decl uf_10 :: (-> T1 int) +decl uf_1 :: (-> T2 T2 T1) +decl ?x66!20 :: T2 +#2511 := ?x66!20 +decl ?x67!19 :: T2 +#2510 := ?x67!19 +#2516 := (uf_1 ?x67!19 ?x66!20) +#2517 := (uf_10 #2516) +#1320 := -1::int +#2524 := (* -1::int #2517) +decl uf_22 :: T3 +#230 := uf_22 +#2514 := (uf_4 uf_22 ?x67!19) +#2520 := (* -1::int #2514) +#3094 := (+ #2520 #2524) +#2512 := (uf_4 uf_22 ?x66!20) +#3095 := (+ #2512 #3094) +#3096 := (<= #3095 0::int) +decl uf_9 :: int +#56 := uf_9 +#2525 := (+ uf_9 #2524) +#2526 := (<= #2525 0::int) +#2521 := (+ uf_9 #2520) +#2522 := (<= #2521 0::int) +#3693 := (or #2522 #2526 #3096) +#3698 := (not #3693) +#10 := (:var 1 T2) +#90 := (uf_1 #11 #10) +#4379 := (pattern #90) +#238 := (uf_4 uf_22 #10) +#1720 := (* -1::int #238) +#235 := (uf_4 uf_22 #11) +#1721 := (+ #235 #1720) +#91 := (uf_10 #90) +#1727 := (+ #91 #1721) +#1750 := (>= #1727 0::int) +#1707 := (* -1::int #235) +#1708 := (+ uf_9 #1707) +#1709 := (<= #1708 0::int) +#1343 := (* -1::int #91) +#1346 := (+ uf_9 #1343) +#1347 := (<= #1346 0::int) +#3661 := (or #1347 #1709 #1750) +#4633 := (forall (vars (?x66 T2) (?x67 T2)) (:pat #4379) #3661) +#4638 := (not #4633) +decl uf_11 :: T2 +#67 := uf_11 +#250 := (uf_4 uf_22 uf_11) +#251 := (= #250 0::int) +#4641 := (or #251 #4638) +#4644 := (not #4641) +#4647 := (or #4644 #3698) +#4650 := (not #4647) +#4609 := (pattern #235) +decl ?x65!18 :: (-> T2 T2) +#2487 := (?x65!18 #11) +#2490 := (uf_1 #2487 #11) +#2491 := (uf_10 #2490) +#3064 := (* -1::int #2491) +#2488 := (uf_4 uf_22 #2487) +#3047 := (* -1::int #2488) +#3065 := (+ #3047 #3064) +#3066 := (+ #235 #3065) +#3067 := (= #3066 0::int) +#3631 := (not #3067) +#3048 := (+ #235 #3047) +#3049 := (<= #3048 0::int) +#3632 := (or #3049 #3631) +#3633 := (not #3632) +#68 := (= #11 uf_11) +#3639 := (or #68 #1709 #3633) +#4625 := (forall (vars (?x64 T2)) (:pat #4609) #3639) +#4630 := (not #4625) +#4653 := (or #4630 #4650) +#4656 := (not #4653) +decl ?x64!17 :: T2 +#2447 := ?x64!17 +#2451 := (uf_1 #11 ?x64!17) +#4610 := (pattern #2451) +#2452 := (uf_10 #2451) +#2448 := (uf_4 uf_22 ?x64!17) +#2449 := (* -1::int #2448) +#3017 := (+ #2449 #2452) +#3018 := (+ #235 #3017) +#3021 := (= #3018 0::int) +#3595 := (not #3021) +#2450 := (+ #235 #2449) +#2455 := (>= #2450 0::int) +#3596 := (or #2455 #3595) +#4611 := (forall (vars (?x65 T2)) (:pat #4609 #4610) #3596) +#4616 := (not #4611) +#2993 := (= uf_11 ?x64!17) +#2459 := (+ uf_9 #2449) +#2460 := (<= #2459 0::int) +#4619 := (or #2460 #2993 #4616) +#4622 := (not #4619) +#4659 := (or #4622 #4656) +#4662 := (not #4659) +decl uf_6 :: (-> T4 T2 T5) +decl uf_15 :: T4 +#113 := uf_15 +#116 := (uf_6 uf_15 #11) +#4445 := (pattern #116) +#1404 := (* -1::int #110) +#1405 := (+ uf_9 #1404) +#1406 := (<= #1405 0::int) +decl uf_8 :: T5 +#33 := uf_8 +#505 := (= uf_8 #116) +#3581 := (or #505 #1406) +#4601 := (forall (vars (?x41 T2)) (:pat #4445 #4403) #3581) +#4606 := (not #4601) +#933 := (= uf_14 uf_22) +#1053 := (not #933) +decl uf_19 :: T4 +#225 := uf_19 +#930 := (= uf_15 uf_19) +#1071 := (not #930) +decl uf_24 :: T3 +#233 := uf_24 +decl uf_23 :: T3 +#232 := uf_23 +#234 := (= uf_23 uf_24) +#1044 := (not #234) +decl uf_21 :: T2 +#228 := uf_21 +decl uf_20 :: T2 +#227 := uf_20 +#229 := (= uf_20 uf_21) +#1062 := (not #229) +#4665 := (or #1062 #1044 #1071 #1053 #4606 #4662) +#4668 := (not #4665) +#2309 := (uf_1 #11 ?x52!15) +#4514 := (pattern #2309) +decl uf_18 :: (-> T2 int) +#158 := (uf_18 #11) +#4454 := (pattern #158) +decl uf_17 :: T4 +#149 := uf_17 +#168 := (uf_6 uf_17 #11) +#4480 := (pattern #168) +#2310 := (uf_10 #2309) +#2306 := (uf_18 ?x52!15) +#2307 := (* -1::int #2306) +#2917 := (+ #2307 #2310) +#2918 := (+ #158 #2917) +#2921 := (= #2918 0::int) +#3474 := (not #2921) +#2308 := (+ #158 #2307) +#2313 := (>= #2308 0::int) +#630 := (= uf_8 #168) +#636 := (not #630) +#3475 := (or #636 #2313 #3474) +#4515 := (forall (vars (?x53 T2)) (:pat #4480 #4454 #4514) #3475) +#4520 := (not #4515) +#180 := (uf_18 #10) +#1505 := (* -1::int #180) +#1506 := (+ #158 #1505) +#1536 := (+ #91 #1506) +#1534 := (>= #1536 0::int) +#3466 := (or #636 #1347 #1534) +#4506 := (forall (vars (?x50 T2) (?x51 T2)) (:pat #4379) #3466) +#4511 := (not #4506) +#2893 := (= uf_11 ?x52!15) +#2317 := (+ uf_9 #2307) +#2318 := (<= #2317 0::int) +#4523 := (or #2318 #2893 #4511 #4520) +#4526 := (not #4523) +decl ?x50!14 :: T2 +#2275 := ?x50!14 +decl ?x51!13 :: T2 +#2274 := ?x51!13 +#2280 := (uf_1 ?x51!13 ?x50!14) +#2281 := (uf_10 #2280) +#2284 := (* -1::int #2281) +#2278 := (uf_18 ?x51!13) +#2879 := (* -1::int #2278) +#2880 := (+ #2879 #2284) +#2276 := (uf_18 ?x50!14) +#2881 := (+ #2276 #2880) +#2882 := (<= #2881 0::int) +#2288 := (uf_6 uf_17 ?x51!13) +#2289 := (= uf_8 #2288) +#3429 := (not #2289) +#2285 := (+ uf_9 #2284) +#2286 := (<= #2285 0::int) +#3444 := (or #2286 #3429 #2882) +#3449 := (not #3444) +#4529 := (or #3449 #4526) +#4532 := (not #4529) +#4497 := (pattern #158 #180) +#1504 := (>= #1506 0::int) +#176 := (uf_6 uf_17 #10) +#648 := (= uf_8 #176) +#3406 := (not #648) +#3421 := (or #630 #3406 #1504) +#4498 := (forall (vars (?x48 T2) (?x49 T2)) (:pat #4497) #3421) +#4503 := (not #4498) +#4535 := (or #4503 #4532) +#4538 := (not #4535) +decl ?x49!11 :: T2 +#2247 := ?x49!11 +#2251 := (uf_18 ?x49!11) +#2853 := (* -1::int #2251) +decl ?x48!12 :: T2 +#2248 := ?x48!12 +#2249 := (uf_18 ?x48!12) +#2854 := (+ #2249 #2853) +#2855 := (<= #2854 0::int) +#2256 := (uf_6 uf_17 ?x49!11) +#2257 := (= uf_8 #2256) +#2254 := (uf_6 uf_17 ?x48!12) +#2255 := (= uf_8 #2254) +#3383 := (not #2255) +#3398 := (or #3383 #2257 #2855) +#3403 := (not #3398) +#4541 := (or #3403 #4538) +#4544 := (not #4541) +#1495 := (>= #158 0::int) +#4489 := (forall (vars (?x47 T2)) (:pat #4454) #1495) +#4494 := (not #4489) +#4547 := (or #4494 #4544) +#4550 := (not #4547) +decl ?x47!10 :: T2 +#2232 := ?x47!10 +#2233 := (uf_18 ?x47!10) +#2234 := (>= #2233 0::int) +#2235 := (not #2234) +#4553 := (or #2235 #4550) +#4556 := (not #4553) +#172 := (uf_18 uf_11) +#173 := (= #172 0::int) +#1492 := (not #173) +#4559 := (or #1492 #4556) +#4562 := (not #4559) +#4565 := (or #1492 #4562) +#4568 := (not #4565) +#616 := (= #110 #158) +#637 := (or #616 #636) +#4481 := (forall (vars (?x46 T2)) (:pat #4403 #4454 #4480) #637) +#4486 := (not #4481) +#4571 := (or #4486 #4568) +#4574 := (not #4571) +decl ?x46!9 :: T2 +#2207 := ?x46!9 +#2212 := (uf_4 uf_14 ?x46!9) +#2211 := (uf_18 ?x46!9) +#2825 := (= #2211 #2212) +#2208 := (uf_6 uf_17 ?x46!9) +#2209 := (= uf_8 #2208) +#2210 := (not #2209) +#2831 := (or #2210 #2825) +#2836 := (not #2831) +#4577 := (or #2836 #4574) +#4580 := (not #4577) +#1480 := (* -1::int #158) +#1481 := (+ #110 #1480) +#1479 := (>= #1481 0::int) +#4472 := (forall (vars (?x45 T2)) (:pat #4403 #4454) #1479) +#4477 := (not #4472) +#4583 := (or #4477 #4580) +#4586 := (not #4583) +decl ?x45!8 :: T2 +#2189 := ?x45!8 +#2192 := (uf_4 uf_14 ?x45!8) +#2815 := (* -1::int #2192) +#2190 := (uf_18 ?x45!8) +#2816 := (+ #2190 #2815) +#2817 := (<= #2816 0::int) +#2822 := (not #2817) +#4589 := (or #2822 #4586) +#4592 := (not #4589) +decl uf_16 :: T2 +#140 := uf_16 +#152 := (uf_1 uf_16 #11) +#4455 := (pattern #152) +#153 := (uf_10 #152) +#1623 := (+ #153 #1480) +#144 := (uf_4 uf_14 uf_16) +#1624 := (+ #144 #1623) +#1625 := (= #1624 0::int) +#1450 := (* -1::int #153) +#1457 := (+ uf_9 #1450) +#1458 := (<= #1457 0::int) +#1449 := (* -1::int #144) +#1451 := (+ #1449 #1450) +#1452 := (+ #110 #1451) +#1453 := (<= #1452 0::int) +#3375 := (or #1453 #1458 #1625) +#4464 := (forall (vars (?x43 T2)) (:pat #4403 #4455 #4454) #3375) +#4469 := (not #4464) +#3355 := (or #1453 #1458) +#3356 := (not #3355) +#3359 := (or #616 #3356) +#4456 := (forall (vars (?x44 T2)) (:pat #4403 #4454 #4455) #3359) +#4461 := (not #4456) +decl ?x41!16 :: T2 +#2408 := ?x41!16 +#2414 := (uf_6 uf_15 ?x41!16) +#2415 := (= uf_8 #2414) +#2409 := (uf_4 uf_14 ?x41!16) +#2410 := (* -1::int #2409) +#2411 := (+ uf_9 #2410) +#2412 := (<= #2411 0::int) +#1655 := (+ uf_9 #1449) +#1656 := (<= #1655 0::int) +#1638 := (+ #110 #1449) +#1637 := (>= #1638 0::int) +#1644 := (or #505 #1637) +#4446 := (forall (vars (?x42 T2)) (:pat #4445 #4403) #1644) +#4451 := (not #4446) +#141 := (uf_6 uf_15 uf_16) +#585 := (= uf_8 #141) +decl uf_7 :: (-> T4 T2 T5 T4) +#150 := (uf_7 uf_15 uf_16 uf_8) +#151 := (= uf_17 #150) +#876 := (not #151) +#4595 := (or #876 #585 #4451 #1656 #2412 #2415 #4461 #4469 #4592) +#4598 := (not #4595) +#4671 := (or #4598 #4668) +#4674 := (not #4671) +#2152 := (?x40!7 #11) +#2155 := (uf_1 #2152 #11) +#2156 := (uf_10 #2155) +#2790 := (* -1::int #2156) +#2153 := (uf_4 uf_14 #2152) +#2773 := (* -1::int #2153) +#2791 := (+ #2773 #2790) +#2792 := (+ #110 #2791) +#2793 := (= #2792 0::int) +#3339 := (not #2793) +#2774 := (+ #110 #2773) +#2775 := (<= #2774 0::int) +#2161 := (uf_6 uf_15 #2152) +#2162 := (= uf_8 #2161) +#3338 := (not #2162) +#3340 := (or #3338 #2775 #3339) +#3341 := (not #3340) +#3347 := (or #68 #1406 #3341) +#4437 := (forall (vars (?x39 T2)) (:pat #4403) #3347) +#4442 := (not #4437) +decl uf_12 :: (-> T2 int) +#69 := (uf_12 #11) +#4355 := (pattern #69) +decl ?x33!6 :: (-> T2 T2) +#2123 := (?x33!6 #11) +#2127 := (uf_12 #2123) +#2728 := (* -1::int #2127) +#2124 := (uf_1 #2123 #11) +#2125 := (uf_10 #2124) +#2745 := (* -1::int #2125) +#2746 := (+ #2745 #2728) +#2747 := (+ #69 #2746) +#2748 := (= #2747 0::int) +#3311 := (not #2748) +#2729 := (+ #69 #2728) +#2730 := (<= #2729 0::int) +decl up_13 :: (-> T2 bool) +#2133 := (up_13 #2123) +#3310 := (not #2133) +#3312 := (or #3310 #2730 #3311) +#3313 := (not #3312) +#1386 := (* -1::int #69) +#1387 := (+ uf_9 #1386) +#1388 := (<= #1387 0::int) +#3319 := (or #68 #1388 #3313) +#4429 := (forall (vars (?x32 T2)) (:pat #4355) #3319) +#4434 := (not #4429) +#114 := (uf_6 uf_15 #10) +#4420 := (pattern #114 #116) +#120 := (uf_4 uf_14 #10) +#1417 := (* -1::int #120) +#1418 := (+ #110 #1417) +#1416 := (>= #1418 0::int) +#502 := (= uf_8 #114) +#3276 := (not #502) +#3291 := (or #3276 #505 #1416) +#4421 := (forall (vars (?x35 T2) (?x36 T2)) (:pat #4420) #3291) +#4426 := (not #4421) +#1424 := (+ #91 #1418) +#1815 := (>= #1424 0::int) +#508 := (not #505) +#3268 := (or #508 #1347 #1815) +#4412 := (forall (vars (?x37 T2) (?x38 T2)) (:pat #4379) #3268) +#4417 := (not #4412) +#4409 := (not #4404) +#108 := (uf_4 uf_14 uf_11) +#109 := (= #108 0::int) +#1854 := (not #109) +#4677 := (or #1854 #4409 #4417 #4426 #4434 #4442 #4674) +#4680 := (not #4677) +decl ?x32!5 :: T2 +#2081 := ?x32!5 +#2091 := (uf_1 #11 ?x32!5) +#4388 := (pattern #2091) +#77 := (up_13 #11) +#4348 := (pattern #77) +#2082 := (uf_12 ?x32!5) +#2083 := (* -1::int #2082) +#2096 := (+ #69 #2083) +#2097 := (>= #2096 0::int) +#2092 := (uf_10 #2091) +#2093 := (+ #2083 #2092) +#2094 := (+ #69 #2093) +#2095 := (= #2094 0::int) +#3229 := (not #2095) +#78 := (not #77) +#3230 := (or #78 #3229 #2097) +#4389 := (forall (vars (?x33 T2)) (:pat #4348 #4355 #4388) #3230) +#4394 := (not #4389) +#2688 := (= uf_11 ?x32!5) +#2084 := (+ uf_9 #2083) +#2085 := (<= #2084 0::int) +#4397 := (or #2085 #2688 #4394) +#4400 := (not #4397) +#4683 := (or #4400 #4680) +#4686 := (not #4683) +#86 := (uf_12 #10) +#1323 := (* -1::int #86) +#1344 := (+ #1323 #91) +#1345 := (+ #69 #1344) +#1342 := (>= #1345 0::int) +#3221 := (or #78 #1342 #1347) +#4380 := (forall (vars (?x30 T2) (?x31 T2)) (:pat #4379) #3221) +#4385 := (not #4380) +#4689 := (or #4385 #4686) +#4692 := (not #4689) +decl ?x31!3 :: T2 +#2051 := ?x31!3 +#2065 := (uf_12 ?x31!3) +decl ?x30!4 :: T2 +#2052 := ?x30!4 +#2062 := (uf_12 ?x30!4) +#2063 := (* -1::int #2062) +#2660 := (+ #2063 #2065) +#2053 := (uf_1 ?x31!3 ?x30!4) +#2054 := (uf_10 #2053) +#2661 := (+ #2054 #2660) +#2664 := (>= #2661 0::int) +#2059 := (up_13 ?x31!3) +#3184 := (not #2059) +#2055 := (* -1::int #2054) +#2056 := (+ uf_9 #2055) +#2057 := (<= #2056 0::int) +#3199 := (or #2057 #3184 #2664) +#3204 := (not #3199) +#4695 := (or #3204 #4692) +#4698 := (not #4695) +#84 := (up_13 #10) +#4370 := (pattern #77 #84) +#1324 := (+ #69 #1323) +#1322 := (>= #1324 0::int) +#2632 := (not #84) +#3176 := (or #77 #2632 #1322) +#4371 := (forall (vars (?x28 T2) (?x29 T2)) (:pat #4370) #3176) +#4376 := (not #4371) +#4701 := (or #4376 #4698) +#4704 := (not #4701) +decl ?x29!1 :: T2 +#2026 := ?x29!1 +#2030 := (uf_12 ?x29!1) +#2647 := (* -1::int #2030) +decl ?x28!2 :: T2 +#2027 := ?x28!2 +#2028 := (uf_12 ?x28!2) +#2648 := (+ #2028 #2647) +#2649 := (<= #2648 0::int) +#2034 := (up_13 ?x29!1) +#2033 := (up_13 ?x28!2) +#2266 := (not #2033) +#2166 := (or #2266 #2034 #2649) +#6004 := [hypothesis]: #2033 +#4349 := (forall (vars (?x26 T2)) (:pat #4348) #78) +#79 := (forall (vars (?x26 T2)) #78) +#4352 := (iff #79 #4349) +#4350 := (iff #78 #78) +#4351 := [refl]: #4350 +#4353 := [quant-intro #4351]: #4352 +#1965 := (~ #79 #79) +#2002 := (~ #78 #78) +#2003 := [refl]: #2002 +#1966 := [nnf-pos #2003]: #1965 +#70 := (= #69 0::int) +#73 := (not #68) +#1912 := (or #73 #70) +#1915 := (forall (vars (?x24 T2)) #1912) +#1918 := (not #1915) +#1846 := (forall (vars (?x34 T2)) #1843) +#1849 := (not #1846) +#511 := (and #502 #508) +#517 := (not #511) +#1832 := (or #517 #1416) +#1837 := (forall (vars (?x35 T2) (?x36 T2)) #1832) +#1840 := (not #1837) +#1348 := (not #1347) +#1807 := (and #505 #1348) +#1812 := (not #1807) +#1818 := (or #1812 #1815) +#1821 := (forall (vars (?x37 T2) (?x38 T2)) #1818) +#1824 := (not #1821) +#1710 := (not #1709) +#1744 := (and #1348 #1710) +#1747 := (not #1744) +#1753 := (or #1747 #1750) +#1756 := (forall (vars (?x66 T2) (?x67 T2)) #1753) +#1759 := (not #1756) +#1767 := (or #251 #1759) +#1772 := (and #1756 #1767) +#1725 := (= #1727 0::int) +#1719 := (>= #1721 0::int) +#1722 := (not #1719) +#1729 := (and #1722 #1725) +#1732 := (exists (vars (?x65 T2)) #1729) +#1713 := (and #73 #1710) +#1716 := (not #1713) +#1735 := (or #1716 #1732) +#1738 := (forall (vars (?x64 T2)) #1735) +#1741 := (not #1738) +#1775 := (or #1741 #1772) +#1778 := (and #1738 #1775) +#1407 := (not #1406) +#1670 := (and #508 #1407) +#1675 := (exists (vars (?x41 T2)) #1670) +#1796 := (or #1062 #1044 #1071 #1053 #1675 #1778) +#1678 := (not #1675) +#1649 := (forall (vars (?x42 T2)) #1644) +#1652 := (not #1649) +#1459 := (not #1458) +#1454 := (not #1453) +#1462 := (and #1454 #1459) +#1620 := (not #1462) +#1628 := (or #1620 #1625) +#1631 := (forall (vars (?x43 T2)) #1628) +#1634 := (not #1631) +#1561 := (= #1536 0::int) +#1558 := (not #1504) +#1570 := (and #630 #1558 #1561) +#1575 := (exists (vars (?x53 T2)) #1570) +#1547 := (+ uf_9 #1480) +#1548 := (<= #1547 0::int) +#1549 := (not #1548) +#1552 := (and #73 #1549) +#1555 := (not #1552) +#1578 := (or #1555 #1575) +#1581 := (forall (vars (?x52 T2)) #1578) +#1526 := (and #630 #1348) +#1531 := (not #1526) +#1538 := (or #1531 #1534) +#1541 := (forall (vars (?x50 T2) (?x51 T2)) #1538) +#1544 := (not #1541) +#1584 := (or #1544 #1581) +#1587 := (and #1541 #1584) +#656 := (and #636 #648) +#664 := (not #656) +#1512 := (or #664 #1504) +#1517 := (forall (vars (?x48 T2) (?x49 T2)) #1512) +#1520 := (not #1517) +#1590 := (or #1520 #1587) +#1593 := (and #1517 #1590) +#1498 := (forall (vars (?x47 T2)) #1495) +#1501 := (not #1498) +#1596 := (or #1501 #1593) +#1599 := (and #1498 #1596) +#1602 := (or #1492 #1599) +#1605 := (and #173 #1602) +#642 := (forall (vars (?x46 T2)) #637) +#824 := (not #642) +#1608 := (or #824 #1605) +#1611 := (and #642 #1608) +#1484 := (forall (vars (?x45 T2)) #1479) +#1487 := (not #1484) +#1614 := (or #1487 #1611) +#1617 := (and #1484 #1614) +#1468 := (or #616 #1462) +#1473 := (forall (vars (?x44 T2)) #1468) +#1476 := (not #1473) +#1702 := (or #876 #585 #1476 #1617 #1634 #1652 #1656 #1678) +#1801 := (and #1702 #1796) +#1422 := (= #1424 0::int) +#1419 := (not #1416) +#1432 := (and #505 #1419 #1422) +#1437 := (exists (vars (?x40 T2)) #1432) +#1410 := (and #73 #1407) +#1413 := (not #1410) +#1440 := (or #1413 #1437) +#1443 := (forall (vars (?x39 T2)) #1440) +#1446 := (not #1443) +#1389 := (not #1388) +#1392 := (and #73 #1389) +#1395 := (not #1392) +#1370 := (= #1345 0::int) +#1366 := (not #1322) +#1378 := (and #77 #1366 #1370) +#1383 := (exists (vars (?x33 T2)) #1378) +#1398 := (or #1383 #1395) +#1401 := (forall (vars (?x32 T2)) #1398) +#1857 := (not #1401) +#1878 := (or #1854 #1857 #1446 #1801 #1824 #1840 #1849) +#1883 := (and #1401 #1878) +#1351 := (and #77 #1348) +#1354 := (not #1351) +#1357 := (or #1342 #1354) +#1360 := (forall (vars (?x30 T2) (?x31 T2)) #1357) +#1363 := (not #1360) +#1886 := (or #1363 #1883) +#1889 := (and #1360 #1886) +#454 := (and #78 #84) +#460 := (not #454) +#1329 := (or #460 #1322) +#1334 := (forall (vars (?x28 T2) (?x29 T2)) #1329) +#1337 := (not #1334) +#1892 := (or #1337 #1889) +#1895 := (and #1334 #1892) +#1313 := (>= #69 0::int) +#1314 := (forall (vars (?x27 T2)) #1313) +#1317 := (not #1314) +#1898 := (or #1317 #1895) +#1901 := (and #1314 #1898) +#80 := (uf_12 uf_11) +#81 := (= #80 0::int) +#1308 := (not #81) +#1904 := (or #1308 #1901) +#1907 := (and #81 #1904) +#437 := (= uf_9 #69) +#443 := (or #68 #437) +#448 := (forall (vars (?x25 T2)) #443) +#1277 := (not #448) +#1268 := (not #79) +#1930 := (or #1268 #1277 #1907 #1918) +#1935 := (not #1930) +#82 := (<= 0::int #69) +#83 := (forall (vars (?x27 T2)) #82) +#87 := (<= #86 #69) +#85 := (and #84 #78) +#88 := (implies #85 #87) +#89 := (forall (vars (?x28 T2) (?x29 T2)) #88) +#94 := (+ #69 #91) +#95 := (<= #86 #94) +#92 := (< #91 uf_9) +#93 := (and #92 #77) +#96 := (implies #93 #95) +#97 := (forall (vars (?x30 T2) (?x31 T2)) #96) +#101 := (< #69 #86) +#102 := (and #77 #101) +#100 := (= #86 #94) +#103 := (and #100 #102) +#104 := (exists (vars (?x33 T2)) #103) +#98 := (< #69 uf_9) +#99 := (and #98 #73) +#105 := (implies #99 #104) +#106 := (forall (vars (?x32 T2)) #105) +#241 := (< #235 #238) +#239 := (+ #235 #91) +#240 := (= #238 #239) +#242 := (and #240 #241) +#243 := (exists (vars (?x65 T2)) #242) +#236 := (< #235 uf_9) +#237 := (and #236 #73) +#244 := (implies #237 #243) +#245 := (forall (vars (?x64 T2)) #244) +#247 := (<= #238 #239) +#246 := (and #92 #236) +#248 := (implies #246 #247) +#249 := (forall (vars (?x66 T2) (?x67 T2)) #248) +#1 := true +#252 := (implies #251 true) +#253 := (and #252 #251) +#254 := (implies #249 #253) +#255 := (and #254 #249) +#256 := (implies #245 #255) +#257 := (and #256 #245) +#258 := (implies true #257) +#259 := (implies #234 #258) +#231 := (= uf_22 uf_14) +#260 := (implies #231 #259) +#261 := (implies #229 #260) +#226 := (= uf_19 uf_15) +#262 := (implies #226 #261) +#263 := (implies true #262) +#264 := (implies true #263) +#117 := (= #116 uf_8) +#118 := (not #117) +#129 := (< #110 uf_9) +#138 := (and #129 #118) +#139 := (exists (vars (?x41 T2)) #138) +#224 := (not #139) +#265 := (implies #224 #264) +#266 := (implies true #265) +#267 := (implies true #266) +#166 := (<= #158 #110) +#167 := (forall (vars (?x45 T2)) #166) +#163 := (= #158 #110) +#169 := (= #168 uf_8) +#170 := (implies #169 #163) +#171 := (forall (vars (?x46 T2)) #170) +#174 := (<= 0::int #158) +#175 := (forall (vars (?x47 T2)) #174) +#181 := (<= #180 #158) +#178 := (not #169) +#177 := (= #176 uf_8) +#179 := (and #177 #178) +#182 := (implies #179 #181) +#183 := (forall (vars (?x48 T2) (?x49 T2)) #182) +#185 := (+ #158 #91) +#186 := (<= #180 #185) +#184 := (and #92 #169) +#187 := (implies #184 #186) +#188 := (forall (vars (?x50 T2) (?x51 T2)) #187) +#192 := (< #158 #180) +#193 := (and #169 #192) +#191 := (= #180 #185) +#194 := (and #191 #193) +#195 := (exists (vars (?x53 T2)) #194) +#189 := (< #158 uf_9) +#190 := (and #189 #73) +#196 := (implies #190 #195) +#197 := (forall (vars (?x52 T2)) #196) +#198 := (implies false true) +#199 := (implies #197 #198) +#200 := (and #199 #197) +#201 := (implies #188 #200) +#202 := (and #201 #188) +#203 := (implies #183 #202) +#204 := (and #203 #183) +#205 := (implies #175 #204) +#206 := (and #205 #175) +#207 := (implies #173 #206) +#208 := (and #207 #173) +#209 := (implies true #208) +#210 := (implies true #209) +#211 := (implies #171 #210) +#212 := (and #211 #171) +#213 := (implies #167 #212) +#214 := (and #213 #167) +#156 := (< #153 uf_9) +#154 := (+ #144 #153) +#155 := (< #154 #110) +#157 := (and #155 #156) +#162 := (not #157) +#164 := (implies #162 #163) +#165 := (forall (vars (?x44 T2)) #164) +#215 := (implies #165 #214) +#159 := (= #158 #154) +#160 := (implies #157 #159) +#161 := (forall (vars (?x43 T2)) #160) +#216 := (implies #161 #215) +#217 := (implies #151 #216) +#146 := (<= #144 #110) +#147 := (implies #118 #146) +#148 := (forall (vars (?x42 T2)) #147) +#218 := (implies #148 #217) +#145 := (< #144 uf_9) +#219 := (implies #145 #218) +#142 := (= #141 uf_8) +#143 := (not #142) +#220 := (implies #143 #219) +#221 := (implies #139 #220) +#222 := (implies true #221) +#223 := (implies true #222) +#268 := (and #223 #267) +#269 := (implies true #268) +#132 := (< #110 #120) +#133 := (and #117 #132) +#125 := (+ #110 #91) +#131 := (= #120 #125) +#134 := (and #131 #133) +#135 := (exists (vars (?x40 T2)) #134) +#130 := (and #129 #73) +#136 := (implies #130 #135) +#137 := (forall (vars (?x39 T2)) #136) +#270 := (implies #137 #269) +#126 := (<= #120 #125) +#124 := (and #92 #117) +#127 := (implies #124 #126) +#128 := (forall (vars (?x37 T2) (?x38 T2)) #127) +#271 := (implies #128 #270) +#121 := (<= #120 #110) +#115 := (= #114 uf_8) +#119 := (and #115 #118) +#122 := (implies #119 #121) +#123 := (forall (vars (?x35 T2) (?x36 T2)) #122) +#272 := (implies #123 #271) +#111 := (<= 0::int #110) +#112 := (forall (vars (?x34 T2)) #111) +#273 := (implies #112 #272) +#274 := (implies #109 #273) +#275 := (implies true #274) +#276 := (implies true #275) +#277 := (implies #106 #276) +#278 := (and #277 #106) +#279 := (implies #97 #278) +#280 := (and #279 #97) +#281 := (implies #89 #280) +#282 := (and #281 #89) +#283 := (implies #83 #282) +#284 := (and #283 #83) +#285 := (implies #81 #284) +#286 := (and #285 #81) +#287 := (implies true #286) +#288 := (implies #79 #287) +#74 := (= #69 uf_9) +#75 := (implies #73 #74) +#76 := (forall (vars (?x25 T2)) #75) +#289 := (implies #76 #288) +#71 := (implies #68 #70) +#72 := (forall (vars (?x24 T2)) #71) +#290 := (implies #72 #289) +#291 := (implies true #290) +#292 := (implies true #291) +#293 := (not #292) +#1938 := (iff #293 #1935) +#983 := (= 0::int #250) +#939 := (+ #91 #235) +#968 := (<= #238 #939) +#974 := (not #246) +#975 := (or #974 #968) +#980 := (forall (vars (?x66 T2) (?x67 T2)) #975) +#1003 := (not #980) +#1004 := (or #1003 #983) +#1012 := (and #980 #1004) +#942 := (= #238 #939) +#948 := (and #241 #942) +#953 := (exists (vars (?x65 T2)) #948) +#936 := (and #73 #236) +#959 := (not #936) +#960 := (or #959 #953) +#965 := (forall (vars (?x64 T2)) #960) +#1020 := (not #965) +#1021 := (or #1020 #1012) +#1029 := (and #965 #1021) +#1045 := (or #1044 #1029) +#1054 := (or #1053 #1045) +#1063 := (or #1062 #1054) +#1072 := (or #1071 #1063) +#579 := (and #129 #508) +#582 := (exists (vars (?x41 T2)) #579) +#1091 := (or #582 #1072) +#703 := (and #192 #630) +#676 := (+ #91 #158) +#697 := (= #180 #676) +#708 := (and #697 #703) +#711 := (exists (vars (?x53 T2)) #708) +#694 := (and #73 #189) +#717 := (not #694) +#718 := (or #717 #711) +#723 := (forall (vars (?x52 T2)) #718) +#679 := (<= #180 #676) +#673 := (and #92 #630) +#685 := (not #673) +#686 := (or #685 #679) +#691 := (forall (vars (?x50 T2) (?x51 T2)) #686) +#745 := (not #691) +#746 := (or #745 #723) +#754 := (and #691 #746) +#665 := (or #181 #664) +#670 := (forall (vars (?x48 T2) (?x49 T2)) #665) +#762 := (not #670) +#763 := (or #762 #754) +#771 := (and #670 #763) +#779 := (not #175) +#780 := (or #779 #771) +#788 := (and #175 #780) +#645 := (= 0::int #172) +#796 := (not #645) +#797 := (or #796 #788) +#805 := (and #645 #797) +#825 := (or #824 #805) +#833 := (and #642 #825) +#841 := (not #167) +#842 := (or #841 #833) +#850 := (and #167 #842) +#622 := (or #157 #616) +#627 := (forall (vars (?x44 T2)) #622) +#858 := (not #627) +#859 := (or #858 #850) +#602 := (= #154 #158) +#608 := (or #162 #602) +#613 := (forall (vars (?x43 T2)) #608) +#867 := (not #613) +#868 := (or #867 #859) +#877 := (or #876 #868) +#594 := (or #146 #505) +#599 := (forall (vars (?x42 T2)) #594) +#885 := (not #599) +#886 := (or #885 #877) +#894 := (not #145) +#895 := (or #894 #886) +#903 := (or #585 #895) +#911 := (not #582) +#912 := (or #911 #903) +#1107 := (and #912 #1091) +#556 := (and #132 #505) +#529 := (+ #91 #110) +#550 := (= #120 #529) +#561 := (and #550 #556) +#564 := (exists (vars (?x40 T2)) #561) +#547 := (and #73 #129) +#570 := (not #547) +#571 := (or #570 #564) +#576 := (forall (vars (?x39 T2)) #571) +#1120 := (not #576) +#1121 := (or #1120 #1107) +#532 := (<= #120 #529) +#526 := (and #92 #505) +#538 := (not #526) +#539 := (or #538 #532) +#544 := (forall (vars (?x37 T2) (?x38 T2)) #539) +#1129 := (not #544) +#1130 := (or #1129 #1121) +#518 := (or #121 #517) +#523 := (forall (vars (?x35 T2) (?x36 T2)) #518) +#1138 := (not #523) +#1139 := (or #1138 #1130) +#1147 := (not #112) +#1148 := (or #1147 #1139) +#499 := (= 0::int #108) +#1156 := (not #499) +#1157 := (or #1156 #1148) +#484 := (and #73 #98) +#490 := (not #484) +#491 := (or #104 #490) +#496 := (forall (vars (?x32 T2)) #491) +#1176 := (not #496) +#1177 := (or #1176 #1157) +#1185 := (and #496 #1177) +#469 := (and #77 #92) +#475 := (not #469) +#476 := (or #95 #475) +#481 := (forall (vars (?x30 T2) (?x31 T2)) #476) +#1193 := (not #481) +#1194 := (or #1193 #1185) +#1202 := (and #481 #1194) +#461 := (or #87 #460) +#466 := (forall (vars (?x28 T2) (?x29 T2)) #461) +#1210 := (not #466) +#1211 := (or #1210 #1202) +#1219 := (and #466 #1211) +#1227 := (not #83) +#1228 := (or #1227 #1219) +#1236 := (and #83 #1228) +#451 := (= 0::int #80) +#1244 := (not #451) +#1245 := (or #1244 #1236) +#1253 := (and #451 #1245) +#1269 := (or #1268 #1253) +#1278 := (or #1277 #1269) +#423 := (= 0::int #69) +#429 := (or #73 #423) +#434 := (forall (vars (?x24 T2)) #429) +#1286 := (not #434) +#1287 := (or #1286 #1278) +#1303 := (not #1287) +#1936 := (iff #1303 #1935) +#1933 := (iff #1287 #1930) +#1921 := (or #1268 #1907) +#1924 := (or #1277 #1921) +#1927 := (or #1918 #1924) +#1931 := (iff #1927 #1930) +#1932 := [rewrite]: #1931 +#1928 := (iff #1287 #1927) +#1925 := (iff #1278 #1924) +#1922 := (iff #1269 #1921) +#1908 := (iff #1253 #1907) +#1905 := (iff #1245 #1904) +#1902 := (iff #1236 #1901) +#1899 := (iff #1228 #1898) +#1896 := (iff #1219 #1895) +#1893 := (iff #1211 #1892) +#1890 := (iff #1202 #1889) +#1887 := (iff #1194 #1886) +#1884 := (iff #1185 #1883) +#1881 := (iff #1177 #1878) +#1860 := (or #1446 #1801) +#1863 := (or #1824 #1860) +#1866 := (or #1840 #1863) +#1869 := (or #1849 #1866) +#1872 := (or #1854 #1869) +#1875 := (or #1857 #1872) +#1879 := (iff #1875 #1878) +#1880 := [rewrite]: #1879 +#1876 := (iff #1177 #1875) +#1873 := (iff #1157 #1872) +#1870 := (iff #1148 #1869) +#1867 := (iff #1139 #1866) +#1864 := (iff #1130 #1863) +#1861 := (iff #1121 #1860) +#1802 := (iff #1107 #1801) +#1799 := (iff #1091 #1796) +#1781 := (or #1044 #1778) +#1784 := (or #1053 #1781) +#1787 := (or #1062 #1784) +#1790 := (or #1071 #1787) +#1793 := (or #1675 #1790) +#1797 := (iff #1793 #1796) +#1798 := [rewrite]: #1797 +#1794 := (iff #1091 #1793) +#1791 := (iff #1072 #1790) +#1788 := (iff #1063 #1787) +#1785 := (iff #1054 #1784) +#1782 := (iff #1045 #1781) +#1779 := (iff #1029 #1778) +#1776 := (iff #1021 #1775) +#1773 := (iff #1012 #1772) +#1770 := (iff #1004 #1767) +#1764 := (or #1759 #251) +#1768 := (iff #1764 #1767) +#1769 := [rewrite]: #1768 +#1765 := (iff #1004 #1764) +#1762 := (iff #983 #251) +#1763 := [rewrite]: #1762 +#1760 := (iff #1003 #1759) +#1757 := (iff #980 #1756) +#1754 := (iff #975 #1753) +#1751 := (iff #968 #1750) +#1752 := [rewrite]: #1751 +#1748 := (iff #974 #1747) +#1745 := (iff #246 #1744) +#1711 := (iff #236 #1710) +#1712 := [rewrite]: #1711 +#1349 := (iff #92 #1348) +#1350 := [rewrite]: #1349 +#1746 := [monotonicity #1350 #1712]: #1745 +#1749 := [monotonicity #1746]: #1748 +#1755 := [monotonicity #1749 #1752]: #1754 +#1758 := [quant-intro #1755]: #1757 +#1761 := [monotonicity #1758]: #1760 +#1766 := [monotonicity #1761 #1763]: #1765 +#1771 := [trans #1766 #1769]: #1770 +#1774 := [monotonicity #1758 #1771]: #1773 +#1742 := (iff #1020 #1741) +#1739 := (iff #965 #1738) +#1736 := (iff #960 #1735) +#1733 := (iff #953 #1732) +#1730 := (iff #948 #1729) +#1726 := (iff #942 #1725) +#1728 := [rewrite]: #1726 +#1723 := (iff #241 #1722) +#1724 := [rewrite]: #1723 +#1731 := [monotonicity #1724 #1728]: #1730 +#1734 := [quant-intro #1731]: #1733 +#1717 := (iff #959 #1716) +#1714 := (iff #936 #1713) +#1715 := [monotonicity #1712]: #1714 +#1718 := [monotonicity #1715]: #1717 +#1737 := [monotonicity #1718 #1734]: #1736 +#1740 := [quant-intro #1737]: #1739 +#1743 := [monotonicity #1740]: #1742 +#1777 := [monotonicity #1743 #1774]: #1776 +#1780 := [monotonicity #1740 #1777]: #1779 +#1783 := [monotonicity #1780]: #1782 +#1786 := [monotonicity #1783]: #1785 +#1789 := [monotonicity #1786]: #1788 +#1792 := [monotonicity #1789]: #1791 +#1676 := (iff #582 #1675) +#1673 := (iff #579 #1670) +#1667 := (and #1407 #508) +#1671 := (iff #1667 #1670) +#1672 := [rewrite]: #1671 +#1668 := (iff #579 #1667) +#1408 := (iff #129 #1407) +#1409 := [rewrite]: #1408 +#1669 := [monotonicity #1409]: #1668 +#1674 := [trans #1669 #1672]: #1673 +#1677 := [quant-intro #1674]: #1676 +#1795 := [monotonicity #1677 #1792]: #1794 +#1800 := [trans #1795 #1798]: #1799 +#1705 := (iff #912 #1702) +#1681 := (or #1476 #1617) +#1684 := (or #1634 #1681) +#1687 := (or #876 #1684) +#1690 := (or #1652 #1687) +#1693 := (or #1656 #1690) +#1696 := (or #585 #1693) +#1699 := (or #1678 #1696) +#1703 := (iff #1699 #1702) +#1704 := [rewrite]: #1703 +#1700 := (iff #912 #1699) +#1697 := (iff #903 #1696) +#1694 := (iff #895 #1693) +#1691 := (iff #886 #1690) +#1688 := (iff #877 #1687) +#1685 := (iff #868 #1684) +#1682 := (iff #859 #1681) +#1618 := (iff #850 #1617) +#1615 := (iff #842 #1614) +#1612 := (iff #833 #1611) +#1609 := (iff #825 #1608) +#1606 := (iff #805 #1605) +#1603 := (iff #797 #1602) +#1600 := (iff #788 #1599) +#1597 := (iff #780 #1596) +#1594 := (iff #771 #1593) +#1591 := (iff #763 #1590) +#1588 := (iff #754 #1587) +#1585 := (iff #746 #1584) +#1582 := (iff #723 #1581) +#1579 := (iff #718 #1578) +#1576 := (iff #711 #1575) +#1573 := (iff #708 #1570) +#1564 := (and #1558 #630) +#1567 := (and #1561 #1564) +#1571 := (iff #1567 #1570) +#1572 := [rewrite]: #1571 +#1568 := (iff #708 #1567) +#1565 := (iff #703 #1564) +#1559 := (iff #192 #1558) +#1560 := [rewrite]: #1559 +#1566 := [monotonicity #1560]: #1565 +#1562 := (iff #697 #1561) +#1563 := [rewrite]: #1562 +#1569 := [monotonicity #1563 #1566]: #1568 +#1574 := [trans #1569 #1572]: #1573 +#1577 := [quant-intro #1574]: #1576 +#1556 := (iff #717 #1555) +#1553 := (iff #694 #1552) +#1550 := (iff #189 #1549) +#1551 := [rewrite]: #1550 +#1554 := [monotonicity #1551]: #1553 +#1557 := [monotonicity #1554]: #1556 +#1580 := [monotonicity #1557 #1577]: #1579 +#1583 := [quant-intro #1580]: #1582 +#1545 := (iff #745 #1544) +#1542 := (iff #691 #1541) +#1539 := (iff #686 #1538) +#1535 := (iff #679 #1534) +#1537 := [rewrite]: #1535 +#1532 := (iff #685 #1531) +#1529 := (iff #673 #1526) +#1523 := (and #1348 #630) +#1527 := (iff #1523 #1526) +#1528 := [rewrite]: #1527 +#1524 := (iff #673 #1523) +#1525 := [monotonicity #1350]: #1524 +#1530 := [trans #1525 #1528]: #1529 +#1533 := [monotonicity #1530]: #1532 +#1540 := [monotonicity #1533 #1537]: #1539 +#1543 := [quant-intro #1540]: #1542 +#1546 := [monotonicity #1543]: #1545 +#1586 := [monotonicity #1546 #1583]: #1585 +#1589 := [monotonicity #1543 #1586]: #1588 +#1521 := (iff #762 #1520) +#1518 := (iff #670 #1517) +#1515 := (iff #665 #1512) +#1509 := (or #1504 #664) +#1513 := (iff #1509 #1512) +#1514 := [rewrite]: #1513 +#1510 := (iff #665 #1509) +#1507 := (iff #181 #1504) +#1508 := [rewrite]: #1507 +#1511 := [monotonicity #1508]: #1510 +#1516 := [trans #1511 #1514]: #1515 +#1519 := [quant-intro #1516]: #1518 +#1522 := [monotonicity #1519]: #1521 +#1592 := [monotonicity #1522 #1589]: #1591 +#1595 := [monotonicity #1519 #1592]: #1594 +#1502 := (iff #779 #1501) +#1499 := (iff #175 #1498) +#1496 := (iff #174 #1495) +#1497 := [rewrite]: #1496 +#1500 := [quant-intro #1497]: #1499 +#1503 := [monotonicity #1500]: #1502 +#1598 := [monotonicity #1503 #1595]: #1597 +#1601 := [monotonicity #1500 #1598]: #1600 +#1493 := (iff #796 #1492) +#1490 := (iff #645 #173) +#1491 := [rewrite]: #1490 +#1494 := [monotonicity #1491]: #1493 +#1604 := [monotonicity #1494 #1601]: #1603 +#1607 := [monotonicity #1491 #1604]: #1606 +#1610 := [monotonicity #1607]: #1609 +#1613 := [monotonicity #1610]: #1612 +#1488 := (iff #841 #1487) +#1485 := (iff #167 #1484) +#1482 := (iff #166 #1479) +#1483 := [rewrite]: #1482 +#1486 := [quant-intro #1483]: #1485 +#1489 := [monotonicity #1486]: #1488 +#1616 := [monotonicity #1489 #1613]: #1615 +#1619 := [monotonicity #1486 #1616]: #1618 +#1477 := (iff #858 #1476) +#1474 := (iff #627 #1473) +#1471 := (iff #622 #1468) +#1465 := (or #1462 #616) +#1469 := (iff #1465 #1468) +#1470 := [rewrite]: #1469 +#1466 := (iff #622 #1465) +#1463 := (iff #157 #1462) +#1460 := (iff #156 #1459) +#1461 := [rewrite]: #1460 +#1455 := (iff #155 #1454) +#1456 := [rewrite]: #1455 +#1464 := [monotonicity #1456 #1461]: #1463 +#1467 := [monotonicity #1464]: #1466 +#1472 := [trans #1467 #1470]: #1471 +#1475 := [quant-intro #1472]: #1474 +#1478 := [monotonicity #1475]: #1477 +#1683 := [monotonicity #1478 #1619]: #1682 +#1635 := (iff #867 #1634) +#1632 := (iff #613 #1631) +#1629 := (iff #608 #1628) +#1626 := (iff #602 #1625) +#1627 := [rewrite]: #1626 +#1621 := (iff #162 #1620) +#1622 := [monotonicity #1464]: #1621 +#1630 := [monotonicity #1622 #1627]: #1629 +#1633 := [quant-intro #1630]: #1632 +#1636 := [monotonicity #1633]: #1635 +#1686 := [monotonicity #1636 #1683]: #1685 +#1689 := [monotonicity #1686]: #1688 +#1653 := (iff #885 #1652) +#1650 := (iff #599 #1649) +#1647 := (iff #594 #1644) +#1641 := (or #1637 #505) +#1645 := (iff #1641 #1644) +#1646 := [rewrite]: #1645 +#1642 := (iff #594 #1641) +#1639 := (iff #146 #1637) +#1640 := [rewrite]: #1639 +#1643 := [monotonicity #1640]: #1642 +#1648 := [trans #1643 #1646]: #1647 +#1651 := [quant-intro #1648]: #1650 +#1654 := [monotonicity #1651]: #1653 +#1692 := [monotonicity #1654 #1689]: #1691 +#1665 := (iff #894 #1656) +#1657 := (not #1656) +#1660 := (not #1657) +#1663 := (iff #1660 #1656) +#1664 := [rewrite]: #1663 +#1661 := (iff #894 #1660) +#1658 := (iff #145 #1657) +#1659 := [rewrite]: #1658 +#1662 := [monotonicity #1659]: #1661 +#1666 := [trans #1662 #1664]: #1665 +#1695 := [monotonicity #1666 #1692]: #1694 +#1698 := [monotonicity #1695]: #1697 +#1679 := (iff #911 #1678) +#1680 := [monotonicity #1677]: #1679 +#1701 := [monotonicity #1680 #1698]: #1700 +#1706 := [trans #1701 #1704]: #1705 +#1803 := [monotonicity #1706 #1800]: #1802 +#1447 := (iff #1120 #1446) +#1444 := (iff #576 #1443) +#1441 := (iff #571 #1440) +#1438 := (iff #564 #1437) +#1435 := (iff #561 #1432) +#1426 := (and #1419 #505) +#1429 := (and #1422 #1426) +#1433 := (iff #1429 #1432) +#1434 := [rewrite]: #1433 +#1430 := (iff #561 #1429) +#1427 := (iff #556 #1426) +#1420 := (iff #132 #1419) +#1421 := [rewrite]: #1420 +#1428 := [monotonicity #1421]: #1427 +#1423 := (iff #550 #1422) +#1425 := [rewrite]: #1423 +#1431 := [monotonicity #1425 #1428]: #1430 +#1436 := [trans #1431 #1434]: #1435 +#1439 := [quant-intro #1436]: #1438 +#1414 := (iff #570 #1413) +#1411 := (iff #547 #1410) +#1412 := [monotonicity #1409]: #1411 +#1415 := [monotonicity #1412]: #1414 +#1442 := [monotonicity #1415 #1439]: #1441 +#1445 := [quant-intro #1442]: #1444 +#1448 := [monotonicity #1445]: #1447 +#1862 := [monotonicity #1448 #1803]: #1861 +#1825 := (iff #1129 #1824) +#1822 := (iff #544 #1821) +#1819 := (iff #539 #1818) +#1816 := (iff #532 #1815) +#1817 := [rewrite]: #1816 +#1813 := (iff #538 #1812) +#1810 := (iff #526 #1807) +#1804 := (and #1348 #505) +#1808 := (iff #1804 #1807) +#1809 := [rewrite]: #1808 +#1805 := (iff #526 #1804) +#1806 := [monotonicity #1350]: #1805 +#1811 := [trans #1806 #1809]: #1810 +#1814 := [monotonicity #1811]: #1813 +#1820 := [monotonicity #1814 #1817]: #1819 +#1823 := [quant-intro #1820]: #1822 +#1826 := [monotonicity #1823]: #1825 +#1865 := [monotonicity #1826 #1862]: #1864 +#1841 := (iff #1138 #1840) +#1838 := (iff #523 #1837) +#1835 := (iff #518 #1832) +#1829 := (or #1416 #517) +#1833 := (iff #1829 #1832) +#1834 := [rewrite]: #1833 +#1830 := (iff #518 #1829) +#1827 := (iff #121 #1416) +#1828 := [rewrite]: #1827 +#1831 := [monotonicity #1828]: #1830 +#1836 := [trans #1831 #1834]: #1835 +#1839 := [quant-intro #1836]: #1838 +#1842 := [monotonicity #1839]: #1841 +#1868 := [monotonicity #1842 #1865]: #1867 +#1850 := (iff #1147 #1849) +#1847 := (iff #112 #1846) +#1844 := (iff #111 #1843) +#1845 := [rewrite]: #1844 +#1848 := [quant-intro #1845]: #1847 +#1851 := [monotonicity #1848]: #1850 +#1871 := [monotonicity #1851 #1868]: #1870 +#1855 := (iff #1156 #1854) +#1852 := (iff #499 #109) +#1853 := [rewrite]: #1852 +#1856 := [monotonicity #1853]: #1855 +#1874 := [monotonicity #1856 #1871]: #1873 +#1858 := (iff #1176 #1857) +#1402 := (iff #496 #1401) +#1399 := (iff #491 #1398) +#1396 := (iff #490 #1395) +#1393 := (iff #484 #1392) +#1390 := (iff #98 #1389) +#1391 := [rewrite]: #1390 +#1394 := [monotonicity #1391]: #1393 +#1397 := [monotonicity #1394]: #1396 +#1384 := (iff #104 #1383) +#1381 := (iff #103 #1378) +#1372 := (and #77 #1366) +#1375 := (and #1370 #1372) +#1379 := (iff #1375 #1378) +#1380 := [rewrite]: #1379 +#1376 := (iff #103 #1375) +#1373 := (iff #102 #1372) +#1367 := (iff #101 #1366) +#1368 := [rewrite]: #1367 +#1374 := [monotonicity #1368]: #1373 +#1369 := (iff #100 #1370) +#1371 := [rewrite]: #1369 +#1377 := [monotonicity #1371 #1374]: #1376 +#1382 := [trans #1377 #1380]: #1381 +#1385 := [quant-intro #1382]: #1384 +#1400 := [monotonicity #1385 #1397]: #1399 +#1403 := [quant-intro #1400]: #1402 +#1859 := [monotonicity #1403]: #1858 +#1877 := [monotonicity #1859 #1874]: #1876 +#1882 := [trans #1877 #1880]: #1881 +#1885 := [monotonicity #1403 #1882]: #1884 +#1364 := (iff #1193 #1363) +#1361 := (iff #481 #1360) +#1358 := (iff #476 #1357) +#1355 := (iff #475 #1354) +#1352 := (iff #469 #1351) +#1353 := [monotonicity #1350]: #1352 +#1356 := [monotonicity #1353]: #1355 +#1341 := (iff #95 #1342) +#1340 := [rewrite]: #1341 +#1359 := [monotonicity #1340 #1356]: #1358 +#1362 := [quant-intro #1359]: #1361 +#1365 := [monotonicity #1362]: #1364 +#1888 := [monotonicity #1365 #1885]: #1887 +#1891 := [monotonicity #1362 #1888]: #1890 +#1338 := (iff #1210 #1337) +#1335 := (iff #466 #1334) +#1332 := (iff #461 #1329) +#1326 := (or #1322 #460) +#1330 := (iff #1326 #1329) +#1331 := [rewrite]: #1330 +#1327 := (iff #461 #1326) +#1321 := (iff #87 #1322) +#1325 := [rewrite]: #1321 +#1328 := [monotonicity #1325]: #1327 +#1333 := [trans #1328 #1331]: #1332 +#1336 := [quant-intro #1333]: #1335 +#1339 := [monotonicity #1336]: #1338 +#1894 := [monotonicity #1339 #1891]: #1893 +#1897 := [monotonicity #1336 #1894]: #1896 +#1318 := (iff #1227 #1317) +#1315 := (iff #83 #1314) +#1311 := (iff #82 #1313) +#1312 := [rewrite]: #1311 +#1316 := [quant-intro #1312]: #1315 +#1319 := [monotonicity #1316]: #1318 +#1900 := [monotonicity #1319 #1897]: #1899 +#1903 := [monotonicity #1316 #1900]: #1902 +#1309 := (iff #1244 #1308) +#1306 := (iff #451 #81) +#1307 := [rewrite]: #1306 +#1310 := [monotonicity #1307]: #1309 +#1906 := [monotonicity #1310 #1903]: #1905 +#1909 := [monotonicity #1307 #1906]: #1908 +#1923 := [monotonicity #1909]: #1922 +#1926 := [monotonicity #1923]: #1925 +#1919 := (iff #1286 #1918) +#1916 := (iff #434 #1915) +#1913 := (iff #429 #1912) +#1910 := (iff #423 #70) +#1911 := [rewrite]: #1910 +#1914 := [monotonicity #1911]: #1913 +#1917 := [quant-intro #1914]: #1916 +#1920 := [monotonicity #1917]: #1919 +#1929 := [monotonicity #1920 #1926]: #1928 +#1934 := [trans #1929 #1932]: #1933 +#1937 := [monotonicity #1934]: #1936 +#1304 := (iff #293 #1303) +#1301 := (iff #292 #1287) +#1292 := (implies true #1287) +#1295 := (iff #1292 #1287) +#1296 := [rewrite]: #1295 +#1299 := (iff #292 #1292) +#1297 := (iff #291 #1287) +#1293 := (iff #291 #1292) +#1290 := (iff #290 #1287) +#1283 := (implies #434 #1278) +#1288 := (iff #1283 #1287) +#1289 := [rewrite]: #1288 +#1284 := (iff #290 #1283) +#1281 := (iff #289 #1278) +#1274 := (implies #448 #1269) +#1279 := (iff #1274 #1278) +#1280 := [rewrite]: #1279 +#1275 := (iff #289 #1274) +#1272 := (iff #288 #1269) +#1265 := (implies #79 #1253) +#1270 := (iff #1265 #1269) +#1271 := [rewrite]: #1270 +#1266 := (iff #288 #1265) +#1263 := (iff #287 #1253) +#1258 := (implies true #1253) +#1261 := (iff #1258 #1253) +#1262 := [rewrite]: #1261 +#1259 := (iff #287 #1258) +#1256 := (iff #286 #1253) +#1250 := (and #1245 #451) +#1254 := (iff #1250 #1253) +#1255 := [rewrite]: #1254 +#1251 := (iff #286 #1250) +#452 := (iff #81 #451) +#453 := [rewrite]: #452 +#1248 := (iff #285 #1245) +#1241 := (implies #451 #1236) +#1246 := (iff #1241 #1245) +#1247 := [rewrite]: #1246 +#1242 := (iff #285 #1241) +#1239 := (iff #284 #1236) +#1233 := (and #1228 #83) +#1237 := (iff #1233 #1236) +#1238 := [rewrite]: #1237 +#1234 := (iff #284 #1233) +#1231 := (iff #283 #1228) +#1224 := (implies #83 #1219) +#1229 := (iff #1224 #1228) +#1230 := [rewrite]: #1229 +#1225 := (iff #283 #1224) +#1222 := (iff #282 #1219) +#1216 := (and #1211 #466) +#1220 := (iff #1216 #1219) +#1221 := [rewrite]: #1220 +#1217 := (iff #282 #1216) +#467 := (iff #89 #466) +#464 := (iff #88 #461) +#457 := (implies #454 #87) +#462 := (iff #457 #461) +#463 := [rewrite]: #462 +#458 := (iff #88 #457) +#455 := (iff #85 #454) +#456 := [rewrite]: #455 +#459 := [monotonicity #456]: #458 +#465 := [trans #459 #463]: #464 +#468 := [quant-intro #465]: #467 +#1214 := (iff #281 #1211) +#1207 := (implies #466 #1202) +#1212 := (iff #1207 #1211) +#1213 := [rewrite]: #1212 +#1208 := (iff #281 #1207) +#1205 := (iff #280 #1202) +#1199 := (and #1194 #481) +#1203 := (iff #1199 #1202) +#1204 := [rewrite]: #1203 +#1200 := (iff #280 #1199) +#482 := (iff #97 #481) +#479 := (iff #96 #476) +#472 := (implies #469 #95) +#477 := (iff #472 #476) +#478 := [rewrite]: #477 +#473 := (iff #96 #472) +#470 := (iff #93 #469) +#471 := [rewrite]: #470 +#474 := [monotonicity #471]: #473 +#480 := [trans #474 #478]: #479 +#483 := [quant-intro #480]: #482 +#1197 := (iff #279 #1194) +#1190 := (implies #481 #1185) +#1195 := (iff #1190 #1194) +#1196 := [rewrite]: #1195 +#1191 := (iff #279 #1190) +#1188 := (iff #278 #1185) +#1182 := (and #1177 #496) +#1186 := (iff #1182 #1185) +#1187 := [rewrite]: #1186 +#1183 := (iff #278 #1182) +#497 := (iff #106 #496) +#494 := (iff #105 #491) +#487 := (implies #484 #104) +#492 := (iff #487 #491) +#493 := [rewrite]: #492 +#488 := (iff #105 #487) +#485 := (iff #99 #484) +#486 := [rewrite]: #485 +#489 := [monotonicity #486]: #488 +#495 := [trans #489 #493]: #494 +#498 := [quant-intro #495]: #497 +#1180 := (iff #277 #1177) +#1173 := (implies #496 #1157) +#1178 := (iff #1173 #1177) +#1179 := [rewrite]: #1178 +#1174 := (iff #277 #1173) +#1171 := (iff #276 #1157) +#1162 := (implies true #1157) +#1165 := (iff #1162 #1157) +#1166 := [rewrite]: #1165 +#1169 := (iff #276 #1162) +#1167 := (iff #275 #1157) +#1163 := (iff #275 #1162) +#1160 := (iff #274 #1157) +#1153 := (implies #499 #1148) +#1158 := (iff #1153 #1157) +#1159 := [rewrite]: #1158 +#1154 := (iff #274 #1153) +#1151 := (iff #273 #1148) +#1144 := (implies #112 #1139) +#1149 := (iff #1144 #1148) +#1150 := [rewrite]: #1149 +#1145 := (iff #273 #1144) +#1142 := (iff #272 #1139) +#1135 := (implies #523 #1130) +#1140 := (iff #1135 #1139) +#1141 := [rewrite]: #1140 +#1136 := (iff #272 #1135) +#1133 := (iff #271 #1130) +#1126 := (implies #544 #1121) +#1131 := (iff #1126 #1130) +#1132 := [rewrite]: #1131 +#1127 := (iff #271 #1126) +#1124 := (iff #270 #1121) +#1117 := (implies #576 #1107) +#1122 := (iff #1117 #1121) +#1123 := [rewrite]: #1122 +#1118 := (iff #270 #1117) +#1115 := (iff #269 #1107) +#1110 := (implies true #1107) +#1113 := (iff #1110 #1107) +#1114 := [rewrite]: #1113 +#1111 := (iff #269 #1110) +#1108 := (iff #268 #1107) +#1105 := (iff #267 #1091) +#1096 := (implies true #1091) +#1099 := (iff #1096 #1091) +#1100 := [rewrite]: #1099 +#1103 := (iff #267 #1096) +#1101 := (iff #266 #1091) +#1097 := (iff #266 #1096) +#1094 := (iff #265 #1091) +#1088 := (implies #911 #1072) +#1092 := (iff #1088 #1091) +#1093 := [rewrite]: #1092 +#1089 := (iff #265 #1088) +#1086 := (iff #264 #1072) +#1077 := (implies true #1072) +#1080 := (iff #1077 #1072) +#1081 := [rewrite]: #1080 +#1084 := (iff #264 #1077) +#1082 := (iff #263 #1072) +#1078 := (iff #263 #1077) +#1075 := (iff #262 #1072) +#1068 := (implies #930 #1063) +#1073 := (iff #1068 #1072) +#1074 := [rewrite]: #1073 +#1069 := (iff #262 #1068) +#1066 := (iff #261 #1063) +#1059 := (implies #229 #1054) +#1064 := (iff #1059 #1063) +#1065 := [rewrite]: #1064 +#1060 := (iff #261 #1059) +#1057 := (iff #260 #1054) +#1050 := (implies #933 #1045) +#1055 := (iff #1050 #1054) +#1056 := [rewrite]: #1055 +#1051 := (iff #260 #1050) +#1048 := (iff #259 #1045) +#1041 := (implies #234 #1029) +#1046 := (iff #1041 #1045) +#1047 := [rewrite]: #1046 +#1042 := (iff #259 #1041) +#1039 := (iff #258 #1029) +#1034 := (implies true #1029) +#1037 := (iff #1034 #1029) +#1038 := [rewrite]: #1037 +#1035 := (iff #258 #1034) +#1032 := (iff #257 #1029) +#1026 := (and #1021 #965) +#1030 := (iff #1026 #1029) +#1031 := [rewrite]: #1030 +#1027 := (iff #257 #1026) +#966 := (iff #245 #965) +#963 := (iff #244 #960) +#956 := (implies #936 #953) +#961 := (iff #956 #960) +#962 := [rewrite]: #961 +#957 := (iff #244 #956) +#954 := (iff #243 #953) +#951 := (iff #242 #948) +#945 := (and #942 #241) +#949 := (iff #945 #948) +#950 := [rewrite]: #949 +#946 := (iff #242 #945) +#943 := (iff #240 #942) +#940 := (= #239 #939) +#941 := [rewrite]: #940 +#944 := [monotonicity #941]: #943 +#947 := [monotonicity #944]: #946 +#952 := [trans #947 #950]: #951 +#955 := [quant-intro #952]: #954 +#937 := (iff #237 #936) +#938 := [rewrite]: #937 +#958 := [monotonicity #938 #955]: #957 +#964 := [trans #958 #962]: #963 +#967 := [quant-intro #964]: #966 +#1024 := (iff #256 #1021) +#1017 := (implies #965 #1012) +#1022 := (iff #1017 #1021) +#1023 := [rewrite]: #1022 +#1018 := (iff #256 #1017) +#1015 := (iff #255 #1012) +#1009 := (and #1004 #980) +#1013 := (iff #1009 #1012) +#1014 := [rewrite]: #1013 +#1010 := (iff #255 #1009) +#981 := (iff #249 #980) +#978 := (iff #248 #975) +#971 := (implies #246 #968) +#976 := (iff #971 #975) +#977 := [rewrite]: #976 +#972 := (iff #248 #971) +#969 := (iff #247 #968) +#970 := [monotonicity #941]: #969 +#973 := [monotonicity #970]: #972 +#979 := [trans #973 #977]: #978 +#982 := [quant-intro #979]: #981 +#1007 := (iff #254 #1004) +#1000 := (implies #980 #983) +#1005 := (iff #1000 #1004) +#1006 := [rewrite]: #1005 +#1001 := (iff #254 #1000) +#998 := (iff #253 #983) +#993 := (and true #983) +#996 := (iff #993 #983) +#997 := [rewrite]: #996 +#994 := (iff #253 #993) +#984 := (iff #251 #983) +#985 := [rewrite]: #984 +#991 := (iff #252 true) +#986 := (implies #983 true) +#989 := (iff #986 true) +#990 := [rewrite]: #989 +#987 := (iff #252 #986) +#988 := [monotonicity #985]: #987 +#992 := [trans #988 #990]: #991 +#995 := [monotonicity #992 #985]: #994 +#999 := [trans #995 #997]: #998 +#1002 := [monotonicity #982 #999]: #1001 +#1008 := [trans #1002 #1006]: #1007 +#1011 := [monotonicity #1008 #982]: #1010 +#1016 := [trans #1011 #1014]: #1015 +#1019 := [monotonicity #967 #1016]: #1018 +#1025 := [trans #1019 #1023]: #1024 +#1028 := [monotonicity #1025 #967]: #1027 +#1033 := [trans #1028 #1031]: #1032 +#1036 := [monotonicity #1033]: #1035 +#1040 := [trans #1036 #1038]: #1039 +#1043 := [monotonicity #1040]: #1042 +#1049 := [trans #1043 #1047]: #1048 +#934 := (iff #231 #933) +#935 := [rewrite]: #934 +#1052 := [monotonicity #935 #1049]: #1051 +#1058 := [trans #1052 #1056]: #1057 +#1061 := [monotonicity #1058]: #1060 +#1067 := [trans #1061 #1065]: #1066 +#931 := (iff #226 #930) +#932 := [rewrite]: #931 +#1070 := [monotonicity #932 #1067]: #1069 +#1076 := [trans #1070 #1074]: #1075 +#1079 := [monotonicity #1076]: #1078 +#1083 := [trans #1079 #1081]: #1082 +#1085 := [monotonicity #1083]: #1084 +#1087 := [trans #1085 #1081]: #1086 +#928 := (iff #224 #911) +#583 := (iff #139 #582) +#580 := (iff #138 #579) +#509 := (iff #118 #508) +#506 := (iff #117 #505) +#507 := [rewrite]: #506 +#510 := [monotonicity #507]: #509 +#581 := [monotonicity #510]: #580 +#584 := [quant-intro #581]: #583 +#929 := [monotonicity #584]: #928 +#1090 := [monotonicity #929 #1087]: #1089 +#1095 := [trans #1090 #1093]: #1094 +#1098 := [monotonicity #1095]: #1097 +#1102 := [trans #1098 #1100]: #1101 +#1104 := [monotonicity #1102]: #1103 +#1106 := [trans #1104 #1100]: #1105 +#926 := (iff #223 #912) +#917 := (implies true #912) +#920 := (iff #917 #912) +#921 := [rewrite]: #920 +#924 := (iff #223 #917) +#922 := (iff #222 #912) +#918 := (iff #222 #917) +#915 := (iff #221 #912) +#908 := (implies #582 #903) +#913 := (iff #908 #912) +#914 := [rewrite]: #913 +#909 := (iff #221 #908) +#906 := (iff #220 #903) +#588 := (not #585) +#900 := (implies #588 #895) +#904 := (iff #900 #903) +#905 := [rewrite]: #904 +#901 := (iff #220 #900) +#898 := (iff #219 #895) +#891 := (implies #145 #886) +#896 := (iff #891 #895) +#897 := [rewrite]: #896 +#892 := (iff #219 #891) +#889 := (iff #218 #886) +#882 := (implies #599 #877) +#887 := (iff #882 #886) +#888 := [rewrite]: #887 +#883 := (iff #218 #882) +#880 := (iff #217 #877) +#873 := (implies #151 #868) +#878 := (iff #873 #877) +#879 := [rewrite]: #878 +#874 := (iff #217 #873) +#871 := (iff #216 #868) +#864 := (implies #613 #859) +#869 := (iff #864 #868) +#870 := [rewrite]: #869 +#865 := (iff #216 #864) +#862 := (iff #215 #859) +#855 := (implies #627 #850) +#860 := (iff #855 #859) +#861 := [rewrite]: #860 +#856 := (iff #215 #855) +#853 := (iff #214 #850) +#847 := (and #842 #167) +#851 := (iff #847 #850) +#852 := [rewrite]: #851 +#848 := (iff #214 #847) +#845 := (iff #213 #842) +#838 := (implies #167 #833) +#843 := (iff #838 #842) +#844 := [rewrite]: #843 +#839 := (iff #213 #838) +#836 := (iff #212 #833) +#830 := (and #825 #642) +#834 := (iff #830 #833) +#835 := [rewrite]: #834 +#831 := (iff #212 #830) +#643 := (iff #171 #642) +#640 := (iff #170 #637) +#633 := (implies #630 #616) +#638 := (iff #633 #637) +#639 := [rewrite]: #638 +#634 := (iff #170 #633) +#617 := (iff #163 #616) +#618 := [rewrite]: #617 +#631 := (iff #169 #630) +#632 := [rewrite]: #631 +#635 := [monotonicity #632 #618]: #634 +#641 := [trans #635 #639]: #640 +#644 := [quant-intro #641]: #643 +#828 := (iff #211 #825) +#821 := (implies #642 #805) +#826 := (iff #821 #825) +#827 := [rewrite]: #826 +#822 := (iff #211 #821) +#819 := (iff #210 #805) +#810 := (implies true #805) +#813 := (iff #810 #805) +#814 := [rewrite]: #813 +#817 := (iff #210 #810) +#815 := (iff #209 #805) +#811 := (iff #209 #810) +#808 := (iff #208 #805) +#802 := (and #797 #645) +#806 := (iff #802 #805) +#807 := [rewrite]: #806 +#803 := (iff #208 #802) +#646 := (iff #173 #645) +#647 := [rewrite]: #646 +#800 := (iff #207 #797) +#793 := (implies #645 #788) +#798 := (iff #793 #797) +#799 := [rewrite]: #798 +#794 := (iff #207 #793) +#791 := (iff #206 #788) +#785 := (and #780 #175) +#789 := (iff #785 #788) +#790 := [rewrite]: #789 +#786 := (iff #206 #785) +#783 := (iff #205 #780) +#776 := (implies #175 #771) +#781 := (iff #776 #780) +#782 := [rewrite]: #781 +#777 := (iff #205 #776) +#774 := (iff #204 #771) +#768 := (and #763 #670) +#772 := (iff #768 #771) +#773 := [rewrite]: #772 +#769 := (iff #204 #768) +#671 := (iff #183 #670) +#668 := (iff #182 #665) +#661 := (implies #656 #181) +#666 := (iff #661 #665) +#667 := [rewrite]: #666 +#662 := (iff #182 #661) +#659 := (iff #179 #656) +#653 := (and #648 #636) +#657 := (iff #653 #656) +#658 := [rewrite]: #657 +#654 := (iff #179 #653) +#651 := (iff #178 #636) +#652 := [monotonicity #632]: #651 +#649 := (iff #177 #648) +#650 := [rewrite]: #649 +#655 := [monotonicity #650 #652]: #654 +#660 := [trans #655 #658]: #659 +#663 := [monotonicity #660]: #662 +#669 := [trans #663 #667]: #668 +#672 := [quant-intro #669]: #671 +#766 := (iff #203 #763) +#759 := (implies #670 #754) +#764 := (iff #759 #763) +#765 := [rewrite]: #764 +#760 := (iff #203 #759) +#757 := (iff #202 #754) +#751 := (and #746 #691) +#755 := (iff #751 #754) +#756 := [rewrite]: #755 +#752 := (iff #202 #751) +#692 := (iff #188 #691) +#689 := (iff #187 #686) +#682 := (implies #673 #679) +#687 := (iff #682 #686) +#688 := [rewrite]: #687 +#683 := (iff #187 #682) +#680 := (iff #186 #679) +#677 := (= #185 #676) +#678 := [rewrite]: #677 +#681 := [monotonicity #678]: #680 +#674 := (iff #184 #673) +#675 := [monotonicity #632]: #674 +#684 := [monotonicity #675 #681]: #683 +#690 := [trans #684 #688]: #689 +#693 := [quant-intro #690]: #692 +#749 := (iff #201 #746) +#742 := (implies #691 #723) +#747 := (iff #742 #746) +#748 := [rewrite]: #747 +#743 := (iff #201 #742) +#740 := (iff #200 #723) +#735 := (and true #723) +#738 := (iff #735 #723) +#739 := [rewrite]: #738 +#736 := (iff #200 #735) +#724 := (iff #197 #723) +#721 := (iff #196 #718) +#714 := (implies #694 #711) +#719 := (iff #714 #718) +#720 := [rewrite]: #719 +#715 := (iff #196 #714) +#712 := (iff #195 #711) +#709 := (iff #194 #708) +#706 := (iff #193 #703) +#700 := (and #630 #192) +#704 := (iff #700 #703) +#705 := [rewrite]: #704 +#701 := (iff #193 #700) +#702 := [monotonicity #632]: #701 +#707 := [trans #702 #705]: #706 +#698 := (iff #191 #697) +#699 := [monotonicity #678]: #698 +#710 := [monotonicity #699 #707]: #709 +#713 := [quant-intro #710]: #712 +#695 := (iff #190 #694) +#696 := [rewrite]: #695 +#716 := [monotonicity #696 #713]: #715 +#722 := [trans #716 #720]: #721 +#725 := [quant-intro #722]: #724 +#733 := (iff #199 true) +#728 := (implies #723 true) +#731 := (iff #728 true) +#732 := [rewrite]: #731 +#729 := (iff #199 #728) +#726 := (iff #198 true) +#727 := [rewrite]: #726 +#730 := [monotonicity #725 #727]: #729 +#734 := [trans #730 #732]: #733 +#737 := [monotonicity #734 #725]: #736 +#741 := [trans #737 #739]: #740 +#744 := [monotonicity #693 #741]: #743 +#750 := [trans #744 #748]: #749 +#753 := [monotonicity #750 #693]: #752 +#758 := [trans #753 #756]: #757 +#761 := [monotonicity #672 #758]: #760 +#767 := [trans #761 #765]: #766 +#770 := [monotonicity #767 #672]: #769 +#775 := [trans #770 #773]: #774 +#778 := [monotonicity #775]: #777 +#784 := [trans #778 #782]: #783 +#787 := [monotonicity #784]: #786 +#792 := [trans #787 #790]: #791 +#795 := [monotonicity #647 #792]: #794 +#801 := [trans #795 #799]: #800 +#804 := [monotonicity #801 #647]: #803 +#809 := [trans #804 #807]: #808 +#812 := [monotonicity #809]: #811 +#816 := [trans #812 #814]: #815 +#818 := [monotonicity #816]: #817 +#820 := [trans #818 #814]: #819 +#823 := [monotonicity #644 #820]: #822 +#829 := [trans #823 #827]: #828 +#832 := [monotonicity #829 #644]: #831 +#837 := [trans #832 #835]: #836 +#840 := [monotonicity #837]: #839 +#846 := [trans #840 #844]: #845 +#849 := [monotonicity #846]: #848 +#854 := [trans #849 #852]: #853 +#628 := (iff #165 #627) +#625 := (iff #164 #622) +#619 := (implies #162 #616) +#623 := (iff #619 #622) +#624 := [rewrite]: #623 +#620 := (iff #164 #619) +#621 := [monotonicity #618]: #620 +#626 := [trans #621 #624]: #625 +#629 := [quant-intro #626]: #628 +#857 := [monotonicity #629 #854]: #856 +#863 := [trans #857 #861]: #862 +#614 := (iff #161 #613) +#611 := (iff #160 #608) +#605 := (implies #157 #602) +#609 := (iff #605 #608) +#610 := [rewrite]: #609 +#606 := (iff #160 #605) +#603 := (iff #159 #602) +#604 := [rewrite]: #603 +#607 := [monotonicity #604]: #606 +#612 := [trans #607 #610]: #611 +#615 := [quant-intro #612]: #614 +#866 := [monotonicity #615 #863]: #865 +#872 := [trans #866 #870]: #871 +#875 := [monotonicity #872]: #874 +#881 := [trans #875 #879]: #880 +#600 := (iff #148 #599) +#597 := (iff #147 #594) +#591 := (implies #508 #146) +#595 := (iff #591 #594) +#596 := [rewrite]: #595 +#592 := (iff #147 #591) +#593 := [monotonicity #510]: #592 +#598 := [trans #593 #596]: #597 +#601 := [quant-intro #598]: #600 +#884 := [monotonicity #601 #881]: #883 +#890 := [trans #884 #888]: #889 +#893 := [monotonicity #890]: #892 +#899 := [trans #893 #897]: #898 +#589 := (iff #143 #588) +#586 := (iff #142 #585) +#587 := [rewrite]: #586 +#590 := [monotonicity #587]: #589 +#902 := [monotonicity #590 #899]: #901 +#907 := [trans #902 #905]: #906 +#910 := [monotonicity #584 #907]: #909 +#916 := [trans #910 #914]: #915 +#919 := [monotonicity #916]: #918 +#923 := [trans #919 #921]: #922 +#925 := [monotonicity #923]: #924 +#927 := [trans #925 #921]: #926 +#1109 := [monotonicity #927 #1106]: #1108 +#1112 := [monotonicity #1109]: #1111 +#1116 := [trans #1112 #1114]: #1115 +#577 := (iff #137 #576) +#574 := (iff #136 #571) +#567 := (implies #547 #564) +#572 := (iff #567 #571) +#573 := [rewrite]: #572 +#568 := (iff #136 #567) +#565 := (iff #135 #564) +#562 := (iff #134 #561) +#559 := (iff #133 #556) +#553 := (and #505 #132) +#557 := (iff #553 #556) +#558 := [rewrite]: #557 +#554 := (iff #133 #553) +#555 := [monotonicity #507]: #554 +#560 := [trans #555 #558]: #559 +#551 := (iff #131 #550) +#530 := (= #125 #529) +#531 := [rewrite]: #530 +#552 := [monotonicity #531]: #551 +#563 := [monotonicity #552 #560]: #562 +#566 := [quant-intro #563]: #565 +#548 := (iff #130 #547) +#549 := [rewrite]: #548 +#569 := [monotonicity #549 #566]: #568 +#575 := [trans #569 #573]: #574 +#578 := [quant-intro #575]: #577 +#1119 := [monotonicity #578 #1116]: #1118 +#1125 := [trans #1119 #1123]: #1124 +#545 := (iff #128 #544) +#542 := (iff #127 #539) +#535 := (implies #526 #532) +#540 := (iff #535 #539) +#541 := [rewrite]: #540 +#536 := (iff #127 #535) +#533 := (iff #126 #532) +#534 := [monotonicity #531]: #533 +#527 := (iff #124 #526) +#528 := [monotonicity #507]: #527 +#537 := [monotonicity #528 #534]: #536 +#543 := [trans #537 #541]: #542 +#546 := [quant-intro #543]: #545 +#1128 := [monotonicity #546 #1125]: #1127 +#1134 := [trans #1128 #1132]: #1133 +#524 := (iff #123 #523) +#521 := (iff #122 #518) +#514 := (implies #511 #121) +#519 := (iff #514 #518) +#520 := [rewrite]: #519 +#515 := (iff #122 #514) +#512 := (iff #119 #511) +#503 := (iff #115 #502) +#504 := [rewrite]: #503 +#513 := [monotonicity #504 #510]: #512 +#516 := [monotonicity #513]: #515 +#522 := [trans #516 #520]: #521 +#525 := [quant-intro #522]: #524 +#1137 := [monotonicity #525 #1134]: #1136 +#1143 := [trans #1137 #1141]: #1142 +#1146 := [monotonicity #1143]: #1145 +#1152 := [trans #1146 #1150]: #1151 +#500 := (iff #109 #499) +#501 := [rewrite]: #500 +#1155 := [monotonicity #501 #1152]: #1154 +#1161 := [trans #1155 #1159]: #1160 +#1164 := [monotonicity #1161]: #1163 +#1168 := [trans #1164 #1166]: #1167 +#1170 := [monotonicity #1168]: #1169 +#1172 := [trans #1170 #1166]: #1171 +#1175 := [monotonicity #498 #1172]: #1174 +#1181 := [trans #1175 #1179]: #1180 +#1184 := [monotonicity #1181 #498]: #1183 +#1189 := [trans #1184 #1187]: #1188 +#1192 := [monotonicity #483 #1189]: #1191 +#1198 := [trans #1192 #1196]: #1197 +#1201 := [monotonicity #1198 #483]: #1200 +#1206 := [trans #1201 #1204]: #1205 +#1209 := [monotonicity #468 #1206]: #1208 +#1215 := [trans #1209 #1213]: #1214 +#1218 := [monotonicity #1215 #468]: #1217 +#1223 := [trans #1218 #1221]: #1222 +#1226 := [monotonicity #1223]: #1225 +#1232 := [trans #1226 #1230]: #1231 +#1235 := [monotonicity #1232]: #1234 +#1240 := [trans #1235 #1238]: #1239 +#1243 := [monotonicity #453 #1240]: #1242 +#1249 := [trans #1243 #1247]: #1248 +#1252 := [monotonicity #1249 #453]: #1251 +#1257 := [trans #1252 #1255]: #1256 +#1260 := [monotonicity #1257]: #1259 +#1264 := [trans #1260 #1262]: #1263 +#1267 := [monotonicity #1264]: #1266 +#1273 := [trans #1267 #1271]: #1272 +#449 := (iff #76 #448) +#446 := (iff #75 #443) +#440 := (implies #73 #437) +#444 := (iff #440 #443) +#445 := [rewrite]: #444 +#441 := (iff #75 #440) +#438 := (iff #74 #437) +#439 := [rewrite]: #438 +#442 := [monotonicity #439]: #441 +#447 := [trans #442 #445]: #446 +#450 := [quant-intro #447]: #449 +#1276 := [monotonicity #450 #1273]: #1275 +#1282 := [trans #1276 #1280]: #1281 +#435 := (iff #72 #434) +#432 := (iff #71 #429) +#426 := (implies #68 #423) +#430 := (iff #426 #429) +#431 := [rewrite]: #430 +#427 := (iff #71 #426) +#424 := (iff #70 #423) +#425 := [rewrite]: #424 +#428 := [monotonicity #425]: #427 +#433 := [trans #428 #431]: #432 +#436 := [quant-intro #433]: #435 +#1285 := [monotonicity #436 #1282]: #1284 +#1291 := [trans #1285 #1289]: #1290 +#1294 := [monotonicity #1291]: #1293 +#1298 := [trans #1294 #1296]: #1297 +#1300 := [monotonicity #1298]: #1299 +#1302 := [trans #1300 #1296]: #1301 +#1305 := [monotonicity #1302]: #1304 +#1939 := [trans #1305 #1937]: #1938 +#422 := [asserted]: #293 +#1940 := [mp #422 #1939]: #1935 +#1941 := [not-or-elim #1940]: #79 +#2004 := [mp~ #1941 #1966]: #79 +#4354 := [mp #2004 #4353]: #4349 +#5709 := (not #4349) +#5715 := (or #5709 #2266) +#5470 := [quant-inst]: #5715 +#6005 := [unit-resolution #5470 #4354 #6004]: false +#6009 := [lemma #6005]: #2266 +#3870 := (or #2166 #2033) +#3957 := [def-axiom]: #3870 +#7294 := [unit-resolution #3957 #6009]: #2166 +#2421 := (not #2166) +#4707 := (or #2421 #4704) +#4710 := (not #4707) +#4362 := (forall (vars (?x27 T2)) (:pat #4355) #1313) +#4367 := (not #4362) +#4713 := (or #4367 #4710) +#4716 := (not #4713) +decl ?x27!0 :: T2 +#2011 := ?x27!0 +#2012 := (uf_12 ?x27!0) +#2013 := (>= #2012 0::int) +#2014 := (not #2013) +#4719 := (or #2014 #4716) +#4722 := (not #4719) +#4725 := (or #1308 #4722) +#4728 := (not #4725) +#4735 := (forall (vars (?x24 T2)) (:pat #4355) #1912) +#4738 := (iff #1915 #4735) +#4736 := (iff #1912 #1912) +#4737 := [refl]: #4736 +#4739 := [quant-intro #4737]: #4738 +#2238 := (~ #1915 #1915) +#2072 := (~ #1912 #1912) +#2073 := [refl]: #2072 +#2239 := [nnf-pos #2073]: #2238 +#1945 := [not-or-elim #1940]: #1915 +#2265 := [mp~ #1945 #2239]: #1915 +#4740 := [mp #2265 #4739]: #4735 +#5485 := [hypothesis]: #1308 +#5347 := (not #4735) +#5348 := (or #5347 #81) +#4030 := (= uf_11 uf_11) +#4033 := (not #4030) +#4034 := (or #4033 #81) +#5341 := (or #5347 #4034) +#5350 := (iff #5341 #5348) +#5361 := (iff #5348 #5348) +#5363 := [rewrite]: #5361 +#5108 := (iff #4034 #81) +#5046 := (or false #81) +#5106 := (iff #5046 #81) +#5107 := [rewrite]: #5106 +#5104 := (iff #4034 #5046) +#3956 := (iff #4033 false) +#6849 := (not true) +#6852 := (iff #6849 false) +#6853 := [rewrite]: #6852 +#5051 := (iff #4033 #6849) +#4038 := (iff #4030 true) +#5050 := [rewrite]: #4038 +#5052 := [monotonicity #5050]: #5051 +#3939 := [trans #5052 #6853]: #3956 +#5105 := [monotonicity #3939]: #5104 +#5346 := [trans #5105 #5107]: #5108 +#5351 := [monotonicity #5346]: #5350 +#5364 := [trans #5351 #5363]: #5350 +#5349 := [quant-inst]: #5341 +#5484 := [mp #5349 #5364]: #5348 +#5486 := [unit-resolution #5484 #5485 #4740]: false +#5487 := [lemma #5486]: #81 +#4731 := (or #1308 #4728) +#3666 := (forall (vars (?x66 T2) (?x67 T2)) #3661) +#3672 := (not #3666) +#3673 := (or #251 #3672) +#3674 := (not #3673) +#3701 := (or #3674 #3698) +#3708 := (not #3701) +#3644 := (forall (vars (?x64 T2)) #3639) +#3707 := (not #3644) +#3709 := (or #3707 #3708) +#3710 := (not #3709) +#3607 := (forall (vars (?x65 T2)) #3596) +#3613 := (not #3607) +#3614 := (or #2460 #2993 #3613) +#3615 := (not #3614) +#3715 := (or #3615 #3710) +#3722 := (not #3715) +#3592 := (forall (vars (?x41 T2)) #3581) +#3721 := (not #3592) +#3723 := (or #1062 #1044 #1071 #1053 #3721 #3722) +#3724 := (not #3723) +#3486 := (forall (vars (?x53 T2)) #3475) +#3493 := (not #3486) +#3471 := (forall (vars (?x50 T2) (?x51 T2)) #3466) +#3492 := (not #3471) +#3494 := (or #2318 #2893 #3492 #3493) +#3495 := (not #3494) +#3500 := (or #3449 #3495) +#3507 := (not #3500) +#3426 := (forall (vars (?x48 T2) (?x49 T2)) #3421) +#3506 := (not #3426) +#3508 := (or #3506 #3507) +#3509 := (not #3508) +#3514 := (or #3403 #3509) +#3520 := (not #3514) +#3521 := (or #1501 #3520) +#3522 := (not #3521) +#3527 := (or #2235 #3522) +#3533 := (not #3527) +#3534 := (or #1492 #3533) +#3535 := (not #3534) +#3540 := (or #1492 #3535) +#3546 := (not #3540) +#3547 := (or #824 #3546) +#3548 := (not #3547) +#3553 := (or #2836 #3548) +#3559 := (not #3553) +#3560 := (or #1487 #3559) +#3561 := (not #3560) +#3566 := (or #2822 #3561) +#3574 := (not #3566) +#3380 := (forall (vars (?x43 T2)) #3375) +#3573 := (not #3380) +#3362 := (forall (vars (?x44 T2)) #3359) +#3572 := (not #3362) +#3575 := (or #876 #585 #1652 #1656 #2412 #2415 #3572 #3573 #3574) +#3576 := (not #3575) +#3729 := (or #3576 #3724) +#3739 := (not #3729) +#3352 := (forall (vars (?x39 T2)) #3347) +#3738 := (not #3352) +#3324 := (forall (vars (?x32 T2)) #3319) +#3737 := (not #3324) +#3296 := (forall (vars (?x35 T2) (?x36 T2)) #3291) +#3736 := (not #3296) +#3273 := (forall (vars (?x37 T2) (?x38 T2)) #3268) +#3735 := (not #3273) +#3740 := (or #1854 #1849 #3735 #3736 #3737 #3738 #3739) +#3741 := (not #3740) +#3241 := (forall (vars (?x33 T2)) #3230) +#3247 := (not #3241) +#3248 := (or #2085 #2688 #3247) +#3249 := (not #3248) +#3746 := (or #3249 #3741) +#3753 := (not #3746) +#3226 := (forall (vars (?x30 T2) (?x31 T2)) #3221) +#3752 := (not #3226) +#3754 := (or #3752 #3753) +#3755 := (not #3754) +#3760 := (or #3204 #3755) +#3767 := (not #3760) +#3181 := (forall (vars (?x28 T2) (?x29 T2)) #3176) +#3766 := (not #3181) +#3768 := (or #3766 #3767) +#3769 := (not #3768) +#3774 := (or #2421 #3769) +#3780 := (not #3774) +#3781 := (or #1317 #3780) +#3782 := (not #3781) +#3787 := (or #2014 #3782) +#3793 := (not #3787) +#3794 := (or #1308 #3793) +#3795 := (not #3794) +#3800 := (or #1308 #3795) +#4732 := (iff #3800 #4731) +#4729 := (iff #3795 #4728) +#4726 := (iff #3794 #4725) +#4723 := (iff #3793 #4722) +#4720 := (iff #3787 #4719) +#4717 := (iff #3782 #4716) +#4714 := (iff #3781 #4713) +#4711 := (iff #3780 #4710) +#4708 := (iff #3774 #4707) +#4705 := (iff #3769 #4704) +#4702 := (iff #3768 #4701) +#4699 := (iff #3767 #4698) +#4696 := (iff #3760 #4695) +#4693 := (iff #3755 #4692) +#4690 := (iff #3754 #4689) +#4687 := (iff #3753 #4686) +#4684 := (iff #3746 #4683) +#4681 := (iff #3741 #4680) +#4678 := (iff #3740 #4677) +#4675 := (iff #3739 #4674) +#4672 := (iff #3729 #4671) +#4669 := (iff #3724 #4668) +#4666 := (iff #3723 #4665) +#4663 := (iff #3722 #4662) +#4660 := (iff #3715 #4659) +#4657 := (iff #3710 #4656) +#4654 := (iff #3709 #4653) +#4651 := (iff #3708 #4650) +#4648 := (iff #3701 #4647) +#4645 := (iff #3674 #4644) +#4642 := (iff #3673 #4641) +#4639 := (iff #3672 #4638) +#4636 := (iff #3666 #4633) +#4634 := (iff #3661 #3661) +#4635 := [refl]: #4634 +#4637 := [quant-intro #4635]: #4636 +#4640 := [monotonicity #4637]: #4639 +#4643 := [monotonicity #4640]: #4642 +#4646 := [monotonicity #4643]: #4645 +#4649 := [monotonicity #4646]: #4648 +#4652 := [monotonicity #4649]: #4651 +#4631 := (iff #3707 #4630) +#4628 := (iff #3644 #4625) +#4626 := (iff #3639 #3639) +#4627 := [refl]: #4626 +#4629 := [quant-intro #4627]: #4628 +#4632 := [monotonicity #4629]: #4631 +#4655 := [monotonicity #4632 #4652]: #4654 +#4658 := [monotonicity #4655]: #4657 +#4623 := (iff #3615 #4622) +#4620 := (iff #3614 #4619) +#4617 := (iff #3613 #4616) +#4614 := (iff #3607 #4611) +#4612 := (iff #3596 #3596) +#4613 := [refl]: #4612 +#4615 := [quant-intro #4613]: #4614 +#4618 := [monotonicity #4615]: #4617 +#4621 := [monotonicity #4618]: #4620 +#4624 := [monotonicity #4621]: #4623 +#4661 := [monotonicity #4624 #4658]: #4660 +#4664 := [monotonicity #4661]: #4663 +#4607 := (iff #3721 #4606) +#4604 := (iff #3592 #4601) +#4602 := (iff #3581 #3581) +#4603 := [refl]: #4602 +#4605 := [quant-intro #4603]: #4604 +#4608 := [monotonicity #4605]: #4607 +#4667 := [monotonicity #4608 #4664]: #4666 +#4670 := [monotonicity #4667]: #4669 +#4599 := (iff #3576 #4598) +#4596 := (iff #3575 #4595) +#4593 := (iff #3574 #4592) +#4590 := (iff #3566 #4589) +#4587 := (iff #3561 #4586) +#4584 := (iff #3560 #4583) +#4581 := (iff #3559 #4580) +#4578 := (iff #3553 #4577) +#4575 := (iff #3548 #4574) +#4572 := (iff #3547 #4571) +#4569 := (iff #3546 #4568) +#4566 := (iff #3540 #4565) +#4563 := (iff #3535 #4562) +#4560 := (iff #3534 #4559) +#4557 := (iff #3533 #4556) +#4554 := (iff #3527 #4553) +#4551 := (iff #3522 #4550) +#4548 := (iff #3521 #4547) +#4545 := (iff #3520 #4544) +#4542 := (iff #3514 #4541) +#4539 := (iff #3509 #4538) +#4536 := (iff #3508 #4535) +#4533 := (iff #3507 #4532) +#4530 := (iff #3500 #4529) +#4527 := (iff #3495 #4526) +#4524 := (iff #3494 #4523) +#4521 := (iff #3493 #4520) +#4518 := (iff #3486 #4515) +#4516 := (iff #3475 #3475) +#4517 := [refl]: #4516 +#4519 := [quant-intro #4517]: #4518 +#4522 := [monotonicity #4519]: #4521 +#4512 := (iff #3492 #4511) +#4509 := (iff #3471 #4506) +#4507 := (iff #3466 #3466) +#4508 := [refl]: #4507 +#4510 := [quant-intro #4508]: #4509 +#4513 := [monotonicity #4510]: #4512 +#4525 := [monotonicity #4513 #4522]: #4524 +#4528 := [monotonicity #4525]: #4527 +#4531 := [monotonicity #4528]: #4530 +#4534 := [monotonicity #4531]: #4533 +#4504 := (iff #3506 #4503) +#4501 := (iff #3426 #4498) +#4499 := (iff #3421 #3421) +#4500 := [refl]: #4499 +#4502 := [quant-intro #4500]: #4501 +#4505 := [monotonicity #4502]: #4504 +#4537 := [monotonicity #4505 #4534]: #4536 +#4540 := [monotonicity #4537]: #4539 +#4543 := [monotonicity #4540]: #4542 +#4546 := [monotonicity #4543]: #4545 +#4495 := (iff #1501 #4494) +#4492 := (iff #1498 #4489) +#4490 := (iff #1495 #1495) +#4491 := [refl]: #4490 +#4493 := [quant-intro #4491]: #4492 +#4496 := [monotonicity #4493]: #4495 +#4549 := [monotonicity #4496 #4546]: #4548 +#4552 := [monotonicity #4549]: #4551 +#4555 := [monotonicity #4552]: #4554 +#4558 := [monotonicity #4555]: #4557 +#4561 := [monotonicity #4558]: #4560 +#4564 := [monotonicity #4561]: #4563 +#4567 := [monotonicity #4564]: #4566 +#4570 := [monotonicity #4567]: #4569 +#4487 := (iff #824 #4486) +#4484 := (iff #642 #4481) +#4482 := (iff #637 #637) +#4483 := [refl]: #4482 +#4485 := [quant-intro #4483]: #4484 +#4488 := [monotonicity #4485]: #4487 +#4573 := [monotonicity #4488 #4570]: #4572 +#4576 := [monotonicity #4573]: #4575 +#4579 := [monotonicity #4576]: #4578 +#4582 := [monotonicity #4579]: #4581 +#4478 := (iff #1487 #4477) +#4475 := (iff #1484 #4472) +#4473 := (iff #1479 #1479) +#4474 := [refl]: #4473 +#4476 := [quant-intro #4474]: #4475 +#4479 := [monotonicity #4476]: #4478 +#4585 := [monotonicity #4479 #4582]: #4584 +#4588 := [monotonicity #4585]: #4587 +#4591 := [monotonicity #4588]: #4590 +#4594 := [monotonicity #4591]: #4593 +#4470 := (iff #3573 #4469) +#4467 := (iff #3380 #4464) +#4465 := (iff #3375 #3375) +#4466 := [refl]: #4465 +#4468 := [quant-intro #4466]: #4467 +#4471 := [monotonicity #4468]: #4470 +#4462 := (iff #3572 #4461) +#4459 := (iff #3362 #4456) +#4457 := (iff #3359 #3359) +#4458 := [refl]: #4457 +#4460 := [quant-intro #4458]: #4459 +#4463 := [monotonicity #4460]: #4462 +#4452 := (iff #1652 #4451) +#4449 := (iff #1649 #4446) +#4447 := (iff #1644 #1644) +#4448 := [refl]: #4447 +#4450 := [quant-intro #4448]: #4449 +#4453 := [monotonicity #4450]: #4452 +#4597 := [monotonicity #4453 #4463 #4471 #4594]: #4596 +#4600 := [monotonicity #4597]: #4599 +#4673 := [monotonicity #4600 #4670]: #4672 +#4676 := [monotonicity #4673]: #4675 +#4443 := (iff #3738 #4442) +#4440 := (iff #3352 #4437) +#4438 := (iff #3347 #3347) +#4439 := [refl]: #4438 +#4441 := [quant-intro #4439]: #4440 +#4444 := [monotonicity #4441]: #4443 +#4435 := (iff #3737 #4434) +#4432 := (iff #3324 #4429) +#4430 := (iff #3319 #3319) +#4431 := [refl]: #4430 +#4433 := [quant-intro #4431]: #4432 +#4436 := [monotonicity #4433]: #4435 +#4427 := (iff #3736 #4426) +#4424 := (iff #3296 #4421) +#4422 := (iff #3291 #3291) +#4423 := [refl]: #4422 +#4425 := [quant-intro #4423]: #4424 +#4428 := [monotonicity #4425]: #4427 +#4418 := (iff #3735 #4417) +#4415 := (iff #3273 #4412) +#4413 := (iff #3268 #3268) +#4414 := [refl]: #4413 +#4416 := [quant-intro #4414]: #4415 +#4419 := [monotonicity #4416]: #4418 +#4410 := (iff #1849 #4409) +#4407 := (iff #1846 #4404) +#4405 := (iff #1843 #1843) +#4406 := [refl]: #4405 +#4408 := [quant-intro #4406]: #4407 +#4411 := [monotonicity #4408]: #4410 +#4679 := [monotonicity #4411 #4419 #4428 #4436 #4444 #4676]: #4678 +#4682 := [monotonicity #4679]: #4681 +#4401 := (iff #3249 #4400) +#4398 := (iff #3248 #4397) +#4395 := (iff #3247 #4394) +#4392 := (iff #3241 #4389) +#4390 := (iff #3230 #3230) +#4391 := [refl]: #4390 +#4393 := [quant-intro #4391]: #4392 +#4396 := [monotonicity #4393]: #4395 +#4399 := [monotonicity #4396]: #4398 +#4402 := [monotonicity #4399]: #4401 +#4685 := [monotonicity #4402 #4682]: #4684 +#4688 := [monotonicity #4685]: #4687 +#4386 := (iff #3752 #4385) +#4383 := (iff #3226 #4380) +#4381 := (iff #3221 #3221) +#4382 := [refl]: #4381 +#4384 := [quant-intro #4382]: #4383 +#4387 := [monotonicity #4384]: #4386 +#4691 := [monotonicity #4387 #4688]: #4690 +#4694 := [monotonicity #4691]: #4693 +#4697 := [monotonicity #4694]: #4696 +#4700 := [monotonicity #4697]: #4699 +#4377 := (iff #3766 #4376) +#4374 := (iff #3181 #4371) +#4372 := (iff #3176 #3176) +#4373 := [refl]: #4372 +#4375 := [quant-intro #4373]: #4374 +#4378 := [monotonicity #4375]: #4377 +#4703 := [monotonicity #4378 #4700]: #4702 +#4706 := [monotonicity #4703]: #4705 +#4709 := [monotonicity #4706]: #4708 +#4712 := [monotonicity #4709]: #4711 +#4368 := (iff #1317 #4367) +#4365 := (iff #1314 #4362) +#4363 := (iff #1313 #1313) +#4364 := [refl]: #4363 +#4366 := [quant-intro #4364]: #4365 +#4369 := [monotonicity #4366]: #4368 +#4715 := [monotonicity #4369 #4712]: #4714 +#4718 := [monotonicity #4715]: #4717 +#4721 := [monotonicity #4718]: #4720 +#4724 := [monotonicity #4721]: #4723 +#4727 := [monotonicity #4724]: #4726 +#4730 := [monotonicity #4727]: #4729 +#4733 := [monotonicity #4730]: #4732 +#2527 := (not #2526) +#2523 := (not #2522) +#3081 := (and #2523 #2527) +#3084 := (not #3081) +#3101 := (or #3084 #3096) +#3104 := (not #3101) +#2536 := (not #251) +#2546 := (and #2536 #1756) +#3110 := (or #2546 #3104) +#3054 := (not #3049) +#3072 := (and #3054 #3067) +#3075 := (or #1716 #3072) +#3078 := (forall (vars (?x64 T2)) #3075) +#3115 := (and #3078 #3110) +#2456 := (not #2455) +#3024 := (and #2456 #3021) +#3027 := (not #3024) +#3030 := (forall (vars (?x65 T2)) #3027) +#2996 := (not #2993) +#2461 := (not #2460) +#3036 := (and #2461 #2996 #3030) +#3118 := (or #3036 #3115) +#2441 := (not #1670) +#2444 := (forall (vars (?x41 T2)) #2441) +#3121 := (and #229 #234 #930 #933 #2444 #3118) +#2314 := (not #2313) +#2924 := (and #630 #2314 #2921) +#2927 := (not #2924) +#2930 := (forall (vars (?x53 T2)) #2927) +#2896 := (not #2893) +#2319 := (not #2318) +#2939 := (and #1541 #2319 #2896 #2930) +#2287 := (not #2286) +#2866 := (and #2287 #2289) +#2869 := (not #2866) +#2887 := (or #2869 #2882) +#2890 := (not #2887) +#2944 := (or #2890 #2939) +#2947 := (and #1517 #2944) +#2258 := (not #2257) +#2841 := (and #2255 #2258) +#2844 := (not #2841) +#2860 := (or #2844 #2855) +#2863 := (not #2860) +#2950 := (or #2863 #2947) +#2953 := (and #1498 #2950) +#2956 := (or #2235 #2953) +#2959 := (and #173 #2956) +#2962 := (or #1492 #2959) +#2965 := (and #642 #2962) +#2968 := (or #2836 #2965) +#2971 := (and #1484 #2968) +#2974 := (or #2822 #2971) +#2416 := (not #2415) +#2413 := (not #2412) +#2980 := (and #151 #588 #1473 #1631 #1649 #1657 #2413 #2416 #2974) +#3124 := (or #2980 #3121) +#2780 := (not #2775) +#2798 := (and #2162 #2780 #2793) +#2801 := (or #1413 #2798) +#2804 := (forall (vars (?x39 T2)) #2801) +#2735 := (not #2730) +#2753 := (and #2133 #2735 #2748) +#2759 := (or #1395 #2753) +#2764 := (forall (vars (?x32 T2)) #2759) +#3130 := (and #109 #1821 #1837 #1846 #2764 #2804 #3124) +#2691 := (not #2688) +#2098 := (not #2097) +#2679 := (and #77 #2095 #2098) +#2682 := (not #2679) +#2685 := (forall (vars (?x33 T2)) #2682) +#2086 := (not #2085) +#2715 := (and #2086 #2685 #2691) +#3135 := (or #2715 #3130) +#3138 := (and #1360 #3135) +#2058 := (not #2057) +#2667 := (and #2058 #2059) +#2670 := (not #2667) +#2673 := (or #2664 #2670) +#2676 := (not #2673) +#3141 := (or #2676 #3138) +#3144 := (and #1334 #3141) +#2035 := (not #2034) +#2635 := (and #2033 #2035) +#2638 := (not #2635) +#2654 := (or #2638 #2649) +#2657 := (not #2654) +#3147 := (or #2657 #3144) +#3150 := (and #1314 #3147) +#3153 := (or #2014 #3150) +#3156 := (and #81 #3153) +#3159 := (or #1308 #3156) +#3801 := (iff #3159 #3800) +#3798 := (iff #3156 #3795) +#3790 := (and #81 #3787) +#3796 := (iff #3790 #3795) +#3797 := [rewrite]: #3796 +#3791 := (iff #3156 #3790) +#3788 := (iff #3153 #3787) +#3785 := (iff #3150 #3782) +#3777 := (and #1314 #3774) +#3783 := (iff #3777 #3782) +#3784 := [rewrite]: #3783 +#3778 := (iff #3150 #3777) +#3775 := (iff #3147 #3774) +#3772 := (iff #3144 #3769) +#3763 := (and #3181 #3760) +#3770 := (iff #3763 #3769) +#3771 := [rewrite]: #3770 +#3764 := (iff #3144 #3763) +#3761 := (iff #3141 #3760) +#3758 := (iff #3138 #3755) +#3749 := (and #3226 #3746) +#3756 := (iff #3749 #3755) +#3757 := [rewrite]: #3756 +#3750 := (iff #3138 #3749) +#3747 := (iff #3135 #3746) +#3744 := (iff #3130 #3741) +#3732 := (and #109 #3273 #3296 #1846 #3324 #3352 #3729) +#3742 := (iff #3732 #3741) +#3743 := [rewrite]: #3742 +#3733 := (iff #3130 #3732) +#3730 := (iff #3124 #3729) +#3727 := (iff #3121 #3724) +#3718 := (and #229 #234 #930 #933 #3592 #3715) +#3725 := (iff #3718 #3724) +#3726 := [rewrite]: #3725 +#3719 := (iff #3121 #3718) +#3716 := (iff #3118 #3715) +#3713 := (iff #3115 #3710) +#3704 := (and #3644 #3701) +#3711 := (iff #3704 #3710) +#3712 := [rewrite]: #3711 +#3705 := (iff #3115 #3704) +#3702 := (iff #3110 #3701) +#3699 := (iff #3104 #3698) +#3696 := (iff #3101 #3693) +#3679 := (or #2522 #2526) +#3690 := (or #3679 #3096) +#3694 := (iff #3690 #3693) +#3695 := [rewrite]: #3694 +#3691 := (iff #3101 #3690) +#3688 := (iff #3084 #3679) +#3680 := (not #3679) +#3683 := (not #3680) +#3686 := (iff #3683 #3679) +#3687 := [rewrite]: #3686 +#3684 := (iff #3084 #3683) +#3681 := (iff #3081 #3680) +#3682 := [rewrite]: #3681 +#3685 := [monotonicity #3682]: #3684 +#3689 := [trans #3685 #3687]: #3688 +#3692 := [monotonicity #3689]: #3691 +#3697 := [trans #3692 #3695]: #3696 +#3700 := [monotonicity #3697]: #3699 +#3677 := (iff #2546 #3674) +#3669 := (and #2536 #3666) +#3675 := (iff #3669 #3674) +#3676 := [rewrite]: #3675 +#3670 := (iff #2546 #3669) +#3667 := (iff #1756 #3666) +#3664 := (iff #1753 #3661) +#3647 := (or #1347 #1709) +#3658 := (or #3647 #1750) +#3662 := (iff #3658 #3661) +#3663 := [rewrite]: #3662 +#3659 := (iff #1753 #3658) +#3656 := (iff #1747 #3647) +#3648 := (not #3647) +#3651 := (not #3648) +#3654 := (iff #3651 #3647) +#3655 := [rewrite]: #3654 +#3652 := (iff #1747 #3651) +#3649 := (iff #1744 #3648) +#3650 := [rewrite]: #3649 +#3653 := [monotonicity #3650]: #3652 +#3657 := [trans #3653 #3655]: #3656 +#3660 := [monotonicity #3657]: #3659 +#3665 := [trans #3660 #3663]: #3664 +#3668 := [quant-intro #3665]: #3667 +#3671 := [monotonicity #3668]: #3670 +#3678 := [trans #3671 #3676]: #3677 +#3703 := [monotonicity #3678 #3700]: #3702 +#3645 := (iff #3078 #3644) +#3642 := (iff #3075 #3639) +#3620 := (or #68 #1709) +#3636 := (or #3620 #3633) +#3640 := (iff #3636 #3639) +#3641 := [rewrite]: #3640 +#3637 := (iff #3075 #3636) +#3634 := (iff #3072 #3633) +#3635 := [rewrite]: #3634 +#3629 := (iff #1716 #3620) +#3621 := (not #3620) +#3624 := (not #3621) +#3627 := (iff #3624 #3620) +#3628 := [rewrite]: #3627 +#3625 := (iff #1716 #3624) +#3622 := (iff #1713 #3621) +#3623 := [rewrite]: #3622 +#3626 := [monotonicity #3623]: #3625 +#3630 := [trans #3626 #3628]: #3629 +#3638 := [monotonicity #3630 #3635]: #3637 +#3643 := [trans #3638 #3641]: #3642 +#3646 := [quant-intro #3643]: #3645 +#3706 := [monotonicity #3646 #3703]: #3705 +#3714 := [trans #3706 #3712]: #3713 +#3618 := (iff #3036 #3615) +#3610 := (and #2461 #2996 #3607) +#3616 := (iff #3610 #3615) +#3617 := [rewrite]: #3616 +#3611 := (iff #3036 #3610) +#3608 := (iff #3030 #3607) +#3605 := (iff #3027 #3596) +#3597 := (not #3596) +#3600 := (not #3597) +#3603 := (iff #3600 #3596) +#3604 := [rewrite]: #3603 +#3601 := (iff #3027 #3600) +#3598 := (iff #3024 #3597) +#3599 := [rewrite]: #3598 +#3602 := [monotonicity #3599]: #3601 +#3606 := [trans #3602 #3604]: #3605 +#3609 := [quant-intro #3606]: #3608 +#3612 := [monotonicity #3609]: #3611 +#3619 := [trans #3612 #3617]: #3618 +#3717 := [monotonicity #3619 #3714]: #3716 +#3593 := (iff #2444 #3592) +#3590 := (iff #2441 #3581) +#3582 := (not #3581) +#3585 := (not #3582) +#3588 := (iff #3585 #3581) +#3589 := [rewrite]: #3588 +#3586 := (iff #2441 #3585) +#3583 := (iff #1670 #3582) +#3584 := [rewrite]: #3583 +#3587 := [monotonicity #3584]: #3586 +#3591 := [trans #3587 #3589]: #3590 +#3594 := [quant-intro #3591]: #3593 +#3720 := [monotonicity #3594 #3717]: #3719 +#3728 := [trans #3720 #3726]: #3727 +#3579 := (iff #2980 #3576) +#3569 := (and #151 #588 #3362 #3380 #1649 #1657 #2413 #2416 #3566) +#3577 := (iff #3569 #3576) +#3578 := [rewrite]: #3577 +#3570 := (iff #2980 #3569) +#3567 := (iff #2974 #3566) +#3564 := (iff #2971 #3561) +#3556 := (and #1484 #3553) +#3562 := (iff #3556 #3561) +#3563 := [rewrite]: #3562 +#3557 := (iff #2971 #3556) +#3554 := (iff #2968 #3553) +#3551 := (iff #2965 #3548) +#3543 := (and #642 #3540) +#3549 := (iff #3543 #3548) +#3550 := [rewrite]: #3549 +#3544 := (iff #2965 #3543) +#3541 := (iff #2962 #3540) +#3538 := (iff #2959 #3535) +#3530 := (and #173 #3527) +#3536 := (iff #3530 #3535) +#3537 := [rewrite]: #3536 +#3531 := (iff #2959 #3530) +#3528 := (iff #2956 #3527) +#3525 := (iff #2953 #3522) +#3517 := (and #1498 #3514) +#3523 := (iff #3517 #3522) +#3524 := [rewrite]: #3523 +#3518 := (iff #2953 #3517) +#3515 := (iff #2950 #3514) +#3512 := (iff #2947 #3509) +#3503 := (and #3426 #3500) +#3510 := (iff #3503 #3509) +#3511 := [rewrite]: #3510 +#3504 := (iff #2947 #3503) +#3501 := (iff #2944 #3500) +#3498 := (iff #2939 #3495) +#3489 := (and #3471 #2319 #2896 #3486) +#3496 := (iff #3489 #3495) +#3497 := [rewrite]: #3496 +#3490 := (iff #2939 #3489) +#3487 := (iff #2930 #3486) +#3484 := (iff #2927 #3475) +#3476 := (not #3475) +#3479 := (not #3476) +#3482 := (iff #3479 #3475) +#3483 := [rewrite]: #3482 +#3480 := (iff #2927 #3479) +#3477 := (iff #2924 #3476) +#3478 := [rewrite]: #3477 +#3481 := [monotonicity #3478]: #3480 +#3485 := [trans #3481 #3483]: #3484 +#3488 := [quant-intro #3485]: #3487 +#3472 := (iff #1541 #3471) +#3469 := (iff #1538 #3466) +#3452 := (or #636 #1347) +#3463 := (or #3452 #1534) +#3467 := (iff #3463 #3466) +#3468 := [rewrite]: #3467 +#3464 := (iff #1538 #3463) +#3461 := (iff #1531 #3452) +#3453 := (not #3452) +#3456 := (not #3453) +#3459 := (iff #3456 #3452) +#3460 := [rewrite]: #3459 +#3457 := (iff #1531 #3456) +#3454 := (iff #1526 #3453) +#3455 := [rewrite]: #3454 +#3458 := [monotonicity #3455]: #3457 +#3462 := [trans #3458 #3460]: #3461 +#3465 := [monotonicity #3462]: #3464 +#3470 := [trans #3465 #3468]: #3469 +#3473 := [quant-intro #3470]: #3472 +#3491 := [monotonicity #3473 #3488]: #3490 +#3499 := [trans #3491 #3497]: #3498 +#3450 := (iff #2890 #3449) +#3447 := (iff #2887 #3444) +#3430 := (or #2286 #3429) +#3441 := (or #3430 #2882) +#3445 := (iff #3441 #3444) +#3446 := [rewrite]: #3445 +#3442 := (iff #2887 #3441) +#3439 := (iff #2869 #3430) +#3431 := (not #3430) +#3434 := (not #3431) +#3437 := (iff #3434 #3430) +#3438 := [rewrite]: #3437 +#3435 := (iff #2869 #3434) +#3432 := (iff #2866 #3431) +#3433 := [rewrite]: #3432 +#3436 := [monotonicity #3433]: #3435 +#3440 := [trans #3436 #3438]: #3439 +#3443 := [monotonicity #3440]: #3442 +#3448 := [trans #3443 #3446]: #3447 +#3451 := [monotonicity #3448]: #3450 +#3502 := [monotonicity #3451 #3499]: #3501 +#3427 := (iff #1517 #3426) +#3424 := (iff #1512 #3421) +#3407 := (or #630 #3406) +#3418 := (or #3407 #1504) +#3422 := (iff #3418 #3421) +#3423 := [rewrite]: #3422 +#3419 := (iff #1512 #3418) +#3416 := (iff #664 #3407) +#3408 := (not #3407) +#3411 := (not #3408) +#3414 := (iff #3411 #3407) +#3415 := [rewrite]: #3414 +#3412 := (iff #664 #3411) +#3409 := (iff #656 #3408) +#3410 := [rewrite]: #3409 +#3413 := [monotonicity #3410]: #3412 +#3417 := [trans #3413 #3415]: #3416 +#3420 := [monotonicity #3417]: #3419 +#3425 := [trans #3420 #3423]: #3424 +#3428 := [quant-intro #3425]: #3427 +#3505 := [monotonicity #3428 #3502]: #3504 +#3513 := [trans #3505 #3511]: #3512 +#3404 := (iff #2863 #3403) +#3401 := (iff #2860 #3398) +#3384 := (or #3383 #2257) +#3395 := (or #3384 #2855) +#3399 := (iff #3395 #3398) +#3400 := [rewrite]: #3399 +#3396 := (iff #2860 #3395) +#3393 := (iff #2844 #3384) +#3385 := (not #3384) +#3388 := (not #3385) +#3391 := (iff #3388 #3384) +#3392 := [rewrite]: #3391 +#3389 := (iff #2844 #3388) +#3386 := (iff #2841 #3385) +#3387 := [rewrite]: #3386 +#3390 := [monotonicity #3387]: #3389 +#3394 := [trans #3390 #3392]: #3393 +#3397 := [monotonicity #3394]: #3396 +#3402 := [trans #3397 #3400]: #3401 +#3405 := [monotonicity #3402]: #3404 +#3516 := [monotonicity #3405 #3513]: #3515 +#3519 := [monotonicity #3516]: #3518 +#3526 := [trans #3519 #3524]: #3525 +#3529 := [monotonicity #3526]: #3528 +#3532 := [monotonicity #3529]: #3531 +#3539 := [trans #3532 #3537]: #3538 +#3542 := [monotonicity #3539]: #3541 +#3545 := [monotonicity #3542]: #3544 +#3552 := [trans #3545 #3550]: #3551 +#3555 := [monotonicity #3552]: #3554 +#3558 := [monotonicity #3555]: #3557 +#3565 := [trans #3558 #3563]: #3564 +#3568 := [monotonicity #3565]: #3567 +#3381 := (iff #1631 #3380) +#3378 := (iff #1628 #3375) +#3372 := (or #3355 #1625) +#3376 := (iff #3372 #3375) +#3377 := [rewrite]: #3376 +#3373 := (iff #1628 #3372) +#3370 := (iff #1620 #3355) +#3365 := (not #3356) +#3368 := (iff #3365 #3355) +#3369 := [rewrite]: #3368 +#3366 := (iff #1620 #3365) +#3357 := (iff #1462 #3356) +#3358 := [rewrite]: #3357 +#3367 := [monotonicity #3358]: #3366 +#3371 := [trans #3367 #3369]: #3370 +#3374 := [monotonicity #3371]: #3373 +#3379 := [trans #3374 #3377]: #3378 +#3382 := [quant-intro #3379]: #3381 +#3363 := (iff #1473 #3362) +#3360 := (iff #1468 #3359) +#3361 := [monotonicity #3358]: #3360 +#3364 := [quant-intro #3361]: #3363 +#3571 := [monotonicity #3364 #3382 #3568]: #3570 +#3580 := [trans #3571 #3578]: #3579 +#3731 := [monotonicity #3580 #3728]: #3730 +#3353 := (iff #2804 #3352) +#3350 := (iff #2801 #3347) +#3327 := (or #68 #1406) +#3344 := (or #3327 #3341) +#3348 := (iff #3344 #3347) +#3349 := [rewrite]: #3348 +#3345 := (iff #2801 #3344) +#3342 := (iff #2798 #3341) +#3343 := [rewrite]: #3342 +#3336 := (iff #1413 #3327) +#3328 := (not #3327) +#3331 := (not #3328) +#3334 := (iff #3331 #3327) +#3335 := [rewrite]: #3334 +#3332 := (iff #1413 #3331) +#3329 := (iff #1410 #3328) +#3330 := [rewrite]: #3329 +#3333 := [monotonicity #3330]: #3332 +#3337 := [trans #3333 #3335]: #3336 +#3346 := [monotonicity #3337 #3343]: #3345 +#3351 := [trans #3346 #3349]: #3350 +#3354 := [quant-intro #3351]: #3353 +#3325 := (iff #2764 #3324) +#3322 := (iff #2759 #3319) +#3299 := (or #68 #1388) +#3316 := (or #3299 #3313) +#3320 := (iff #3316 #3319) +#3321 := [rewrite]: #3320 +#3317 := (iff #2759 #3316) +#3314 := (iff #2753 #3313) +#3315 := [rewrite]: #3314 +#3308 := (iff #1395 #3299) +#3300 := (not #3299) +#3303 := (not #3300) +#3306 := (iff #3303 #3299) +#3307 := [rewrite]: #3306 +#3304 := (iff #1395 #3303) +#3301 := (iff #1392 #3300) +#3302 := [rewrite]: #3301 +#3305 := [monotonicity #3302]: #3304 +#3309 := [trans #3305 #3307]: #3308 +#3318 := [monotonicity #3309 #3315]: #3317 +#3323 := [trans #3318 #3321]: #3322 +#3326 := [quant-intro #3323]: #3325 +#3297 := (iff #1837 #3296) +#3294 := (iff #1832 #3291) +#3277 := (or #3276 #505) +#3288 := (or #3277 #1416) +#3292 := (iff #3288 #3291) +#3293 := [rewrite]: #3292 +#3289 := (iff #1832 #3288) +#3286 := (iff #517 #3277) +#3278 := (not #3277) +#3281 := (not #3278) +#3284 := (iff #3281 #3277) +#3285 := [rewrite]: #3284 +#3282 := (iff #517 #3281) +#3279 := (iff #511 #3278) +#3280 := [rewrite]: #3279 +#3283 := [monotonicity #3280]: #3282 +#3287 := [trans #3283 #3285]: #3286 +#3290 := [monotonicity #3287]: #3289 +#3295 := [trans #3290 #3293]: #3294 +#3298 := [quant-intro #3295]: #3297 +#3274 := (iff #1821 #3273) +#3271 := (iff #1818 #3268) +#3254 := (or #508 #1347) +#3265 := (or #3254 #1815) +#3269 := (iff #3265 #3268) +#3270 := [rewrite]: #3269 +#3266 := (iff #1818 #3265) +#3263 := (iff #1812 #3254) +#3255 := (not #3254) +#3258 := (not #3255) +#3261 := (iff #3258 #3254) +#3262 := [rewrite]: #3261 +#3259 := (iff #1812 #3258) +#3256 := (iff #1807 #3255) +#3257 := [rewrite]: #3256 +#3260 := [monotonicity #3257]: #3259 +#3264 := [trans #3260 #3262]: #3263 +#3267 := [monotonicity #3264]: #3266 +#3272 := [trans #3267 #3270]: #3271 +#3275 := [quant-intro #3272]: #3274 +#3734 := [monotonicity #3275 #3298 #3326 #3354 #3731]: #3733 +#3745 := [trans #3734 #3743]: #3744 +#3252 := (iff #2715 #3249) +#3244 := (and #2086 #3241 #2691) +#3250 := (iff #3244 #3249) +#3251 := [rewrite]: #3250 +#3245 := (iff #2715 #3244) +#3242 := (iff #2685 #3241) +#3239 := (iff #2682 #3230) +#3231 := (not #3230) +#3234 := (not #3231) +#3237 := (iff #3234 #3230) +#3238 := [rewrite]: #3237 +#3235 := (iff #2682 #3234) +#3232 := (iff #2679 #3231) +#3233 := [rewrite]: #3232 +#3236 := [monotonicity #3233]: #3235 +#3240 := [trans #3236 #3238]: #3239 +#3243 := [quant-intro #3240]: #3242 +#3246 := [monotonicity #3243]: #3245 +#3253 := [trans #3246 #3251]: #3252 +#3748 := [monotonicity #3253 #3745]: #3747 +#3227 := (iff #1360 #3226) +#3224 := (iff #1357 #3221) +#3207 := (or #78 #1347) +#3218 := (or #1342 #3207) +#3222 := (iff #3218 #3221) +#3223 := [rewrite]: #3222 +#3219 := (iff #1357 #3218) +#3216 := (iff #1354 #3207) +#3208 := (not #3207) +#3211 := (not #3208) +#3214 := (iff #3211 #3207) +#3215 := [rewrite]: #3214 +#3212 := (iff #1354 #3211) +#3209 := (iff #1351 #3208) +#3210 := [rewrite]: #3209 +#3213 := [monotonicity #3210]: #3212 +#3217 := [trans #3213 #3215]: #3216 +#3220 := [monotonicity #3217]: #3219 +#3225 := [trans #3220 #3223]: #3224 +#3228 := [quant-intro #3225]: #3227 +#3751 := [monotonicity #3228 #3748]: #3750 +#3759 := [trans #3751 #3757]: #3758 +#3205 := (iff #2676 #3204) +#3202 := (iff #2673 #3199) +#3185 := (or #2057 #3184) +#3196 := (or #2664 #3185) +#3200 := (iff #3196 #3199) +#3201 := [rewrite]: #3200 +#3197 := (iff #2673 #3196) +#3194 := (iff #2670 #3185) +#3186 := (not #3185) +#3189 := (not #3186) +#3192 := (iff #3189 #3185) +#3193 := [rewrite]: #3192 +#3190 := (iff #2670 #3189) +#3187 := (iff #2667 #3186) +#3188 := [rewrite]: #3187 +#3191 := [monotonicity #3188]: #3190 +#3195 := [trans #3191 #3193]: #3194 +#3198 := [monotonicity #3195]: #3197 +#3203 := [trans #3198 #3201]: #3202 +#3206 := [monotonicity #3203]: #3205 +#3762 := [monotonicity #3206 #3759]: #3761 +#3182 := (iff #1334 #3181) +#3179 := (iff #1329 #3176) +#3162 := (or #77 #2632) +#3173 := (or #3162 #1322) +#3177 := (iff #3173 #3176) +#3178 := [rewrite]: #3177 +#3174 := (iff #1329 #3173) +#3171 := (iff #460 #3162) +#3163 := (not #3162) +#3166 := (not #3163) +#3169 := (iff #3166 #3162) +#3170 := [rewrite]: #3169 +#3167 := (iff #460 #3166) +#3164 := (iff #454 #3163) +#3165 := [rewrite]: #3164 +#3168 := [monotonicity #3165]: #3167 +#3172 := [trans #3168 #3170]: #3171 +#3175 := [monotonicity #3172]: #3174 +#3180 := [trans #3175 #3178]: #3179 +#3183 := [quant-intro #3180]: #3182 +#3765 := [monotonicity #3183 #3762]: #3764 +#3773 := [trans #3765 #3771]: #3772 +#2534 := (iff #2657 #2421) +#2199 := (iff #2654 #2166) +#2042 := (or #2266 #2034) +#2297 := (or #2042 #2649) +#2167 := (iff #2297 #2166) +#2198 := [rewrite]: #2167 +#2499 := (iff #2654 #2297) +#2138 := (iff #2638 #2042) +#1973 := (not #2042) +#2218 := (not #1973) +#2018 := (iff #2218 #2042) +#2137 := [rewrite]: #2018 +#2219 := (iff #2638 #2218) +#1974 := (iff #2635 #1973) +#2043 := [rewrite]: #1974 +#2017 := [monotonicity #2043]: #2219 +#2296 := [trans #2017 #2137]: #2138 +#2500 := [monotonicity #2296]: #2499 +#2420 := [trans #2500 #2198]: #2199 +#2535 := [monotonicity #2420]: #2534 +#3776 := [monotonicity #2535 #3773]: #3775 +#3779 := [monotonicity #3776]: #3778 +#3786 := [trans #3779 #3784]: #3785 +#3789 := [monotonicity #3786]: #3788 +#3792 := [monotonicity #3789]: #3791 +#3799 := [trans #3792 #3797]: #3798 +#3802 := [monotonicity #3799]: #3801 +#2513 := (* -1::int #2512) +#2515 := (+ #2514 #2513) +#2518 := (+ #2517 #2515) +#2519 := (>= #2518 0::int) +#2528 := (and #2527 #2523) +#2529 := (not #2528) +#2530 := (or #2529 #2519) +#2531 := (not #2530) +#2550 := (or #2531 #2546) +#2489 := (+ #2488 #1707) +#2492 := (+ #2491 #2489) +#2493 := (= #2492 0::int) +#2494 := (>= #2489 0::int) +#2495 := (not #2494) +#2496 := (and #2495 #2493) +#2501 := (or #1716 #2496) +#2504 := (forall (vars (?x64 T2)) #2501) +#2554 := (and #2504 #2550) +#2453 := (+ #2452 #2450) +#2454 := (= #2453 0::int) +#2457 := (and #2456 #2454) +#2473 := (not #2457) +#2476 := (forall (vars (?x65 T2)) #2473) +#2462 := (= ?x64!17 uf_11) +#2463 := (not #2462) +#2464 := (and #2463 #2461) +#2465 := (not #2464) +#2470 := (not #2465) +#2480 := (and #2470 #2476) +#2558 := (or #2480 #2554) +#2438 := (not #1053) +#2435 := (not #1071) +#2432 := (not #1044) +#2429 := (not #1062) +#2562 := (and #2429 #2432 #2435 #2438 #2444 #2558) +#2417 := (and #2416 #2413) +#2311 := (+ #2310 #2308) +#2312 := (= #2311 0::int) +#2315 := (and #630 #2314 #2312) +#2332 := (not #2315) +#2335 := (forall (vars (?x53 T2)) #2332) +#2320 := (= ?x52!15 uf_11) +#2321 := (not #2320) +#2322 := (and #2321 #2319) +#2323 := (not #2322) +#2329 := (not #2323) +#2339 := (and #2329 #2335) +#2344 := (and #1541 #2339) +#2277 := (* -1::int #2276) +#2279 := (+ #2278 #2277) +#2282 := (+ #2281 #2279) +#2283 := (>= #2282 0::int) +#2290 := (and #2289 #2287) +#2291 := (not #2290) +#2292 := (or #2291 #2283) +#2293 := (not #2292) +#2348 := (or #2293 #2344) +#2352 := (and #1517 #2348) +#2250 := (* -1::int #2249) +#2252 := (+ #2251 #2250) +#2253 := (>= #2252 0::int) +#2259 := (and #2258 #2255) +#2260 := (not #2259) +#2261 := (or #2260 #2253) +#2262 := (not #2261) +#2356 := (or #2262 #2352) +#2360 := (and #1498 #2356) +#2364 := (or #2235 #2360) +#2229 := (not #1492) +#2368 := (and #2229 #2364) +#2372 := (or #1492 #2368) +#2376 := (and #642 #2372) +#2213 := (= #2212 #2211) +#2214 := (or #2213 #2210) +#2215 := (not #2214) +#2380 := (or #2215 #2376) +#2384 := (and #1484 #2380) +#2191 := (* -1::int #2190) +#2193 := (+ #2192 #2191) +#2194 := (>= #2193 0::int) +#2195 := (not #2194) +#2388 := (or #2195 #2384) +#2177 := (not #876) +#2425 := (and #2177 #588 #1473 #2388 #1631 #1649 #1657 #2417) +#2566 := (or #2425 #2562) +#2154 := (+ #2153 #1404) +#2157 := (+ #2156 #2154) +#2158 := (= #2157 0::int) +#2159 := (>= #2154 0::int) +#2160 := (not #2159) +#2163 := (and #2162 #2160 #2158) +#2168 := (or #1413 #2163) +#2171 := (forall (vars (?x39 T2)) #2168) +#2126 := (+ #1386 #2125) +#2128 := (+ #2127 #2126) +#2129 := (= #2128 0::int) +#2130 := (+ #2127 #1386) +#2131 := (>= #2130 0::int) +#2132 := (not #2131) +#2134 := (and #2133 #2132 #2129) +#2141 := (or #2134 #1395) +#2144 := (forall (vars (?x32 T2)) #2141) +#2120 := (not #1854) +#2591 := (and #2120 #2144 #2171 #2566 #1821 #1837 #1846) +#2087 := (= ?x32!5 uf_11) +#2088 := (not #2087) +#2089 := (and #2088 #2086) +#2090 := (not #2089) +#2112 := (not #2090) +#2099 := (and #77 #2098 #2095) +#2105 := (not #2099) +#2108 := (forall (vars (?x33 T2)) #2105) +#2115 := (and #2108 #2112) +#2595 := (or #2115 #2591) +#2599 := (and #1360 #2595) +#2060 := (and #2059 #2058) +#2061 := (not #2060) +#2064 := (+ #2063 #2054) +#2066 := (+ #2065 #2064) +#2067 := (>= #2066 0::int) +#2068 := (or #2067 #2061) +#2069 := (not #2068) +#2603 := (or #2069 #2599) +#2607 := (and #1334 #2603) +#2029 := (* -1::int #2028) +#2031 := (+ #2030 #2029) +#2032 := (>= #2031 0::int) +#2036 := (and #2035 #2033) +#2037 := (not #2036) +#2038 := (or #2037 #2032) +#2039 := (not #2038) +#2611 := (or #2039 #2607) +#2615 := (and #1314 #2611) +#2619 := (or #2014 #2615) +#1969 := (not #1308) +#2623 := (and #1969 #2619) +#2627 := (or #1308 #2623) +#3160 := (iff #2627 #3159) +#3157 := (iff #2623 #3156) +#3154 := (iff #2619 #3153) +#3151 := (iff #2615 #3150) +#3148 := (iff #2611 #3147) +#3145 := (iff #2607 #3144) +#3142 := (iff #2603 #3141) +#3139 := (iff #2599 #3138) +#3136 := (iff #2595 #3135) +#3133 := (iff #2591 #3130) +#3127 := (and #109 #2764 #2804 #3124 #1821 #1837 #1846) +#3131 := (iff #3127 #3130) +#3132 := [rewrite]: #3131 +#3128 := (iff #2591 #3127) +#3125 := (iff #2566 #3124) +#3122 := (iff #2562 #3121) +#3119 := (iff #2558 #3118) +#3116 := (iff #2554 #3115) +#3113 := (iff #2550 #3110) +#3107 := (or #3104 #2546) +#3111 := (iff #3107 #3110) +#3112 := [rewrite]: #3111 +#3108 := (iff #2550 #3107) +#3105 := (iff #2531 #3104) +#3102 := (iff #2530 #3101) +#3099 := (iff #2519 #3096) +#3087 := (+ #2514 #2517) +#3088 := (+ #2513 #3087) +#3091 := (>= #3088 0::int) +#3097 := (iff #3091 #3096) +#3098 := [rewrite]: #3097 +#3092 := (iff #2519 #3091) +#3089 := (= #2518 #3088) +#3090 := [rewrite]: #3089 +#3093 := [monotonicity #3090]: #3092 +#3100 := [trans #3093 #3098]: #3099 +#3085 := (iff #2529 #3084) +#3082 := (iff #2528 #3081) +#3083 := [rewrite]: #3082 +#3086 := [monotonicity #3083]: #3085 +#3103 := [monotonicity #3086 #3100]: #3102 +#3106 := [monotonicity #3103]: #3105 +#3109 := [monotonicity #3106]: #3108 +#3114 := [trans #3109 #3112]: #3113 +#3079 := (iff #2504 #3078) +#3076 := (iff #2501 #3075) +#3073 := (iff #2496 #3072) +#3070 := (iff #2493 #3067) +#3057 := (+ #2488 #2491) +#3058 := (+ #1707 #3057) +#3061 := (= #3058 0::int) +#3068 := (iff #3061 #3067) +#3069 := [rewrite]: #3068 +#3062 := (iff #2493 #3061) +#3059 := (= #2492 #3058) +#3060 := [rewrite]: #3059 +#3063 := [monotonicity #3060]: #3062 +#3071 := [trans #3063 #3069]: #3070 +#3055 := (iff #2495 #3054) +#3052 := (iff #2494 #3049) +#3041 := (+ #1707 #2488) +#3044 := (>= #3041 0::int) +#3050 := (iff #3044 #3049) +#3051 := [rewrite]: #3050 +#3045 := (iff #2494 #3044) +#3042 := (= #2489 #3041) +#3043 := [rewrite]: #3042 +#3046 := [monotonicity #3043]: #3045 +#3053 := [trans #3046 #3051]: #3052 +#3056 := [monotonicity #3053]: #3055 +#3074 := [monotonicity #3056 #3071]: #3073 +#3077 := [monotonicity #3074]: #3076 +#3080 := [quant-intro #3077]: #3079 +#3117 := [monotonicity #3080 #3114]: #3116 +#3039 := (iff #2480 #3036) +#3002 := (and #2461 #2996) +#3033 := (and #3002 #3030) +#3037 := (iff #3033 #3036) +#3038 := [rewrite]: #3037 +#3034 := (iff #2480 #3033) +#3031 := (iff #2476 #3030) +#3028 := (iff #2473 #3027) +#3025 := (iff #2457 #3024) +#3022 := (iff #2454 #3021) +#3019 := (= #2453 #3018) +#3020 := [rewrite]: #3019 +#3023 := [monotonicity #3020]: #3022 +#3026 := [monotonicity #3023]: #3025 +#3029 := [monotonicity #3026]: #3028 +#3032 := [quant-intro #3029]: #3031 +#3015 := (iff #2470 #3002) +#3007 := (not #3002) +#3010 := (not #3007) +#3013 := (iff #3010 #3002) +#3014 := [rewrite]: #3013 +#3011 := (iff #2470 #3010) +#3008 := (iff #2465 #3007) +#3005 := (iff #2464 #3002) +#2999 := (and #2996 #2461) +#3003 := (iff #2999 #3002) +#3004 := [rewrite]: #3003 +#3000 := (iff #2464 #2999) +#2997 := (iff #2463 #2996) +#2994 := (iff #2462 #2993) +#2995 := [rewrite]: #2994 +#2998 := [monotonicity #2995]: #2997 +#3001 := [monotonicity #2998]: #3000 +#3006 := [trans #3001 #3004]: #3005 +#3009 := [monotonicity #3006]: #3008 +#3012 := [monotonicity #3009]: #3011 +#3016 := [trans #3012 #3014]: #3015 +#3035 := [monotonicity #3016 #3032]: #3034 +#3040 := [trans #3035 #3038]: #3039 +#3120 := [monotonicity #3040 #3117]: #3119 +#2991 := (iff #2438 #933) +#2992 := [rewrite]: #2991 +#2989 := (iff #2435 #930) +#2990 := [rewrite]: #2989 +#2987 := (iff #2432 #234) +#2988 := [rewrite]: #2987 +#2985 := (iff #2429 #229) +#2986 := [rewrite]: #2985 +#3123 := [monotonicity #2986 #2988 #2990 #2992 #3120]: #3122 +#2983 := (iff #2425 #2980) +#2977 := (and #151 #588 #1473 #2974 #1631 #1649 #1657 #2417) +#2981 := (iff #2977 #2980) +#2982 := [rewrite]: #2981 +#2978 := (iff #2425 #2977) +#2975 := (iff #2388 #2974) +#2972 := (iff #2384 #2971) +#2969 := (iff #2380 #2968) +#2966 := (iff #2376 #2965) +#2963 := (iff #2372 #2962) +#2960 := (iff #2368 #2959) +#2957 := (iff #2364 #2956) +#2954 := (iff #2360 #2953) +#2951 := (iff #2356 #2950) +#2948 := (iff #2352 #2947) +#2945 := (iff #2348 #2944) +#2942 := (iff #2344 #2939) +#2902 := (and #2319 #2896) +#2933 := (and #2902 #2930) +#2936 := (and #1541 #2933) +#2940 := (iff #2936 #2939) +#2941 := [rewrite]: #2940 +#2937 := (iff #2344 #2936) +#2934 := (iff #2339 #2933) +#2931 := (iff #2335 #2930) +#2928 := (iff #2332 #2927) +#2925 := (iff #2315 #2924) +#2922 := (iff #2312 #2921) +#2919 := (= #2311 #2918) +#2920 := [rewrite]: #2919 +#2923 := [monotonicity #2920]: #2922 +#2926 := [monotonicity #2923]: #2925 +#2929 := [monotonicity #2926]: #2928 +#2932 := [quant-intro #2929]: #2931 +#2915 := (iff #2329 #2902) +#2907 := (not #2902) +#2910 := (not #2907) +#2913 := (iff #2910 #2902) +#2914 := [rewrite]: #2913 +#2911 := (iff #2329 #2910) +#2908 := (iff #2323 #2907) +#2905 := (iff #2322 #2902) +#2899 := (and #2896 #2319) +#2903 := (iff #2899 #2902) +#2904 := [rewrite]: #2903 +#2900 := (iff #2322 #2899) +#2897 := (iff #2321 #2896) +#2894 := (iff #2320 #2893) +#2895 := [rewrite]: #2894 +#2898 := [monotonicity #2895]: #2897 +#2901 := [monotonicity #2898]: #2900 +#2906 := [trans #2901 #2904]: #2905 +#2909 := [monotonicity #2906]: #2908 +#2912 := [monotonicity #2909]: #2911 +#2916 := [trans #2912 #2914]: #2915 +#2935 := [monotonicity #2916 #2932]: #2934 +#2938 := [monotonicity #2935]: #2937 +#2943 := [trans #2938 #2941]: #2942 +#2891 := (iff #2293 #2890) +#2888 := (iff #2292 #2887) +#2885 := (iff #2283 #2882) +#2872 := (+ #2278 #2281) +#2873 := (+ #2277 #2872) +#2876 := (>= #2873 0::int) +#2883 := (iff #2876 #2882) +#2884 := [rewrite]: #2883 +#2877 := (iff #2283 #2876) +#2874 := (= #2282 #2873) +#2875 := [rewrite]: #2874 +#2878 := [monotonicity #2875]: #2877 +#2886 := [trans #2878 #2884]: #2885 +#2870 := (iff #2291 #2869) +#2867 := (iff #2290 #2866) +#2868 := [rewrite]: #2867 +#2871 := [monotonicity #2868]: #2870 +#2889 := [monotonicity #2871 #2886]: #2888 +#2892 := [monotonicity #2889]: #2891 +#2946 := [monotonicity #2892 #2943]: #2945 +#2949 := [monotonicity #2946]: #2948 +#2864 := (iff #2262 #2863) +#2861 := (iff #2261 #2860) +#2858 := (iff #2253 #2855) +#2847 := (+ #2250 #2251) +#2850 := (>= #2847 0::int) +#2856 := (iff #2850 #2855) +#2857 := [rewrite]: #2856 +#2851 := (iff #2253 #2850) +#2848 := (= #2252 #2847) +#2849 := [rewrite]: #2848 +#2852 := [monotonicity #2849]: #2851 +#2859 := [trans #2852 #2857]: #2858 +#2845 := (iff #2260 #2844) +#2842 := (iff #2259 #2841) +#2843 := [rewrite]: #2842 +#2846 := [monotonicity #2843]: #2845 +#2862 := [monotonicity #2846 #2859]: #2861 +#2865 := [monotonicity #2862]: #2864 +#2952 := [monotonicity #2865 #2949]: #2951 +#2955 := [monotonicity #2952]: #2954 +#2958 := [monotonicity #2955]: #2957 +#2839 := (iff #2229 #173) +#2840 := [rewrite]: #2839 +#2961 := [monotonicity #2840 #2958]: #2960 +#2964 := [monotonicity #2961]: #2963 +#2967 := [monotonicity #2964]: #2966 +#2837 := (iff #2215 #2836) +#2834 := (iff #2214 #2831) +#2828 := (or #2825 #2210) +#2832 := (iff #2828 #2831) +#2833 := [rewrite]: #2832 +#2829 := (iff #2214 #2828) +#2826 := (iff #2213 #2825) +#2827 := [rewrite]: #2826 +#2830 := [monotonicity #2827]: #2829 +#2835 := [trans #2830 #2833]: #2834 +#2838 := [monotonicity #2835]: #2837 +#2970 := [monotonicity #2838 #2967]: #2969 +#2973 := [monotonicity #2970]: #2972 +#2823 := (iff #2195 #2822) +#2820 := (iff #2194 #2817) +#2809 := (+ #2191 #2192) +#2812 := (>= #2809 0::int) +#2818 := (iff #2812 #2817) +#2819 := [rewrite]: #2818 +#2813 := (iff #2194 #2812) +#2810 := (= #2193 #2809) +#2811 := [rewrite]: #2810 +#2814 := [monotonicity #2811]: #2813 +#2821 := [trans #2814 #2819]: #2820 +#2824 := [monotonicity #2821]: #2823 +#2976 := [monotonicity #2824 #2973]: #2975 +#2807 := (iff #2177 #151) +#2808 := [rewrite]: #2807 +#2979 := [monotonicity #2808 #2976]: #2978 +#2984 := [trans #2979 #2982]: #2983 +#3126 := [monotonicity #2984 #3123]: #3125 +#2805 := (iff #2171 #2804) +#2802 := (iff #2168 #2801) +#2799 := (iff #2163 #2798) +#2796 := (iff #2158 #2793) +#2783 := (+ #2153 #2156) +#2784 := (+ #1404 #2783) +#2787 := (= #2784 0::int) +#2794 := (iff #2787 #2793) +#2795 := [rewrite]: #2794 +#2788 := (iff #2158 #2787) +#2785 := (= #2157 #2784) +#2786 := [rewrite]: #2785 +#2789 := [monotonicity #2786]: #2788 +#2797 := [trans #2789 #2795]: #2796 +#2781 := (iff #2160 #2780) +#2778 := (iff #2159 #2775) +#2767 := (+ #1404 #2153) +#2770 := (>= #2767 0::int) +#2776 := (iff #2770 #2775) +#2777 := [rewrite]: #2776 +#2771 := (iff #2159 #2770) +#2768 := (= #2154 #2767) +#2769 := [rewrite]: #2768 +#2772 := [monotonicity #2769]: #2771 +#2779 := [trans #2772 #2777]: #2778 +#2782 := [monotonicity #2779]: #2781 +#2800 := [monotonicity #2782 #2797]: #2799 +#2803 := [monotonicity #2800]: #2802 +#2806 := [quant-intro #2803]: #2805 +#2765 := (iff #2144 #2764) +#2762 := (iff #2141 #2759) +#2756 := (or #2753 #1395) +#2760 := (iff #2756 #2759) +#2761 := [rewrite]: #2760 +#2757 := (iff #2141 #2756) +#2754 := (iff #2134 #2753) +#2751 := (iff #2129 #2748) +#2738 := (+ #2125 #2127) +#2739 := (+ #1386 #2738) +#2742 := (= #2739 0::int) +#2749 := (iff #2742 #2748) +#2750 := [rewrite]: #2749 +#2743 := (iff #2129 #2742) +#2740 := (= #2128 #2739) +#2741 := [rewrite]: #2740 +#2744 := [monotonicity #2741]: #2743 +#2752 := [trans #2744 #2750]: #2751 +#2736 := (iff #2132 #2735) +#2733 := (iff #2131 #2730) +#2722 := (+ #1386 #2127) +#2725 := (>= #2722 0::int) +#2731 := (iff #2725 #2730) +#2732 := [rewrite]: #2731 +#2726 := (iff #2131 #2725) +#2723 := (= #2130 #2722) +#2724 := [rewrite]: #2723 +#2727 := [monotonicity #2724]: #2726 +#2734 := [trans #2727 #2732]: #2733 +#2737 := [monotonicity #2734]: #2736 +#2755 := [monotonicity #2737 #2752]: #2754 +#2758 := [monotonicity #2755]: #2757 +#2763 := [trans #2758 #2761]: #2762 +#2766 := [quant-intro #2763]: #2765 +#2720 := (iff #2120 #109) +#2721 := [rewrite]: #2720 +#3129 := [monotonicity #2721 #2766 #2806 #3126]: #3128 +#3134 := [trans #3129 #3132]: #3133 +#2718 := (iff #2115 #2715) +#2697 := (and #2086 #2691) +#2712 := (and #2685 #2697) +#2716 := (iff #2712 #2715) +#2717 := [rewrite]: #2716 +#2713 := (iff #2115 #2712) +#2710 := (iff #2112 #2697) +#2702 := (not #2697) +#2705 := (not #2702) +#2708 := (iff #2705 #2697) +#2709 := [rewrite]: #2708 +#2706 := (iff #2112 #2705) +#2703 := (iff #2090 #2702) +#2700 := (iff #2089 #2697) +#2694 := (and #2691 #2086) +#2698 := (iff #2694 #2697) +#2699 := [rewrite]: #2698 +#2695 := (iff #2089 #2694) +#2692 := (iff #2088 #2691) +#2689 := (iff #2087 #2688) +#2690 := [rewrite]: #2689 +#2693 := [monotonicity #2690]: #2692 +#2696 := [monotonicity #2693]: #2695 +#2701 := [trans #2696 #2699]: #2700 +#2704 := [monotonicity #2701]: #2703 +#2707 := [monotonicity #2704]: #2706 +#2711 := [trans #2707 #2709]: #2710 +#2686 := (iff #2108 #2685) +#2683 := (iff #2105 #2682) +#2680 := (iff #2099 #2679) +#2681 := [rewrite]: #2680 +#2684 := [monotonicity #2681]: #2683 +#2687 := [quant-intro #2684]: #2686 +#2714 := [monotonicity #2687 #2711]: #2713 +#2719 := [trans #2714 #2717]: #2718 +#3137 := [monotonicity #2719 #3134]: #3136 +#3140 := [monotonicity #3137]: #3139 +#2677 := (iff #2069 #2676) +#2674 := (iff #2068 #2673) +#2671 := (iff #2061 #2670) +#2668 := (iff #2060 #2667) +#2669 := [rewrite]: #2668 +#2672 := [monotonicity #2669]: #2671 +#2665 := (iff #2067 #2664) +#2662 := (= #2066 #2661) +#2663 := [rewrite]: #2662 +#2666 := [monotonicity #2663]: #2665 +#2675 := [monotonicity #2666 #2672]: #2674 +#2678 := [monotonicity #2675]: #2677 +#3143 := [monotonicity #2678 #3140]: #3142 +#3146 := [monotonicity #3143]: #3145 +#2658 := (iff #2039 #2657) +#2655 := (iff #2038 #2654) +#2652 := (iff #2032 #2649) +#2641 := (+ #2029 #2030) +#2644 := (>= #2641 0::int) +#2650 := (iff #2644 #2649) +#2651 := [rewrite]: #2650 +#2645 := (iff #2032 #2644) +#2642 := (= #2031 #2641) +#2643 := [rewrite]: #2642 +#2646 := [monotonicity #2643]: #2645 +#2653 := [trans #2646 #2651]: #2652 +#2639 := (iff #2037 #2638) +#2636 := (iff #2036 #2635) +#2637 := [rewrite]: #2636 +#2640 := [monotonicity #2637]: #2639 +#2656 := [monotonicity #2640 #2653]: #2655 +#2659 := [monotonicity #2656]: #2658 +#3149 := [monotonicity #2659 #3146]: #3148 +#3152 := [monotonicity #3149]: #3151 +#3155 := [monotonicity #3152]: #3154 +#2633 := (iff #1969 #81) +#2634 := [rewrite]: #2633 +#3158 := [monotonicity #2634 #3155]: #3157 +#3161 := [monotonicity #3158]: #3160 +#1943 := (not #1907) +#2628 := (~ #1943 #2627) +#2624 := (not #1904) +#2625 := (~ #2624 #2623) +#2620 := (not #1901) +#2621 := (~ #2620 #2619) +#2616 := (not #1898) +#2617 := (~ #2616 #2615) +#2612 := (not #1895) +#2613 := (~ #2612 #2611) +#2608 := (not #1892) +#2609 := (~ #2608 #2607) +#2604 := (not #1889) +#2605 := (~ #2604 #2603) +#2600 := (not #1886) +#2601 := (~ #2600 #2599) +#2596 := (not #1883) +#2597 := (~ #2596 #2595) +#2592 := (not #1878) +#2593 := (~ #2592 #2591) +#2588 := (not #1849) +#2589 := (~ #2588 #1846) +#2586 := (~ #1846 #1846) +#2584 := (~ #1843 #1843) +#2585 := [refl]: #2584 +#2587 := [nnf-pos #2585]: #2586 +#2590 := [nnf-neg #2587]: #2589 +#2581 := (not #1840) +#2582 := (~ #2581 #1837) +#2579 := (~ #1837 #1837) +#2577 := (~ #1832 #1832) +#2578 := [refl]: #2577 +#2580 := [nnf-pos #2578]: #2579 +#2583 := [nnf-neg #2580]: #2582 +#2574 := (not #1824) +#2575 := (~ #2574 #1821) +#2572 := (~ #1821 #1821) +#2570 := (~ #1818 #1818) +#2571 := [refl]: #2570 +#2573 := [nnf-pos #2571]: #2572 +#2576 := [nnf-neg #2573]: #2575 +#2567 := (not #1801) +#2568 := (~ #2567 #2566) +#2563 := (not #1796) +#2564 := (~ #2563 #2562) +#2559 := (not #1778) +#2560 := (~ #2559 #2558) +#2555 := (not #1775) +#2556 := (~ #2555 #2554) +#2551 := (not #1772) +#2552 := (~ #2551 #2550) +#2547 := (not #1767) +#2548 := (~ #2547 #2546) +#2543 := (not #1759) +#2544 := (~ #2543 #1756) +#2541 := (~ #1756 #1756) +#2539 := (~ #1753 #1753) +#2540 := [refl]: #2539 +#2542 := [nnf-pos #2540]: #2541 +#2545 := [nnf-neg #2542]: #2544 +#2537 := (~ #2536 #2536) +#2538 := [refl]: #2537 +#2549 := [nnf-neg #2538 #2545]: #2548 +#2532 := (~ #1759 #2531) +#2533 := [sk]: #2532 +#2553 := [nnf-neg #2533 #2549]: #2552 +#2507 := (not #1741) +#2508 := (~ #2507 #2504) +#2505 := (~ #1738 #2504) +#2502 := (~ #1735 #2501) +#2497 := (~ #1732 #2496) +#2498 := [sk]: #2497 +#2485 := (~ #1716 #1716) +#2486 := [refl]: #2485 +#2503 := [monotonicity #2486 #2498]: #2502 +#2506 := [nnf-pos #2503]: #2505 +#2509 := [nnf-neg #2506]: #2508 +#2557 := [nnf-neg #2509 #2553]: #2556 +#2483 := (~ #1741 #2480) +#2458 := (exists (vars (?x65 T2)) #2457) +#2466 := (or #2465 #2458) +#2467 := (not #2466) +#2481 := (~ #2467 #2480) +#2477 := (not #2458) +#2478 := (~ #2477 #2476) +#2474 := (~ #2473 #2473) +#2475 := [refl]: #2474 +#2479 := [nnf-neg #2475]: #2478 +#2471 := (~ #2470 #2470) +#2472 := [refl]: #2471 +#2482 := [nnf-neg #2472 #2479]: #2481 +#2468 := (~ #1741 #2467) +#2469 := [sk]: #2468 +#2484 := [trans #2469 #2482]: #2483 +#2561 := [nnf-neg #2484 #2557]: #2560 +#2445 := (~ #1678 #2444) +#2442 := (~ #2441 #2441) +#2443 := [refl]: #2442 +#2446 := [nnf-neg #2443]: #2445 +#2439 := (~ #2438 #2438) +#2440 := [refl]: #2439 +#2436 := (~ #2435 #2435) +#2437 := [refl]: #2436 +#2433 := (~ #2432 #2432) +#2434 := [refl]: #2433 +#2430 := (~ #2429 #2429) +#2431 := [refl]: #2430 +#2565 := [nnf-neg #2431 #2434 #2437 #2440 #2446 #2561]: #2564 +#2426 := (not #1702) +#2427 := (~ #2426 #2425) +#2422 := (not #1678) +#2423 := (~ #2422 #2417) +#2418 := (~ #1675 #2417) +#2419 := [sk]: #2418 +#2424 := [nnf-neg #2419]: #2423 +#2406 := (~ #1657 #1657) +#2407 := [refl]: #2406 +#2403 := (not #1652) +#2404 := (~ #2403 #1649) +#2401 := (~ #1649 #1649) +#2399 := (~ #1644 #1644) +#2400 := [refl]: #2399 +#2402 := [nnf-pos #2400]: #2401 +#2405 := [nnf-neg #2402]: #2404 +#2396 := (not #1634) +#2397 := (~ #2396 #1631) +#2394 := (~ #1631 #1631) +#2392 := (~ #1628 #1628) +#2393 := [refl]: #2392 +#2395 := [nnf-pos #2393]: #2394 +#2398 := [nnf-neg #2395]: #2397 +#2389 := (not #1617) +#2390 := (~ #2389 #2388) +#2385 := (not #1614) +#2386 := (~ #2385 #2384) +#2381 := (not #1611) +#2382 := (~ #2381 #2380) +#2377 := (not #1608) +#2378 := (~ #2377 #2376) +#2373 := (not #1605) +#2374 := (~ #2373 #2372) +#2369 := (not #1602) +#2370 := (~ #2369 #2368) +#2365 := (not #1599) +#2366 := (~ #2365 #2364) +#2361 := (not #1596) +#2362 := (~ #2361 #2360) +#2357 := (not #1593) +#2358 := (~ #2357 #2356) +#2353 := (not #1590) +#2354 := (~ #2353 #2352) +#2349 := (not #1587) +#2350 := (~ #2349 #2348) +#2345 := (not #1584) +#2346 := (~ #2345 #2344) +#2326 := (not #1581) +#2342 := (~ #2326 #2339) +#2316 := (exists (vars (?x53 T2)) #2315) +#2324 := (or #2323 #2316) +#2325 := (not #2324) +#2340 := (~ #2325 #2339) +#2336 := (not #2316) +#2337 := (~ #2336 #2335) +#2333 := (~ #2332 #2332) +#2334 := [refl]: #2333 +#2338 := [nnf-neg #2334]: #2337 +#2330 := (~ #2329 #2329) +#2331 := [refl]: #2330 +#2341 := [nnf-neg #2331 #2338]: #2340 +#2327 := (~ #2326 #2325) +#2328 := [sk]: #2327 +#2343 := [trans #2328 #2341]: #2342 +#2302 := (not #1544) +#2303 := (~ #2302 #1541) +#2300 := (~ #1541 #1541) +#2298 := (~ #1538 #1538) +#2299 := [refl]: #2298 +#2301 := [nnf-pos #2299]: #2300 +#2304 := [nnf-neg #2301]: #2303 +#2347 := [nnf-neg #2304 #2343]: #2346 +#2294 := (~ #1544 #2293) +#2295 := [sk]: #2294 +#2351 := [nnf-neg #2295 #2347]: #2350 +#2271 := (not #1520) +#2272 := (~ #2271 #1517) +#2269 := (~ #1517 #1517) +#2267 := (~ #1512 #1512) +#2268 := [refl]: #2267 +#2270 := [nnf-pos #2268]: #2269 +#2273 := [nnf-neg #2270]: #2272 +#2355 := [nnf-neg #2273 #2351]: #2354 +#2263 := (~ #1520 #2262) +#2264 := [sk]: #2263 +#2359 := [nnf-neg #2264 #2355]: #2358 +#2244 := (not #1501) +#2245 := (~ #2244 #1498) +#2242 := (~ #1498 #1498) +#2240 := (~ #1495 #1495) +#2241 := [refl]: #2240 +#2243 := [nnf-pos #2241]: #2242 +#2246 := [nnf-neg #2243]: #2245 +#2363 := [nnf-neg #2246 #2359]: #2362 +#2236 := (~ #1501 #2235) +#2237 := [sk]: #2236 +#2367 := [nnf-neg #2237 #2363]: #2366 +#2230 := (~ #2229 #2229) +#2231 := [refl]: #2230 +#2371 := [nnf-neg #2231 #2367]: #2370 +#2227 := (~ #1492 #1492) +#2228 := [refl]: #2227 +#2375 := [nnf-neg #2228 #2371]: #2374 +#2224 := (not #824) +#2225 := (~ #2224 #642) +#2222 := (~ #642 #642) +#2220 := (~ #637 #637) +#2221 := [refl]: #2220 +#2223 := [nnf-pos #2221]: #2222 +#2226 := [nnf-neg #2223]: #2225 +#2379 := [nnf-neg #2226 #2375]: #2378 +#2216 := (~ #824 #2215) +#2217 := [sk]: #2216 +#2383 := [nnf-neg #2217 #2379]: #2382 +#2204 := (not #1487) +#2205 := (~ #2204 #1484) +#2202 := (~ #1484 #1484) +#2200 := (~ #1479 #1479) +#2201 := [refl]: #2200 +#2203 := [nnf-pos #2201]: #2202 +#2206 := [nnf-neg #2203]: #2205 +#2387 := [nnf-neg #2206 #2383]: #2386 +#2196 := (~ #1487 #2195) +#2197 := [sk]: #2196 +#2391 := [nnf-neg #2197 #2387]: #2390 +#2186 := (not #1476) +#2187 := (~ #2186 #1473) +#2184 := (~ #1473 #1473) +#2182 := (~ #1468 #1468) +#2183 := [refl]: #2182 +#2185 := [nnf-pos #2183]: #2184 +#2188 := [nnf-neg #2185]: #2187 +#2180 := (~ #588 #588) +#2181 := [refl]: #2180 +#2178 := (~ #2177 #2177) +#2179 := [refl]: #2178 +#2428 := [nnf-neg #2179 #2181 #2188 #2391 #2398 #2405 #2407 #2424]: #2427 +#2569 := [nnf-neg #2428 #2565]: #2568 +#2174 := (not #1446) +#2175 := (~ #2174 #2171) +#2172 := (~ #1443 #2171) +#2169 := (~ #1440 #2168) +#2164 := (~ #1437 #2163) +#2165 := [sk]: #2164 +#2150 := (~ #1413 #1413) +#2151 := [refl]: #2150 +#2170 := [monotonicity #2151 #2165]: #2169 +#2173 := [nnf-pos #2170]: #2172 +#2176 := [nnf-neg #2173]: #2175 +#2147 := (not #1857) +#2148 := (~ #2147 #2144) +#2145 := (~ #1401 #2144) +#2142 := (~ #1398 #2141) +#2139 := (~ #1395 #1395) +#2140 := [refl]: #2139 +#2135 := (~ #1383 #2134) +#2136 := [sk]: #2135 +#2143 := [monotonicity #2136 #2140]: #2142 +#2146 := [nnf-pos #2143]: #2145 +#2149 := [nnf-neg #2146]: #2148 +#2121 := (~ #2120 #2120) +#2122 := [refl]: #2121 +#2594 := [nnf-neg #2122 #2149 #2176 #2569 #2576 #2583 #2590]: #2593 +#2118 := (~ #1857 #2115) +#2100 := (exists (vars (?x33 T2)) #2099) +#2101 := (or #2100 #2090) +#2102 := (not #2101) +#2116 := (~ #2102 #2115) +#2113 := (~ #2112 #2112) +#2114 := [refl]: #2113 +#2109 := (not #2100) +#2110 := (~ #2109 #2108) +#2106 := (~ #2105 #2105) +#2107 := [refl]: #2106 +#2111 := [nnf-neg #2107]: #2110 +#2117 := [nnf-neg #2111 #2114]: #2116 +#2103 := (~ #1857 #2102) +#2104 := [sk]: #2103 +#2119 := [trans #2104 #2117]: #2118 +#2598 := [nnf-neg #2119 #2594]: #2597 +#2078 := (not #1363) +#2079 := (~ #2078 #1360) +#2076 := (~ #1360 #1360) +#2074 := (~ #1357 #1357) +#2075 := [refl]: #2074 +#2077 := [nnf-pos #2075]: #2076 +#2080 := [nnf-neg #2077]: #2079 +#2602 := [nnf-neg #2080 #2598]: #2601 +#2070 := (~ #1363 #2069) +#2071 := [sk]: #2070 +#2606 := [nnf-neg #2071 #2602]: #2605 +#2048 := (not #1337) +#2049 := (~ #2048 #1334) +#2046 := (~ #1334 #1334) +#2044 := (~ #1329 #1329) +#2045 := [refl]: #2044 +#2047 := [nnf-pos #2045]: #2046 +#2050 := [nnf-neg #2047]: #2049 +#2610 := [nnf-neg #2050 #2606]: #2609 +#2040 := (~ #1337 #2039) +#2041 := [sk]: #2040 +#2614 := [nnf-neg #2041 #2610]: #2613 +#2023 := (not #1317) +#2024 := (~ #2023 #1314) +#2021 := (~ #1314 #1314) +#2019 := (~ #1313 #1313) +#2020 := [refl]: #2019 +#2022 := [nnf-pos #2020]: #2021 +#2025 := [nnf-neg #2022]: #2024 +#2618 := [nnf-neg #2025 #2614]: #2617 +#2015 := (~ #1317 #2014) +#2016 := [sk]: #2015 +#2622 := [nnf-neg #2016 #2618]: #2621 +#1970 := (~ #1969 #1969) +#2010 := [refl]: #1970 +#2626 := [nnf-neg #2010 #2622]: #2625 +#2008 := (~ #1308 #1308) +#2009 := [refl]: #2008 +#2629 := [nnf-neg #2009 #2626]: #2628 +#1944 := [not-or-elim #1940]: #1943 +#2630 := [mp~ #1944 #2629]: #2627 +#2631 := [mp #2630 #3161]: #3159 +#3803 := [mp #2631 #3802]: #3800 +#4734 := [mp #3803 #4733]: #4731 +#7295 := [unit-resolution #4734 #5487]: #4728 +#4058 := (or #4725 #4719) +#4059 := [def-axiom]: #4058 +#7296 := [unit-resolution #4059 #7295]: #4719 +#373 := (<= uf_9 0::int) +#374 := (not #373) +#57 := (< 0::int uf_9) +#375 := (iff #57 #374) +#376 := [rewrite]: #375 +#369 := [asserted]: #57 +#377 := [mp #369 #376]: #374 +#5901 := (* -1::int #2012) +#5891 := (+ uf_9 #5901) +#5892 := (<= #5891 0::int) +#5472 := (= uf_9 #2012) +#5745 := (= uf_11 ?x27!0) +#5918 := (not #5745) +#5916 := (= #2012 0::int) +#5836 := (not #5916) +#5835 := [hypothesis]: #2014 +#5837 := (or #5836 #2013) +#5896 := [th-lemma]: #5837 +#5922 := [unit-resolution #5896 #5835]: #5836 +#5981 := (or #5347 #5918 #5916) +#5477 := (= ?x27!0 uf_11) +#5917 := (not #5477) +#5890 := (or #5917 #5916) +#5982 := (or #5347 #5890) +#5987 := (iff #5982 #5981) +#5915 := (or #5918 #5916) +#5984 := (or #5347 #5915) +#5985 := (iff #5984 #5981) +#5986 := [rewrite]: #5985 +#6000 := (iff #5982 #5984) +#5921 := (iff #5890 #5915) +#5919 := (iff #5917 #5918) +#5743 := (iff #5477 #5745) +#5746 := [rewrite]: #5743 +#5920 := [monotonicity #5746]: #5919 +#5980 := [monotonicity #5920]: #5921 +#5979 := [monotonicity #5980]: #6000 +#5988 := [trans #5979 #5986]: #5987 +#5983 := [quant-inst]: #5982 +#6010 := [mp #5983 #5988]: #5981 +#5923 := [unit-resolution #6010 #4740 #5922]: #5918 +#5748 := (or #5472 #5745) +#4356 := (forall (vars (?x25 T2)) (:pat #4355) #443) +#4359 := (iff #448 #4356) +#4357 := (iff #443 #443) +#4358 := [refl]: #4357 +#4360 := [quant-intro #4358]: #4359 +#1967 := (~ #448 #448) +#2005 := (~ #443 #443) +#2006 := [refl]: #2005 +#1968 := [nnf-pos #2006]: #1967 +#1942 := [not-or-elim #1940]: #448 +#2007 := [mp~ #1942 #1968]: #448 +#4361 := [mp #2007 #4360]: #4356 +#5821 := (not #4356) +#5827 := (or #5821 #5472 #5745) +#5737 := (or #5477 #5472) +#5828 := (or #5821 #5737) +#5894 := (iff #5828 #5827) +#5830 := (or #5821 #5748) +#5846 := (iff #5830 #5827) +#5847 := [rewrite]: #5846 +#5826 := (iff #5828 #5830) +#5756 := (iff #5737 #5748) +#5736 := (or #5745 #5472) +#5752 := (iff #5736 #5748) +#5753 := [rewrite]: #5752 +#5747 := (iff #5737 #5736) +#5744 := [monotonicity #5746]: #5747 +#5834 := [trans #5744 #5753]: #5756 +#5831 := [monotonicity #5834]: #5826 +#5895 := [trans #5831 #5847]: #5894 +#5829 := [quant-inst]: #5828 +#5900 := [mp #5829 #5895]: #5827 +#5924 := [unit-resolution #5900 #4361]: #5748 +#5989 := [unit-resolution #5924 #5923]: #5472 +#6012 := (not #5472) +#6013 := (or #6012 #5892) +#6014 := [th-lemma]: #6013 +#6042 := [unit-resolution #6014 #5989]: #5892 +#6011 := (<= #2012 0::int) +#6043 := (or #6011 #2013) +#6044 := [th-lemma]: #6043 +#6045 := [unit-resolution #6044 #5835]: #6011 +#6046 := [th-lemma #6045 #6042 #377]: false +#6041 := [lemma #6046]: #2013 +#4053 := (or #4722 #2014 #4716) +#4054 := [def-axiom]: #4053 +#7297 := [unit-resolution #4054 #6041 #7296]: #4716 +#4077 := (or #4713 #4707) +#4078 := [def-axiom]: #4077 +#7298 := [unit-resolution #4078 #7297]: #4707 +#4071 := (or #4710 #2421 #4704) +#4073 := [def-axiom]: #4071 +#7299 := [unit-resolution #4073 #7298 #7294]: #4704 +#4098 := (or #4701 #4695) +#4099 := [def-axiom]: #4098 +#7300 := [unit-resolution #4099 #7299]: #4695 +#6817 := [hypothesis]: #2059 +#6051 := (or #5709 #3184) +#6081 := [quant-inst]: #6051 +#6818 := [unit-resolution #6081 #4354 #6817]: false +#6839 := [lemma #6818]: #3184 +#3960 := (or #3199 #2059) +#3964 := [def-axiom]: #3960 +#7301 := [unit-resolution #3964 #6839]: #3199 +#4094 := (or #4698 #3204 #4692) +#4095 := [def-axiom]: #4094 +#7302 := [unit-resolution #4095 #7301 #7300]: #4692 +#4108 := (or #4689 #4683) +#4129 := [def-axiom]: #4108 +#7303 := [unit-resolution #4129 #7302]: #4683 +#6633 := (= uf_9 #2082) +#6706 := (not #6633) +#6684 := [hypothesis]: #4400 +#4274 := (or #4397 #2086) +#3948 := [def-axiom]: #4274 +#6685 := [unit-resolution #3948 #6684]: #2086 +#6798 := (or #6706 #2085) +#6840 := [th-lemma]: #6798 +#6835 := [unit-resolution #6840 #6685]: #6706 +#3949 := (or #4397 #2691) +#4281 := [def-axiom]: #3949 +#6841 := [unit-resolution #4281 #6684]: #2691 +#6695 := (or #5821 #2688 #6633) +#6680 := (or #2087 #6633) +#6696 := (or #5821 #6680) +#6732 := (iff #6696 #6695) +#6548 := (or #2688 #6633) +#6694 := (or #5821 #6548) +#6699 := (iff #6694 #6695) +#6705 := [rewrite]: #6699 +#6697 := (iff #6696 #6694) +#6608 := (iff #6680 #6548) +#6665 := [monotonicity #2690]: #6608 +#6698 := [monotonicity #6665]: #6697 +#6728 := [trans #6698 #6705]: #6732 +#6547 := [quant-inst]: #6696 +#6734 := [mp #6547 #6728]: #6695 +#6842 := [unit-resolution #6734 #4361 #6841 #6835]: false +#6843 := [lemma #6842]: #4397 +#4116 := (or #4686 #4400 #4680) +#4117 := [def-axiom]: #4116 +#7304 := [unit-resolution #4117 #6843 #7303]: #4680 +#4149 := (or #4677 #4404) +#4145 := [def-axiom]: #4149 +#7305 := [unit-resolution #4145 #7304]: #4404 +#24124 := (or #4409 #20405) +#25438 := [quant-inst]: #24124 +#29563 := [unit-resolution #25438 #7305]: #20405 +#15997 := (* -1::int #15996) +#15993 := (uf_1 #15992 ?x52!15) +#15994 := (uf_10 #15993) +#15995 := (* -1::int #15994) +#16012 := (+ #15995 #15997) +#15362 := (uf_4 uf_14 ?x52!15) +#16013 := (+ #15362 #16012) +#22006 := (>= #16013 0::int) +#16016 := (= #16013 0::int) +#16019 := (not #16016) +#16004 := (uf_6 uf_15 #15992) +#16005 := (= uf_8 #16004) +#16006 := (not #16005) +#16002 := (+ #15362 #15997) +#16003 := (<= #16002 0::int) +#16025 := (or #16003 #16006 #16019) +#16030 := (not #16025) +#15397 := (* -1::int #15362) +#16009 := (+ uf_9 #15397) +#16010 := (<= #16009 0::int) +#34324 := (not #16010) +#15398 := (+ #2306 #15397) +#14218 := (>= #15398 0::int) +#15367 := (= #2306 #15362) +decl uf_3 :: (-> T1 T2) +#20604 := (uf_3 #15993) +#32578 := (uf_6 uf_15 #20604) +#32576 := (= uf_8 #32578) +#5319 := (uf_6 #150 uf_16) +#5314 := (= uf_8 #5319) +decl uf_2 :: (-> T1 T2) +#6008 := (uf_1 uf_16 uf_11) +#7128 := (uf_2 #6008) +#32602 := (= #7128 #20604) +#32591 := (ite #32602 #5314 #32576) +#7203 := (uf_7 uf_15 #7128 #5319) +#32554 := (uf_6 #7203 #20604) +#32538 := (= uf_8 #32554) +#32564 := (iff #32538 #32591) +#30 := (:var 1 T5) +#20 := (:var 2 T2) +#29 := (:var 3 T4) +#31 := (uf_7 #29 #20 #30) +#32 := (uf_6 #31 #11) +#4314 := (pattern #32) +#36 := (uf_6 #29 #11) +#335 := (= uf_8 #36) +#35 := (= #30 uf_8) +#24 := (= #11 #20) +#338 := (ite #24 #35 #335) +#34 := (= #32 uf_8) +#341 := (iff #34 #338) +#4315 := (forall (vars (?x10 T4) (?x11 T2) (?x12 T5) (?x13 T2)) (:pat #4314) #341) +#344 := (forall (vars (?x10 T4) (?x11 T2) (?x12 T5) (?x13 T2)) #341) +#4318 := (iff #344 #4315) +#4316 := (iff #341 #341) +#4317 := [refl]: #4316 +#4319 := [quant-intro #4317]: #4318 +#1953 := (~ #344 #344) +#1989 := (~ #341 #341) +#1990 := [refl]: #1989 +#1954 := [nnf-pos #1990]: #1953 +#37 := (= #36 uf_8) +#38 := (ite #24 #35 #37) +#39 := (iff #34 #38) +#40 := (forall (vars (?x10 T4) (?x11 T2) (?x12 T5) (?x13 T2)) #39) +#345 := (iff #40 #344) +#342 := (iff #39 #341) +#339 := (iff #38 #338) +#336 := (iff #37 #335) +#337 := [rewrite]: #336 +#340 := [monotonicity #337]: #339 +#343 := [monotonicity #340]: #342 +#346 := [quant-intro #343]: #345 +#333 := [asserted]: #40 +#349 := [mp #333 #346]: #344 +#1991 := [mp~ #349 #1954]: #344 +#4320 := [mp #1991 #4319]: #4315 +#7026 := (not #4315) +#32561 := (or #7026 #32564) +#6089 := (= #5319 uf_8) +#32580 := (= #20604 #7128) +#32553 := (ite #32580 #6089 #32576) +#32556 := (= #32554 uf_8) +#32579 := (iff #32556 #32553) +#32603 := (or #7026 #32579) +#32548 := (iff #32603 #32561) +#32606 := (iff #32561 #32561) +#32631 := [rewrite]: #32606 +#32629 := (iff #32579 #32564) +#32535 := (iff #32553 #32591) +#6101 := (iff #6089 #5314) +#6102 := [rewrite]: #6101 +#32601 := (iff #32580 #32602) +#32563 := [rewrite]: #32601 +#32560 := [monotonicity #32563 #6102]: #32535 +#32537 := (iff #32556 #32538) +#32557 := [rewrite]: #32537 +#32565 := [monotonicity #32557 #32560]: #32629 +#32587 := [monotonicity #32565]: #32548 +#32604 := [trans #32587 #32631]: #32548 +#32558 := [quant-inst]: #32603 +#32623 := [mp #32558 #32604]: #32561 +#32954 := [unit-resolution #32623 #4320]: #32564 +#30313 := (not #32538) +#8488 := (uf_6 uf_17 ?x52!15) +#9319 := (= uf_8 #8488) +#9846 := (not #9319) +#32968 := (iff #9846 #30313) +#32955 := (iff #9319 #32538) +#32985 := (iff #32538 #9319) +#32980 := (= #32554 #8488) +#32983 := (= #20604 ?x52!15) +#20605 := (= ?x52!15 #20604) +#12 := (uf_1 #10 #11) +#4294 := (pattern #12) +#13 := (uf_3 #12) +#317 := (= #11 #13) +#4295 := (forall (vars (?x2 T2) (?x3 T2)) (:pat #4294) #317) +#321 := (forall (vars (?x2 T2) (?x3 T2)) #317) +#4298 := (iff #321 #4295) +#4296 := (iff #317 #317) +#4297 := [refl]: #4296 +#4299 := [quant-intro #4297]: #4298 +#1948 := (~ #321 #321) +#1980 := (~ #317 #317) +#1981 := [refl]: #1980 +#1946 := [nnf-pos #1981]: #1948 +#14 := (= #13 #11) +#15 := (forall (vars (?x2 T2) (?x3 T2)) #14) +#322 := (iff #15 #321) +#319 := (iff #14 #317) +#320 := [rewrite]: #319 +#323 := [quant-intro #320]: #322 +#316 := [asserted]: #15 +#326 := [mp #316 #323]: #321 +#1982 := [mp~ #326 #1946]: #321 +#4300 := [mp #1982 #4299]: #4295 +#5378 := (not #4295) +#27981 := (or #5378 #20605) +#27945 := [quant-inst]: #27981 +#32953 := [unit-resolution #27945 #4300]: #20605 +#32984 := [symm #32953]: #32983 +#8596 := (= #7203 uf_17) +#8594 := (= #150 uf_17) +#4147 := (or #4677 #109) +#4148 := [def-axiom]: #4147 +#7307 := [unit-resolution #4148 #7304]: #109 +#4150 := (or #4677 #4412) +#4130 := [def-axiom]: #4150 +#8027 := [unit-resolution #4130 #7304]: #4412 +#4137 := (or #4677 #4437) +#4132 := [def-axiom]: #4137 +#8030 := [unit-resolution #4132 #7304]: #4437 +#5284 := (or #4665 #4442 #4417 #1854) +#4776 := (uf_4 uf_14 ?x64!17) +#4777 := (* -1::int #4776) +#4778 := (+ uf_9 #4777) +#4779 := (<= #4778 0::int) +#4845 := (?x40!7 ?x64!17) +#4941 := (uf_6 uf_15 #4845) +#4942 := (= uf_8 #4941) +#4943 := (not #4942) +#4848 := (uf_4 uf_14 #4845) +#4849 := (* -1::int #4848) +#4939 := (+ #4776 #4849) +#4940 := (<= #4939 0::int) +#4846 := (uf_1 #4845 ?x64!17) +#4847 := (uf_10 #4846) +#4842 := (* -1::int #4847) +#4926 := (+ #4842 #4849) +#4927 := (+ #4776 #4926) +#4930 := (= #4927 0::int) +#4932 := (not #4930) +#5006 := (or #4932 #4940 #4943) +#5242 := [hypothesis]: #4668 +#4159 := (or #4665 #933) +#4154 := [def-axiom]: #4159 +#5243 := [unit-resolution #4154 #5242]: #933 +#4134 := (or #4665 #4659) +#4135 := [def-axiom]: #4134 +#5244 := [unit-resolution #4135 #5242]: #4659 +#5245 := [hypothesis]: #109 +#5247 := (= #250 #108) +#5246 := [symm #5243]: #231 +#5248 := [monotonicity #5246]: #5247 +#5249 := [trans #5248 #5245]: #251 +#4193 := (or #4641 #2536) +#4198 := [def-axiom]: #4193 +#5250 := [unit-resolution #4198 #5249]: #4641 +#5087 := [hypothesis]: #4412 +#4160 := (or #4665 #4601) +#4133 := [def-axiom]: #4160 +#5230 := [unit-resolution #4133 #5242]: #4601 +#5100 := (or #3693 #4417 #4606 #1053) +#4862 := (uf_4 uf_14 ?x67!19) +#3931 := (uf_4 uf_14 ?x66!20) +#3932 := (* -1::int #3931) +#4955 := (+ #3932 #4862) +#4956 := (+ #2517 #4955) +#4959 := (>= #4956 0::int) +#4866 := (uf_6 uf_15 ?x67!19) +#4867 := (= uf_8 #4866) +#4863 := (* -1::int #4862) +#4864 := (+ uf_9 #4863) +#4865 := (<= #4864 0::int) +#5068 := (not #4865) +#5072 := [hypothesis]: #3698 +#4189 := (or #3693 #2523) +#4186 := [def-axiom]: #4189 +#5073 := [unit-resolution #4186 #5072]: #2523 +#5061 := (+ #2514 #4863) +#5063 := (>= #5061 0::int) +#5060 := (= #2514 #4862) +#5075 := (= #4862 #2514) +#5074 := [hypothesis]: #933 +#5076 := [monotonicity #5074]: #5075 +#5077 := [symm #5076]: #5060 +#5078 := (not #5060) +#5079 := (or #5078 #5063) +#5080 := [th-lemma]: #5079 +#5081 := [unit-resolution #5080 #5077]: #5063 +#5069 := (not #5063) +#5070 := (or #5068 #5069 #2522) +#5064 := [hypothesis]: #2523 +#5065 := [hypothesis]: #4865 +#5066 := [hypothesis]: #5063 +#5067 := [th-lemma #5066 #5065 #5064]: false +#5071 := [lemma #5067]: #5070 +#5082 := [unit-resolution #5071 #5081 #5073]: #5068 +#4869 := (or #4865 #4867) +#5083 := [hypothesis]: #4601 +#4872 := (or #4606 #4865 #4867) +#4868 := (or #4867 #4865) +#4873 := (or #4606 #4868) +#4880 := (iff #4873 #4872) +#4875 := (or #4606 #4869) +#4878 := (iff #4875 #4872) +#4879 := [rewrite]: #4878 +#4876 := (iff #4873 #4875) +#4870 := (iff #4868 #4869) +#4871 := [rewrite]: #4870 +#4877 := [monotonicity #4871]: #4876 +#4881 := [trans #4877 #4879]: #4880 +#4874 := [quant-inst]: #4873 +#4882 := [mp #4874 #4881]: #4872 +#5084 := [unit-resolution #4882 #5083]: #4869 +#5085 := [unit-resolution #5084 #5082]: #4867 +#4953 := (not #4867) +#5088 := (or #4953 #4959) +#4190 := (or #3693 #2527) +#4170 := [def-axiom]: #4190 +#5086 := [unit-resolution #4170 #5072]: #2527 +#4970 := (or #4417 #2526 #4953 #4959) +#4948 := (+ #4862 #3932) +#4949 := (+ #2517 #4948) +#4952 := (>= #4949 0::int) +#4954 := (or #4953 #2526 #4952) +#4971 := (or #4417 #4954) +#4978 := (iff #4971 #4970) +#4965 := (or #2526 #4953 #4959) +#4973 := (or #4417 #4965) +#4976 := (iff #4973 #4970) +#4977 := [rewrite]: #4976 +#4974 := (iff #4971 #4973) +#4968 := (iff #4954 #4965) +#4962 := (or #4953 #2526 #4959) +#4966 := (iff #4962 #4965) +#4967 := [rewrite]: #4966 +#4963 := (iff #4954 #4962) +#4960 := (iff #4952 #4959) +#4957 := (= #4949 #4956) +#4958 := [rewrite]: #4957 +#4961 := [monotonicity #4958]: #4960 +#4964 := [monotonicity #4961]: #4963 +#4969 := [trans #4964 #4967]: #4968 +#4975 := [monotonicity #4969]: #4974 +#4979 := [trans #4975 #4977]: #4978 +#4972 := [quant-inst]: #4971 +#4980 := [mp #4972 #4979]: #4970 +#5089 := [unit-resolution #4980 #5087 #5086]: #5088 +#5090 := [unit-resolution #5089 #5085]: #4959 +#4171 := (not #3096) +#4173 := (or #3693 #4171) +#4174 := [def-axiom]: #4173 +#5091 := [unit-resolution #4174 #5072]: #4171 +#5054 := (+ #2512 #3932) +#5058 := (<= #5054 0::int) +#5053 := (= #2512 #3931) +#5092 := (= #3931 #2512) +#5093 := [monotonicity #5074]: #5092 +#5094 := [symm #5093]: #5053 +#5095 := (not #5053) +#5096 := (or #5095 #5058) +#5097 := [th-lemma]: #5096 +#5098 := [unit-resolution #5097 #5094]: #5058 +#5099 := [th-lemma #5098 #5091 #5081 #5090]: false +#5101 := [lemma #5099]: #5100 +#5231 := [unit-resolution #5101 #5230 #5087 #5243]: #3693 +#4181 := (or #4650 #4644 #3698) +#4182 := [def-axiom]: #4181 +#5232 := [unit-resolution #4182 #5231 #5250]: #4650 +#4161 := (or #4653 #4647) +#4162 := [def-axiom]: #4161 +#5233 := [unit-resolution #4162 #5232]: #4653 +#4169 := (or #4662 #4622 #4656) +#4155 := [def-axiom]: #4169 +#5234 := [unit-resolution #4155 #5233 #5244]: #4622 +#4194 := (or #4619 #4611) +#4195 := [def-axiom]: #4194 +#5229 := [unit-resolution #4195 #5234]: #4611 +#5713 := (or #5006 #4616 #1053) +#5123 := (uf_4 uf_22 #4845) +#5136 := (* -1::int #5123) +#5137 := (+ #2448 #5136) +#5138 := (<= #5137 0::int) +#5150 := (+ #4842 #5136) +#5151 := (+ #2448 #5150) +#5152 := (= #5151 0::int) +#5382 := (+ #4848 #5136) +#5387 := (>= #5382 0::int) +#5381 := (= #4848 #5123) +#5643 := (= #5123 #4848) +#5642 := [symm #5074]: #231 +#5644 := [monotonicity #5642]: #5643 +#5645 := [symm #5644]: #5381 +#5646 := (not #5381) +#5647 := (or #5646 #5387) +#5648 := [th-lemma]: #5647 +#5649 := [unit-resolution #5648 #5645]: #5387 +#5119 := (+ #2448 #4777) +#5121 := (>= #5119 0::int) +#5118 := (= #2448 #4776) +#5629 := (= #4776 #2448) +#5630 := [monotonicity #5074]: #5629 +#5631 := [symm #5630]: #5118 +#5632 := (not #5118) +#5633 := (or #5632 #5121) +#5628 := [th-lemma]: #5633 +#5634 := [unit-resolution #5628 #5631]: #5121 +#5040 := (>= #4927 0::int) +#5005 := (not #5006) +#5635 := [hypothesis]: #5005 +#5042 := (or #5006 #4930) +#5043 := [def-axiom]: #5042 +#5636 := [unit-resolution #5043 #5635]: #4930 +#5637 := (or #4932 #5040) +#5638 := [th-lemma]: #5637 +#5653 := [unit-resolution #5638 #5636]: #5040 +#5386 := (<= #5382 0::int) +#5654 := (or #5646 #5386) +#5675 := [th-lemma]: #5654 +#5676 := [unit-resolution #5675 #5645]: #5386 +#5120 := (<= #5119 0::int) +#5677 := (or #5632 #5120) +#5678 := [th-lemma]: #5677 +#5679 := [unit-resolution #5678 #5631]: #5120 +#5034 := (<= #4927 0::int) +#5674 := (or #4932 #5034) +#5680 := [th-lemma]: #5674 +#5681 := [unit-resolution #5680 #5636]: #5034 +#5562 := (not #5387) +#5567 := (not #5121) +#5566 := (not #5040) +#5779 := (not #5386) +#5778 := (not #5120) +#5777 := (not #5034) +#5572 := (or #5152 #5777 #5778 #5779 #5566 #5567 #5562) +#5774 := [hypothesis]: #5386 +#5775 := [hypothesis]: #5120 +#5776 := [hypothesis]: #5034 +#5157 := (not #5152) +#5772 := [hypothesis]: #5157 +#5175 := (>= #5151 0::int) +#5563 := [hypothesis]: #5387 +#5564 := [hypothesis]: #5121 +#5565 := [hypothesis]: #5040 +#5568 := (or #5175 #5566 #5567 #5562) +#5569 := [th-lemma]: #5568 +#5570 := [unit-resolution #5569 #5565 #5564 #5563]: #5175 +#5784 := (not #5175) +#5788 := (or #5784 #5152 #5777 #5778 #5779) +#5773 := [hypothesis]: #5175 +#5174 := (<= #5151 0::int) +#5780 := (or #5174 #5777 #5778 #5779) +#5781 := [th-lemma]: #5780 +#5782 := [unit-resolution #5781 #5776 #5775 #5774]: #5174 +#5783 := (not #5174) +#5785 := (or #5152 #5783 #5784) +#5786 := [th-lemma]: #5785 +#5787 := [unit-resolution #5786 #5782 #5773 #5772]: false +#5789 := [lemma #5787]: #5788 +#5571 := [unit-resolution #5789 #5570 #5772 #5776 #5775 #5774]: false +#5641 := [lemma #5571]: #5572 +#5682 := [unit-resolution #5641 #5681 #5679 #5676 #5653 #5634 #5649]: #5152 +#5160 := (or #5138 #5157) +#5683 := [hypothesis]: #4611 +#5163 := (or #4616 #5138 #5157) +#5122 := (+ #2449 #4847) +#5124 := (+ #5123 #5122) +#5125 := (= #5124 0::int) +#5126 := (not #5125) +#5127 := (+ #5123 #2449) +#5128 := (>= #5127 0::int) +#5129 := (or #5128 #5126) +#5164 := (or #4616 #5129) +#5171 := (iff #5164 #5163) +#5166 := (or #4616 #5160) +#5169 := (iff #5166 #5163) +#5170 := [rewrite]: #5169 +#5167 := (iff #5164 #5166) +#5161 := (iff #5129 #5160) +#5158 := (iff #5126 #5157) +#5155 := (iff #5125 #5152) +#5143 := (+ #4847 #5123) +#5144 := (+ #2449 #5143) +#5147 := (= #5144 0::int) +#5153 := (iff #5147 #5152) +#5154 := [rewrite]: #5153 +#5148 := (iff #5125 #5147) +#5145 := (= #5124 #5144) +#5146 := [rewrite]: #5145 +#5149 := [monotonicity #5146]: #5148 +#5156 := [trans #5149 #5154]: #5155 +#5159 := [monotonicity #5156]: #5158 +#5141 := (iff #5128 #5138) +#5130 := (+ #2449 #5123) +#5133 := (>= #5130 0::int) +#5139 := (iff #5133 #5138) +#5140 := [rewrite]: #5139 +#5134 := (iff #5128 #5133) +#5131 := (= #5127 #5130) +#5132 := [rewrite]: #5131 +#5135 := [monotonicity #5132]: #5134 +#5142 := [trans #5135 #5140]: #5141 +#5162 := [monotonicity #5142 #5159]: #5161 +#5168 := [monotonicity #5162]: #5167 +#5172 := [trans #5168 #5170]: #5171 +#5165 := [quant-inst]: #5164 +#5173 := [mp #5165 #5172]: #5163 +#5684 := [unit-resolution #5173 #5683]: #5160 +#5710 := [unit-resolution #5684 #5682]: #5138 +#5041 := (not #4940) +#5044 := (or #5006 #5041) +#5045 := [def-axiom]: #5044 +#5711 := [unit-resolution #5045 #5635]: #5041 +#5712 := [th-lemma #5634 #5711 #5649 #5710]: false +#5714 := [lemma #5712]: #5713 +#5235 := [unit-resolution #5714 #5229 #5243]: #5006 +#5238 := (or #4779 #5005) +#4191 := (or #4619 #2996) +#4192 := [def-axiom]: #4191 +#5236 := [unit-resolution #4192 #5234]: #2996 +#5237 := [hypothesis]: #4437 +#5023 := (or #4442 #2993 #4779 #5005) +#4850 := (+ #4849 #4842) +#4851 := (+ #4776 #4850) +#4852 := (= #4851 0::int) +#4938 := (not #4852) +#4944 := (or #4943 #4940 #4938) +#4945 := (not #4944) +#4946 := (or #2462 #4779 #4945) +#5025 := (or #4442 #4946) +#5031 := (iff #5025 #5023) +#5013 := (or #2993 #4779 #5005) +#5027 := (or #4442 #5013) +#5024 := (iff #5027 #5023) +#5030 := [rewrite]: #5024 +#5028 := (iff #5025 #5027) +#5014 := (iff #4946 #5013) +#5011 := (iff #4945 #5005) +#5009 := (iff #4944 #5006) +#4935 := (or #4943 #4940 #4932) +#5007 := (iff #4935 #5006) +#5008 := [rewrite]: #5007 +#4950 := (iff #4944 #4935) +#4933 := (iff #4938 #4932) +#4925 := (iff #4852 #4930) +#4928 := (= #4851 #4927) +#4929 := [rewrite]: #4928 +#4931 := [monotonicity #4929]: #4925 +#4934 := [monotonicity #4931]: #4933 +#4951 := [monotonicity #4934]: #4950 +#5010 := [trans #4951 #5008]: #5009 +#5012 := [monotonicity #5010]: #5011 +#5015 := [monotonicity #2995 #5012]: #5014 +#5029 := [monotonicity #5015]: #5028 +#5032 := [trans #5029 #5030]: #5031 +#5026 := [quant-inst]: #5025 +#5033 := [mp #5026 #5032]: #5023 +#5239 := [unit-resolution #5033 #5237 #5236]: #5238 +#5254 := [unit-resolution #5239 #5235]: #4779 +#4200 := (or #4619 #2461) +#4207 := [def-axiom]: #4200 +#5255 := [unit-resolution #4207 #5234]: #2461 +#5280 := [monotonicity #5243]: #5629 +#5281 := [symm #5280]: #5118 +#5282 := [unit-resolution #5628 #5281]: #5121 +#5283 := [th-lemma #5282 #5255 #5254]: false +#5279 := [lemma #5283]: #5284 +#8031 := [unit-resolution #5279 #8030 #8027 #7307]: #4665 +#4138 := (or #4677 #4671) +#4106 := [def-axiom]: #4138 +#7357 := [unit-resolution #4106 #7304]: #4671 +#4143 := (or #4674 #4598 #4668) +#4144 := [def-axiom]: #4143 +#7389 := [unit-resolution #4144 #7357 #8031]: #4598 +#4125 := (or #4595 #151) +#4126 := [def-axiom]: #4125 +#8578 := [unit-resolution #4126 #7389]: #151 +#8595 := [symm #8578]: #8594 +#8592 := (= #7203 #150) +#48 := (:var 0 T5) +#47 := (:var 2 T4) +#49 := (uf_7 #47 #10 #48) +#4329 := (pattern #49) +#360 := (= uf_8 #48) +#50 := (uf_6 #49 #10) +#356 := (= uf_8 #50) +#363 := (iff #356 #360) +#4330 := (forall (vars (?x17 T4) (?x18 T2) (?x19 T5)) (:pat #4329) #363) +#366 := (forall (vars (?x17 T4) (?x18 T2) (?x19 T5)) #363) +#4333 := (iff #366 #4330) +#4331 := (iff #363 #363) +#4332 := [refl]: #4331 +#4334 := [quant-intro #4332]: #4333 +#1957 := (~ #366 #366) +#1995 := (~ #363 #363) +#1996 := [refl]: #1995 +#1958 := [nnf-pos #1996]: #1957 +#52 := (= #48 uf_8) +#51 := (= #50 uf_8) +#53 := (iff #51 #52) +#54 := (forall (vars (?x17 T4) (?x18 T2) (?x19 T5)) #53) +#367 := (iff #54 #366) +#364 := (iff #53 #363) +#361 := (iff #52 #360) +#362 := [rewrite]: #361 +#358 := (iff #51 #356) +#359 := [rewrite]: #358 +#365 := [monotonicity #359 #362]: #364 +#368 := [quant-intro #365]: #367 +#355 := [asserted]: #54 +#371 := [mp #355 #368]: #366 +#1997 := [mp~ #371 #1958]: #366 +#4335 := [mp #1997 #4334]: #4330 +#7014 := (not #4330) +#7015 := (or #7014 #5314) +#5318 := (= uf_8 uf_8) +#5320 := (iff #5314 #5318) +#7018 := (or #7014 #5320) +#7020 := (iff #7018 #7015) +#7022 := (iff #7015 #7015) +#7023 := [rewrite]: #7022 +#5344 := (iff #5320 #5314) +#5323 := (iff #5314 true) +#5342 := (iff #5323 #5314) +#5343 := [rewrite]: #5342 +#5324 := (iff #5320 #5323) +#5321 := (iff #5318 true) +#5322 := [rewrite]: #5321 +#5340 := [monotonicity #5322]: #5324 +#5345 := [trans #5340 #5343]: #5344 +#7021 := [monotonicity #5345]: #7020 +#7024 := [trans #7021 #7023]: #7020 +#7019 := [quant-inst]: #7018 +#7025 := [mp #7019 #7024]: #7015 +#8579 := [unit-resolution #7025 #4335]: #5314 +#8580 := [symm #8579]: #6089 +#8035 := (= #7128 uf_16) +#7129 := (= uf_16 #7128) +#16 := (uf_2 #12) +#325 := (= #10 #16) +#4301 := (forall (vars (?x4 T2) (?x5 T2)) (:pat #4294) #325) +#329 := (forall (vars (?x4 T2) (?x5 T2)) #325) +#4304 := (iff #329 #4301) +#4302 := (iff #325 #325) +#4303 := [refl]: #4302 +#4305 := [quant-intro #4303]: #4304 +#1949 := (~ #329 #329) +#1983 := (~ #325 #325) +#1984 := [refl]: #1983 +#1950 := [nnf-pos #1984]: #1949 +#17 := (= #16 #10) +#18 := (forall (vars (?x4 T2) (?x5 T2)) #17) +#330 := (iff #18 #329) +#327 := (iff #17 #325) +#328 := [rewrite]: #327 +#331 := [quant-intro #328]: #330 +#324 := [asserted]: #18 +#334 := [mp #324 #331]: #329 +#1985 := [mp~ #334 #1950]: #329 +#4306 := [mp #1985 #4305]: #4301 +#7136 := (not #4301) +#7154 := (or #7136 #7129) +#7155 := [quant-inst]: #7154 +#8034 := [unit-resolution #7155 #4306]: #7129 +#8036 := [symm #8034]: #8035 +#8593 := [monotonicity #8036 #8580]: #8592 +#8591 := [trans #8593 #8595]: #8596 +#32982 := [monotonicity #8591 #32984]: #32980 +#32970 := [monotonicity #32982]: #32985 +#32965 := [symm #32970]: #32955 +#32971 := [monotonicity #32965]: #32968 +#32956 := (not #15367) +#32950 := [hypothesis]: #32956 +#15373 := (or #9846 #15367) +#8693 := (= #144 #2212) +#8633 := (= #2212 #144) +#6559 := (= ?x46!9 uf_16) +#7707 := (= ?x46!9 #7128) +#6330 := (uf_6 uf_15 ?x46!9) +#6365 := (= uf_8 #6330) +#7717 := (ite #7707 #5314 #6365) +#7711 := (uf_6 #7203 ?x46!9) +#7714 := (= uf_8 #7711) +#7720 := (iff #7714 #7717) +#8351 := (or #7026 #7720) +#7708 := (ite #7707 #6089 #6365) +#7712 := (= #7711 uf_8) +#7713 := (iff #7712 #7708) +#8361 := (or #7026 #7713) +#8363 := (iff #8361 #8351) +#8365 := (iff #8351 #8351) +#8366 := [rewrite]: #8365 +#7721 := (iff #7713 #7720) +#7718 := (iff #7708 #7717) +#7719 := [monotonicity #6102]: #7718 +#7715 := (iff #7712 #7714) +#7716 := [rewrite]: #7715 +#7722 := [monotonicity #7716 #7719]: #7721 +#8364 := [monotonicity #7722]: #8363 +#8367 := [trans #8364 #8366]: #8363 +#8362 := [quant-inst]: #8361 +#8368 := [mp #8362 #8367]: #8351 +#8576 := [unit-resolution #8368 #4320]: #7720 +#8601 := (= #2208 #7711) +#8597 := (= #7711 #2208) +#8598 := [monotonicity #8591]: #8597 +#8625 := [symm #8598]: #8601 +#8571 := [hypothesis]: #2836 +#4285 := (or #2831 #2209) +#4275 := [def-axiom]: #4285 +#8577 := [unit-resolution #4275 #8571]: #2209 +#8626 := [trans #8577 #8625]: #7714 +#8404 := (not #7714) +#8401 := (not #7720) +#8405 := (or #8401 #8404 #7717) +#8400 := [def-axiom]: #8405 +#8627 := [unit-resolution #8400 #8626 #8576]: #7717 +#6393 := (uf_1 uf_16 ?x46!9) +#6394 := (uf_10 #6393) +#6337 := (* -1::int #2212) +#6411 := (+ #6337 #6394) +#6412 := (+ #144 #6411) +#6413 := (>= #6412 0::int) +#8287 := (not #6413) +#6395 := (* -1::int #6394) +#6396 := (+ uf_9 #6395) +#6397 := (<= #6396 0::int) +#6421 := (or #6397 #6413) +#6426 := (not #6421) +#3935 := (not #2825) +#3940 := (or #2831 #3935) +#4276 := [def-axiom]: #3940 +#8572 := [unit-resolution #4276 #8571]: #3935 +#4212 := (or #4595 #4456) +#4213 := [def-axiom]: #4212 +#8574 := [unit-resolution #4213 #7389]: #4456 +#8302 := (or #4461 #2825 #6426) +#6398 := (+ #1449 #6395) +#6399 := (+ #2212 #6398) +#6400 := (<= #6399 0::int) +#6401 := (or #6400 #6397) +#6402 := (not #6401) +#6403 := (or #2213 #6402) +#8303 := (or #4461 #6403) +#8310 := (iff #8303 #8302) +#6429 := (or #2825 #6426) +#8305 := (or #4461 #6429) +#8308 := (iff #8305 #8302) +#8309 := [rewrite]: #8308 +#8306 := (iff #8303 #8305) +#6430 := (iff #6403 #6429) +#6427 := (iff #6402 #6426) +#6424 := (iff #6401 #6421) +#6418 := (or #6413 #6397) +#6422 := (iff #6418 #6421) +#6423 := [rewrite]: #6422 +#6419 := (iff #6401 #6418) +#6416 := (iff #6400 #6413) +#6404 := (+ #2212 #6395) +#6405 := (+ #1449 #6404) +#6408 := (<= #6405 0::int) +#6414 := (iff #6408 #6413) +#6415 := [rewrite]: #6414 +#6409 := (iff #6400 #6408) +#6406 := (= #6399 #6405) +#6407 := [rewrite]: #6406 +#6410 := [monotonicity #6407]: #6409 +#6417 := [trans #6410 #6415]: #6416 +#6420 := [monotonicity #6417]: #6419 +#6425 := [trans #6420 #6423]: #6424 +#6428 := [monotonicity #6425]: #6427 +#6431 := [monotonicity #2827 #6428]: #6430 +#8307 := [monotonicity #6431]: #8306 +#8311 := [trans #8307 #8309]: #8310 +#8304 := [quant-inst]: #8303 +#8312 := [mp #8304 #8311]: #8302 +#8628 := [unit-resolution #8312 #8574 #8572]: #6426 +#8288 := (or #6421 #8287) +#8289 := [def-axiom]: #8288 +#8629 := [unit-resolution #8289 #8628]: #8287 +#8369 := (not #7717) +#9187 := (or #7707 #6413 #8369) +#7754 := (uf_1 #7128 ?x46!9) +#7838 := (uf_3 #7754) +#9086 := (uf_4 uf_14 #7838) +#9087 := (* -1::int #9086) +#7168 := (uf_4 uf_14 #7128) +#9088 := (+ #7168 #9087) +#9089 := (>= #9088 0::int) +#9090 := (uf_6 uf_15 #7838) +#9091 := (= uf_8 #9090) +#9139 := (= #6330 #9090) +#9135 := (= #9090 #6330) +#9133 := (= #7838 ?x46!9) +#7839 := (= ?x46!9 #7838) +#8519 := (or #5378 #7839) +#8257 := [quant-inst]: #8519 +#9132 := [unit-resolution #8257 #4300]: #7839 +#9134 := [symm #9132]: #9133 +#9136 := [monotonicity #9134]: #9135 +#9140 := [symm #9136]: #9139 +#9129 := [hypothesis]: #7717 +#7733 := (not #7707) +#9130 := [hypothesis]: #7733 +#8347 := (or #8369 #7707 #6365) +#8354 := [def-axiom]: #8347 +#9131 := [unit-resolution #8354 #9130 #9129]: #6365 +#9141 := [trans #9131 #9140]: #9091 +#9092 := (not #9091) +#9154 := (or #9089 #9092) +#7246 := (uf_6 uf_15 #7128) +#7247 := (= uf_8 #7246) +#9149 := (not #7247) +#9150 := (iff #588 #9149) +#9147 := (iff #585 #7247) +#9145 := (iff #7247 #585) +#9143 := (= #7246 #141) +#9144 := [monotonicity #8036]: #9143 +#9146 := [monotonicity #9144]: #9145 +#9148 := [symm #9146]: #9147 +#9151 := [monotonicity #9148]: #9150 +#4127 := (or #4595 #588) +#4220 := [def-axiom]: #4127 +#9142 := [unit-resolution #4220 #7389]: #588 +#9152 := [mp #9142 #9151]: #9149 +#4076 := (or #4677 #4421) +#4131 := [def-axiom]: #4076 +#9153 := [unit-resolution #4131 #7304]: #4421 +#9097 := (or #4426 #7247 #9089 #9092) +#9093 := (or #9092 #7247 #9089) +#9098 := (or #4426 #9093) +#9105 := (iff #9098 #9097) +#9094 := (or #7247 #9089 #9092) +#9100 := (or #4426 #9094) +#9103 := (iff #9100 #9097) +#9104 := [rewrite]: #9103 +#9101 := (iff #9098 #9100) +#9095 := (iff #9093 #9094) +#9096 := [rewrite]: #9095 +#9102 := [monotonicity #9096]: #9101 +#9106 := [trans #9102 #9104]: #9105 +#9099 := [quant-inst]: #9098 +#9107 := [mp #9099 #9106]: #9097 +#9155 := [unit-resolution #9107 #9153 #9152]: #9154 +#9156 := [unit-resolution #9155 #9141]: #9089 +#9157 := [hypothesis]: #8287 +#7755 := (uf_10 #7754) +#7756 := (* -1::int #7755) +#8520 := (+ #6394 #7756) +#8524 := (>= #8520 0::int) +#8517 := (= #6394 #7755) +#9160 := (= #7755 #6394) +#9158 := (= #7754 #6393) +#9159 := [monotonicity #8036]: #9158 +#9161 := [monotonicity #9159]: #9160 +#9162 := [symm #9161]: #8517 +#9163 := (not #8517) +#9164 := (or #9163 #8524) +#9165 := [th-lemma]: #9164 +#9166 := [unit-resolution #9165 #9162]: #8524 +#8514 := (>= #7755 0::int) +#7798 := (<= #7755 0::int) +#7799 := (not #7798) +#7804 := (or #7707 #7799) +#59 := (uf_10 #12) +#409 := (<= #59 0::int) +#410 := (not #409) +#58 := (= #10 #11) +#413 := (or #58 #410) +#4342 := (forall (vars (?x22 T2) (?x23 T2)) (:pat #4294) #413) +#416 := (forall (vars (?x22 T2) (?x23 T2)) #413) +#4345 := (iff #416 #4342) +#4343 := (iff #413 #413) +#4344 := [refl]: #4343 +#4346 := [quant-intro #4344]: #4345 +#1963 := (~ #416 #416) +#1962 := (~ #413 #413) +#2000 := [refl]: #1962 +#1964 := [nnf-pos #2000]: #1963 +#64 := (< 0::int #59) +#63 := (not #58) +#65 := (implies #63 #64) +#66 := (forall (vars (?x22 T2) (?x23 T2)) #65) +#419 := (iff #66 #416) +#403 := (or #58 #64) +#406 := (forall (vars (?x22 T2) (?x23 T2)) #403) +#417 := (iff #406 #416) +#414 := (iff #403 #413) +#411 := (iff #64 #410) +#412 := [rewrite]: #411 +#415 := [monotonicity #412]: #414 +#418 := [quant-intro #415]: #417 +#407 := (iff #66 #406) +#404 := (iff #65 #403) +#405 := [rewrite]: #404 +#408 := [quant-intro #405]: #407 +#420 := [trans #408 #418]: #419 +#402 := [asserted]: #66 +#421 := [mp #402 #420]: #416 +#2001 := [mp~ #421 #1964]: #416 +#4347 := [mp #2001 #4346]: #4342 +#7093 := (not #4342) +#8435 := (or #7093 #7707 #7799) +#7800 := (= #7128 ?x46!9) +#7801 := (or #7800 #7799) +#8436 := (or #7093 #7801) +#8447 := (iff #8436 #8435) +#8437 := (or #7093 #7804) +#8443 := (iff #8437 #8435) +#8444 := [rewrite]: #8443 +#8438 := (iff #8436 #8437) +#7805 := (iff #7801 #7804) +#7802 := (iff #7800 #7707) +#7803 := [rewrite]: #7802 +#7806 := [monotonicity #7803]: #7805 +#8439 := [monotonicity #7806]: #8438 +#8448 := [trans #8439 #8444]: #8447 +#8434 := [quant-inst]: #8436 +#8482 := [mp #8434 #8448]: #8435 +#9167 := [unit-resolution #8482 #4347]: #7804 +#9168 := [unit-resolution #9167 #9130]: #7799 +#9169 := (or #8514 #7798) +#9170 := [th-lemma]: #9169 +#9171 := [unit-resolution #9170 #9168]: #8514 +#9126 := (+ #2212 #9087) +#9127 := (<= #9126 0::int) +#9125 := (= #2212 #9086) +#9172 := (= #9086 #2212) +#9173 := [monotonicity #9134]: #9172 +#9174 := [symm #9173]: #9125 +#9175 := (not #9125) +#9176 := (or #9175 #9127) +#9177 := [th-lemma]: #9176 +#9178 := [unit-resolution #9177 #9174]: #9127 +#7162 := (* -1::int #7168) +#7568 := (+ #144 #7162) +#7572 := (>= #7568 0::int) +#7233 := (= #144 #7168) +#9179 := (= #7168 #144) +#9180 := [monotonicity #8036]: #9179 +#9181 := [symm #9180]: #7233 +#9182 := (not #7233) +#9183 := (or #9182 #7572) +#9184 := [th-lemma]: #9183 +#9185 := [unit-resolution #9184 #9181]: #7572 +#9186 := [th-lemma #9185 #9178 #9171 #9166 #9157 #9156]: false +#9188 := [lemma #9186]: #9187 +#8624 := [unit-resolution #9188 #8629 #8627]: #7707 +#8630 := [trans #8624 #8036]: #6559 +#8634 := [monotonicity #8630]: #8633 +#8694 := [symm #8634]: #8693 +#8695 := (= #2211 #144) +#5856 := (uf_18 uf_16) +#8641 := (= #5856 #144) +#5857 := (= #144 #5856) +#5844 := (uf_1 uf_16 uf_16) +#5845 := (uf_10 #5844) +#5864 := (>= #5845 0::int) +#5848 := (* -1::int #5845) +#5849 := (+ uf_9 #5848) +#5850 := (<= #5849 0::int) +#5872 := (or #5850 #5864) +#7965 := (uf_1 #7128 #7128) +#7966 := (uf_10 #7965) +#7967 := (* -1::int #7966) +#8029 := (+ #5845 #7967) +#8033 := (>= #8029 0::int) +#8028 := (= #5845 #7966) +#8039 := (= #5844 #7965) +#8037 := (= #7965 #5844) +#8038 := [monotonicity #8036 #8036]: #8037 +#8040 := [symm #8038]: #8039 +#8041 := [monotonicity #8040]: #8028 +#8042 := (not #8028) +#8043 := (or #8042 #8033) +#8044 := [th-lemma]: #8043 +#8045 := [unit-resolution #8044 #8041]: #8033 +#7976 := (>= #7966 0::int) +#7998 := (= #7966 0::int) +#60 := (= #59 0::int) +#393 := (or #63 #60) +#4336 := (forall (vars (?x20 T2) (?x21 T2)) (:pat #4294) #393) +#396 := (forall (vars (?x20 T2) (?x21 T2)) #393) +#4339 := (iff #396 #4336) +#4337 := (iff #393 #393) +#4338 := [refl]: #4337 +#4340 := [quant-intro #4338]: #4339 +#1959 := (~ #396 #396) +#1998 := (~ #393 #393) +#1999 := [refl]: #1998 +#1960 := [nnf-pos #1999]: #1959 +#61 := (implies #58 #60) +#62 := (forall (vars (?x20 T2) (?x21 T2)) #61) +#399 := (iff #62 #396) +#372 := (= 0::int #59) +#383 := (or #63 #372) +#388 := (forall (vars (?x20 T2) (?x21 T2)) #383) +#397 := (iff #388 #396) +#394 := (iff #383 #393) +#391 := (iff #372 #60) +#392 := [rewrite]: #391 +#395 := [monotonicity #392]: #394 +#398 := [quant-intro #395]: #397 +#389 := (iff #62 #388) +#386 := (iff #61 #383) +#380 := (implies #58 #372) +#384 := (iff #380 #383) +#385 := [rewrite]: #384 +#381 := (iff #61 #380) +#378 := (iff #60 #372) +#379 := [rewrite]: #378 +#382 := [monotonicity #379]: #381 +#387 := [trans #382 #385]: #386 +#390 := [quant-intro #387]: #389 +#400 := [trans #390 #398]: #399 +#370 := [asserted]: #62 +#401 := [mp #370 #400]: #396 +#1961 := [mp~ #401 #1960]: #396 +#4341 := [mp #1961 #4340]: #4336 +#6863 := (not #4336) +#8012 := (or #6863 #7998) +#7248 := (= #7128 #7128) +#7999 := (not #7248) +#8000 := (or #7999 #7998) +#8013 := (or #6863 #8000) +#8015 := (iff #8013 #8012) +#8017 := (iff #8012 #8012) +#8018 := [rewrite]: #8017 +#8010 := (iff #8000 #7998) +#8005 := (or false #7998) +#8008 := (iff #8005 #7998) +#8009 := [rewrite]: #8008 +#8006 := (iff #8000 #8005) +#8003 := (iff #7999 false) +#8001 := (iff #7999 #6849) +#7256 := (iff #7248 true) +#7257 := [rewrite]: #7256 +#8002 := [monotonicity #7257]: #8001 +#8004 := [trans #8002 #6853]: #8003 +#8007 := [monotonicity #8004]: #8006 +#8011 := [trans #8007 #8009]: #8010 +#8016 := [monotonicity #8011]: #8015 +#8019 := [trans #8016 #8018]: #8015 +#8014 := [quant-inst]: #8013 +#8020 := [mp #8014 #8019]: #8012 +#8046 := [unit-resolution #8020 #4341]: #7998 +#8047 := (not #7998) +#8048 := (or #8047 #7976) +#8049 := [th-lemma]: #8048 +#8050 := [unit-resolution #8049 #8046]: #7976 +#6974 := (not #5864) +#8051 := [hypothesis]: #6974 +#8052 := [th-lemma #8051 #8050 #8045]: false +#8053 := [lemma #8052]: #5864 +#6975 := (or #5872 #6974) +#6976 := [def-axiom]: #6975 +#8573 := [unit-resolution #6976 #8053]: #5872 +#5877 := (not #5872) +#5880 := (or #5857 #5877) +#6955 := (or #4461 #5857 #5877) +#5851 := (+ #1449 #5848) +#5852 := (+ #144 #5851) +#5853 := (<= #5852 0::int) +#5854 := (or #5853 #5850) +#5855 := (not #5854) +#5858 := (or #5857 #5855) +#6956 := (or #4461 #5858) +#6967 := (iff #6956 #6955) +#6962 := (or #4461 #5880) +#6965 := (iff #6962 #6955) +#6966 := [rewrite]: #6965 +#6963 := (iff #6956 #6962) +#5881 := (iff #5858 #5880) +#5878 := (iff #5855 #5877) +#5875 := (iff #5854 #5872) +#5869 := (or #5864 #5850) +#5873 := (iff #5869 #5872) +#5874 := [rewrite]: #5873 +#5870 := (iff #5854 #5869) +#5867 := (iff #5853 #5864) +#5861 := (<= #5848 0::int) +#5865 := (iff #5861 #5864) +#5866 := [rewrite]: #5865 +#5862 := (iff #5853 #5861) +#5859 := (= #5852 #5848) +#5860 := [rewrite]: #5859 +#5863 := [monotonicity #5860]: #5862 +#5868 := [trans #5863 #5866]: #5867 +#5871 := [monotonicity #5868]: #5870 +#5876 := [trans #5871 #5874]: #5875 +#5879 := [monotonicity #5876]: #5878 +#5882 := [monotonicity #5879]: #5881 +#6964 := [monotonicity #5882]: #6963 +#6968 := [trans #6964 #6966]: #6967 +#6961 := [quant-inst]: #6956 +#6969 := [mp #6961 #6968]: #6955 +#8575 := [unit-resolution #6969 #8574]: #5880 +#8570 := [unit-resolution #8575 #8573]: #5857 +#8642 := [symm #8570]: #8641 +#8631 := (= #2211 #5856) +#8632 := [monotonicity #8630]: #8631 +#8696 := [trans #8632 #8642]: #8695 +#8697 := [trans #8696 #8694]: #2825 +#8698 := [unit-resolution #8572 #8697]: false +#8699 := [lemma #8698]: #2831 +#4203 := (or #4595 #4589) +#4204 := [def-axiom]: #4203 +#9435 := [unit-resolution #4204 #7389]: #4589 +#4209 := (or #4595 #4464) +#4214 := [def-axiom]: #4209 +#7390 := [unit-resolution #4214 #7389]: #4464 +#6363 := (or #2817 #4469 #4461) +#6139 := (uf_1 uf_16 ?x45!8) +#6140 := (uf_10 #6139) +#6165 := (+ #2191 #6140) +#6166 := (+ #144 #6165) +#6192 := (>= #6166 0::int) +#6169 := (= #6166 0::int) +#6144 := (* -1::int #6140) +#6145 := (+ uf_9 #6144) +#6146 := (<= #6145 0::int) +#6226 := (not #6146) +#6158 := (+ #2815 #6140) +#6159 := (+ #144 #6158) +#6160 := (>= #6159 0::int) +#6203 := (or #6146 #6160) +#6208 := (not #6203) +#6197 := (= #2190 #2192) +#6342 := (not #6197) +#6341 := [hypothesis]: #2822 +#6345 := (or #6342 #2817) +#6346 := [th-lemma]: #6345 +#6347 := [unit-resolution #6346 #6341]: #6342 +#6348 := [hypothesis]: #4456 +#6214 := (or #4461 #6197 #6208) +#6147 := (+ #1449 #6144) +#6148 := (+ #2192 #6147) +#6149 := (<= #6148 0::int) +#6193 := (or #6149 #6146) +#6194 := (not #6193) +#6195 := (= #2192 #2190) +#6196 := (or #6195 #6194) +#6215 := (or #4461 #6196) +#6222 := (iff #6215 #6214) +#6211 := (or #6197 #6208) +#6217 := (or #4461 #6211) +#6220 := (iff #6217 #6214) +#6221 := [rewrite]: #6220 +#6218 := (iff #6215 #6217) +#6212 := (iff #6196 #6211) +#6209 := (iff #6194 #6208) +#6206 := (iff #6193 #6203) +#6200 := (or #6160 #6146) +#6204 := (iff #6200 #6203) +#6205 := [rewrite]: #6204 +#6201 := (iff #6193 #6200) +#6163 := (iff #6149 #6160) +#6151 := (+ #2192 #6144) +#6152 := (+ #1449 #6151) +#6155 := (<= #6152 0::int) +#6161 := (iff #6155 #6160) +#6162 := [rewrite]: #6161 +#6156 := (iff #6149 #6155) +#6153 := (= #6148 #6152) +#6154 := [rewrite]: #6153 +#6157 := [monotonicity #6154]: #6156 +#6164 := [trans #6157 #6162]: #6163 +#6202 := [monotonicity #6164]: #6201 +#6207 := [trans #6202 #6205]: #6206 +#6210 := [monotonicity #6207]: #6209 +#6198 := (iff #6195 #6197) +#6199 := [rewrite]: #6198 +#6213 := [monotonicity #6199 #6210]: #6212 +#6219 := [monotonicity #6213]: #6218 +#6223 := [trans #6219 #6221]: #6222 +#6216 := [quant-inst]: #6215 +#6224 := [mp #6216 #6223]: #6214 +#6349 := [unit-resolution #6224 #6348 #6347]: #6208 +#6227 := (or #6203 #6226) +#6228 := [def-axiom]: #6227 +#6350 := [unit-resolution #6228 #6349]: #6226 +#6229 := (not #6160) +#6230 := (or #6203 #6229) +#6231 := [def-axiom]: #6230 +#6351 := [unit-resolution #6231 #6349]: #6229 +#6175 := (or #6146 #6160 #6169) +#6352 := [hypothesis]: #4464 +#6180 := (or #4469 #6146 #6160 #6169) +#6141 := (+ #6140 #2191) +#6142 := (+ #144 #6141) +#6143 := (= #6142 0::int) +#6150 := (or #6149 #6146 #6143) +#6181 := (or #4469 #6150) +#6188 := (iff #6181 #6180) +#6183 := (or #4469 #6175) +#6186 := (iff #6183 #6180) +#6187 := [rewrite]: #6186 +#6184 := (iff #6181 #6183) +#6178 := (iff #6150 #6175) +#6172 := (or #6160 #6146 #6169) +#6176 := (iff #6172 #6175) +#6177 := [rewrite]: #6176 +#6173 := (iff #6150 #6172) +#6170 := (iff #6143 #6169) +#6167 := (= #6142 #6166) +#6168 := [rewrite]: #6167 +#6171 := [monotonicity #6168]: #6170 +#6174 := [monotonicity #6164 #6171]: #6173 +#6179 := [trans #6174 #6177]: #6178 +#6185 := [monotonicity #6179]: #6184 +#6189 := [trans #6185 #6187]: #6188 +#6182 := [quant-inst]: #6181 +#6190 := [mp #6182 #6189]: #6180 +#6353 := [unit-resolution #6190 #6352]: #6175 +#6354 := [unit-resolution #6353 #6351 #6350]: #6169 +#6355 := (not #6169) +#6356 := (or #6355 #6192) +#6357 := [th-lemma]: #6356 +#6358 := [unit-resolution #6357 #6354]: #6192 +#6225 := (>= #2816 0::int) +#6359 := (or #6225 #2817) +#6360 := [th-lemma]: #6359 +#6361 := [unit-resolution #6360 #6341]: #6225 +#6362 := [th-lemma #6361 #6351 #6358]: false +#6364 := [lemma #6362]: #6363 +#9436 := [unit-resolution #6364 #7390 #8574]: #2817 +#4123 := (or #4592 #2822 #4586) +#4124 := [def-axiom]: #4123 +#9437 := [unit-resolution #4124 #9436 #9435]: #4586 +#4215 := (or #4583 #4577) +#4216 := [def-axiom]: #4215 +#9438 := [unit-resolution #4216 #9437]: #4577 +#4111 := (or #4580 #2836 #4574) +#4070 := [def-axiom]: #4111 +#9439 := [unit-resolution #4070 #9438]: #4577 +#9422 := [unit-resolution #9439 #8699]: #4574 +#4068 := (or #4571 #4481) +#4069 := [def-axiom]: #4068 +#9423 := [unit-resolution #4069 #9422]: #4481 +#10877 := (or #4486 #9846 #15367) +#15363 := (= #15362 #2306) +#15366 := (or #15363 #9846) +#14529 := (or #4486 #15366) +#14658 := (iff #14529 #10877) +#14681 := (or #4486 #15373) +#14554 := (iff #14681 #10877) +#14690 := [rewrite]: #14554 +#12540 := (iff #14529 #14681) +#15376 := (iff #15366 #15373) +#15370 := (or #15367 #9846) +#15374 := (iff #15370 #15373) +#15375 := [rewrite]: #15374 +#15371 := (iff #15366 #15370) +#15368 := (iff #15363 #15367) +#15369 := [rewrite]: #15368 +#15372 := [monotonicity #15369]: #15371 +#15377 := [trans #15372 #15375]: #15376 +#14677 := [monotonicity #15377]: #12540 +#14520 := [trans #14677 #14690]: #14658 +#14692 := [quant-inst]: #14529 +#10887 := [mp #14692 #14520]: #10877 +#32978 := [unit-resolution #10887 #9423]: #15373 +#32979 := [unit-resolution #32978 #32950]: #9846 +#32981 := [mp #32979 #32971]: #30313 +#18779 := (= ?x52!15 #7128) +#32989 := (iff #18779 #32602) +#32770 := (iff #32602 #18779) +#25219 := (= #7128 ?x52!15) +#25223 := (iff #25219 #18779) +#29787 := [commutativity]: #25223 +#32974 := (iff #32602 #25219) +#32975 := [monotonicity #32984]: #32974 +#32986 := [trans #32975 #29787]: #32770 +#32991 := [symm #32986]: #32989 +#15413 := (uf_1 uf_16 ?x52!15) +#15414 := (uf_10 #15413) +#15439 := (+ #2307 #15414) +#15440 := (+ #144 #15439) +#15443 := (= #15440 0::int) +#15432 := (+ #15397 #15414) +#15433 := (+ #144 #15432) +#15434 := (>= #15433 0::int) +#15418 := (* -1::int #15414) +#15419 := (+ uf_9 #15418) +#15420 := (<= #15419 0::int) +#15473 := (or #15420 #15434) +#15478 := (not #15473) +#15481 := (or #15367 #15478) +#14730 := (or #4461 #15367 #15478) +#15421 := (+ #1449 #15418) +#15422 := (+ #15362 #15421) +#15423 := (<= #15422 0::int) +#15467 := (or #15423 #15420) +#15468 := (not #15467) +#15469 := (or #15363 #15468) +#14738 := (or #4461 #15469) +#14986 := (iff #14738 #14730) +#14556 := (or #4461 #15481) +#14896 := (iff #14556 #14730) +#15034 := [rewrite]: #14896 +#14832 := (iff #14738 #14556) +#15482 := (iff #15469 #15481) +#15479 := (iff #15468 #15478) +#15476 := (iff #15467 #15473) +#15470 := (or #15434 #15420) +#15474 := (iff #15470 #15473) +#15475 := [rewrite]: #15474 +#15471 := (iff #15467 #15470) +#15437 := (iff #15423 #15434) +#15425 := (+ #15362 #15418) +#15426 := (+ #1449 #15425) +#15429 := (<= #15426 0::int) +#15435 := (iff #15429 #15434) +#15436 := [rewrite]: #15435 +#15430 := (iff #15423 #15429) +#15427 := (= #15422 #15426) +#15428 := [rewrite]: #15427 +#15431 := [monotonicity #15428]: #15430 +#15438 := [trans #15431 #15436]: #15437 +#15472 := [monotonicity #15438]: #15471 +#15477 := [trans #15472 #15475]: #15476 +#15480 := [monotonicity #15477]: #15479 +#15483 := [monotonicity #15369 #15480]: #15482 +#15031 := [monotonicity #15483]: #14832 +#14985 := [trans #15031 #15034]: #14986 +#12501 := [quant-inst]: #14738 +#15678 := [mp #12501 #14985]: #14730 +#32969 := [unit-resolution #15678 #8574]: #15481 +#32952 := [unit-resolution #32969 #32950]: #15478 +#29629 := (or #15473 #15443) +#25301 := (not #15443) +#29623 := [hypothesis]: #25301 +#15187 := (not #15420) +#29624 := [hypothesis]: #15478 +#14833 := (or #15473 #15187) +#15233 := [def-axiom]: #14833 +#29625 := [unit-resolution #15233 #29624]: #15187 +#8899 := (not #15434) +#15050 := (or #15473 #8899) +#15028 := [def-axiom]: #15050 +#29626 := [unit-resolution #15028 #29624]: #8899 +#15449 := (or #15420 #15434 #15443) +#12503 := (or #4469 #15420 #15434 #15443) +#15415 := (+ #15414 #2307) +#15416 := (+ #144 #15415) +#15417 := (= #15416 0::int) +#15424 := (or #15423 #15420 #15417) +#12502 := (or #4469 #15424) +#14824 := (iff #12502 #12503) +#14693 := (or #4469 #15449) +#14698 := (iff #14693 #12503) +#14734 := [rewrite]: #14698 +#14675 := (iff #12502 #14693) +#15452 := (iff #15424 #15449) +#15446 := (or #15434 #15420 #15443) +#15450 := (iff #15446 #15449) +#15451 := [rewrite]: #15450 +#15447 := (iff #15424 #15446) +#15444 := (iff #15417 #15443) +#15441 := (= #15416 #15440) +#15442 := [rewrite]: #15441 +#15445 := [monotonicity #15442]: #15444 +#15448 := [monotonicity #15438 #15445]: #15447 +#15453 := [trans #15448 #15451]: #15452 +#14648 := [monotonicity #15453]: #14675 +#14687 := [trans #14648 #14734]: #14824 +#14736 := [quant-inst]: #12502 +#13488 := [mp #14736 #14687]: #12503 +#29627 := [unit-resolution #13488 #7390]: #15449 +#29628 := [unit-resolution #29627 #29626 #29625 #29623]: false +#29630 := [lemma #29628]: #29629 +#32972 := [unit-resolution #29630 #32952]: #15443 +#29799 := (or #25301 #18779) +#7126 := (uf_3 #6008) +#15598 := (uf_1 #7126 ?x52!15) +#27533 := (uf_3 #15598) +#28710 := (uf_1 #7128 #27533) +#28711 := (uf_10 #28710) +#28714 := (* -1::int #28711) +#28814 := (+ #15414 #28714) +#28506 := (>= #28814 0::int) +#28505 := (= #15414 #28711) +#29767 := (= #28711 #15414) +#29765 := (= #28710 #15413) +#29763 := (= #27533 ?x52!15) +#27534 := (= ?x52!15 #27533) +#27563 := (or #5378 #27534) +#27564 := [quant-inst]: #27563 +#29762 := [unit-resolution #27564 #4300]: #27534 +#29764 := [symm #29762]: #29763 +#29766 := [monotonicity #8036 #29764]: #29765 +#29768 := [monotonicity #29766]: #29767 +#29769 := [symm #29768]: #28505 +#29770 := (not #28505) +#29771 := (or #29770 #28506) +#29772 := [th-lemma]: #29771 +#29773 := [unit-resolution #29772 #29769]: #28506 +#5902 := (* -1::int #5856) +#6232 := (+ #144 #5902) +#6233 := (>= #6232 0::int) +#4218 := (or #4583 #4472) +#4120 := [def-axiom]: #4218 +#14110 := [unit-resolution #4120 #9437]: #4472 +#6959 := (or #4477 #6233) +#6960 := [quant-inst]: #6959 +#12934 := [unit-resolution #6960 #14110]: #6233 +#7167 := (uf_18 #7128) +#8141 := (* -1::int #7167) +#10187 := (+ #5856 #8141) +#7462 := (>= #10187 0::int) +#10181 := (= #5856 #7167) +#14102 := (= #7167 #5856) +#14103 := [monotonicity #8036]: #14102 +#14104 := [symm #14103]: #10181 +#14105 := (not #10181) +#25236 := (or #14105 #7462) +#25243 := [th-lemma]: #25236 +#25235 := [unit-resolution #25243 #14104]: #7462 +#14406 := (<= #15440 0::int) +#25300 := [hypothesis]: #15443 +#25302 := (or #25301 #14406) +#25303 := [th-lemma]: #25302 +#25304 := [unit-resolution #25303 #25300]: #14406 +#15344 := (+ #2306 #8141) +#15518 := (<= #15344 0::int) +#7164 := (uf_6 uf_17 #7128) +#7165 := (= uf_8 #7164) +#25365 := (= #5319 #7164) +#25361 := (= #7164 #5319) +#25364 := [monotonicity #8578 #8036]: #25361 +#25366 := [symm #25364]: #25365 +#25367 := [trans #8579 #25366]: #7165 +#15503 := (uf_1 #7128 ?x52!15) +#15504 := (uf_10 #15503) +#15530 := (* -1::int #15504) +#15531 := (+ #8141 #15530) +#15532 := (+ #2306 #15531) +#15533 := (= #15532 0::int) +#25324 := (or #25301 #15533) +#15538 := (not #15533) +#25256 := [hypothesis]: #15538 +#14445 := (>= #15532 0::int) +#14444 := (+ #15414 #15530) +#14494 := (>= #14444 0::int) +#14488 := (= #15414 #15504) +#25275 := (= #15504 #15414) +#25257 := (= #15503 #15413) +#25274 := [monotonicity #8036]: #25257 +#25270 := [monotonicity #25274]: #25275 +#25276 := [symm #25270]: #14488 +#25277 := (not #14488) +#25278 := (or #25277 #14494) +#25279 := [th-lemma]: #25278 +#25280 := [unit-resolution #25279 #25276]: #14494 +#25306 := (not #14406) +#25305 := (not #14494) +#13409 := (not #6233) +#25299 := (not #7462) +#25307 := (or #14445 #25299 #13409 #25305 #25306) +#25308 := [th-lemma]: #25307 +#25309 := [unit-resolution #25308 #25304 #25235 #12934 #25280]: #14445 +#14391 := (<= #15532 0::int) +#14441 := (<= #14444 0::int) +#25316 := (or #25277 #14441) +#25317 := [th-lemma]: #25316 +#25315 := [unit-resolution #25317 #25276]: #14441 +#6970 := (<= #6232 0::int) +#14098 := (not #5857) +#14099 := (or #14098 #6970) +#14100 := [th-lemma]: #14099 +#14101 := [unit-resolution #14100 #8570]: #6970 +#10188 := (<= #10187 0::int) +#14106 := (or #14105 #10188) +#14107 := [th-lemma]: #14106 +#14108 := [unit-resolution #14107 #14104]: #10188 +#14435 := (>= #15440 0::int) +#25318 := (or #25301 #14435) +#25319 := [th-lemma]: #25318 +#25320 := [unit-resolution #25319 #25300]: #14435 +#25333 := (not #14435) +#25332 := (not #14441) +#12642 := (not #6970) +#25331 := (not #10188) +#25334 := (or #14391 #25331 #12642 #25332 #25333) +#25335 := [th-lemma]: #25334 +#25336 := [unit-resolution #25335 #25320 #14108 #14101 #25315]: #14391 +#25314 := (not #14445) +#25337 := (not #14391) +#25321 := (or #15533 #25337 #25314) +#25322 := [th-lemma]: #25321 +#25323 := [unit-resolution #25322 #25336 #25309 #25256]: false +#25313 := [lemma #25323]: #25324 +#29774 := [unit-resolution #25313 #25300]: #15533 +#7166 := (not #7165) +#15541 := (or #7166 #15518 #15538) +#6002 := (+ #108 #1449) +#7864 := (<= #6002 0::int) +#23377 := (= #108 #144) +#12197 := (= #144 #108) +#5945 := (= uf_16 uf_11) +#5947 := (= uf_11 uf_16) +#5928 := (?x40!7 uf_16) +#5932 := (uf_4 uf_14 #5928) +#5933 := (* -1::int #5932) +#5929 := (uf_1 #5928 uf_16) +#5930 := (uf_10 #5929) +#5931 := (* -1::int #5930) +#5950 := (+ #5931 #5933) +#5951 := (+ #144 #5950) +#5954 := (= #5951 0::int) +#5957 := (not #5954) +#5940 := (uf_6 uf_15 #5928) +#5941 := (= uf_8 #5940) +#5942 := (not #5941) +#5938 := (+ #144 #5933) +#5939 := (<= #5938 0::int) +#5963 := (or #5939 #5942 #5957) +#6003 := (>= #6002 0::int) +#9539 := (not #7864) +#23470 := [hypothesis]: #9539 +#23505 := (or #7864 #6003) +#23465 := [th-lemma]: #23505 +#23464 := [unit-resolution #23465 #23470]: #6003 +#9672 := (not #6003) +#18009 := (or #9672 #5939) +#7466 := (>= #5932 0::int) +#8252 := (not #7466) +#8253 := [hypothesis]: #8252 +#8212 := (or #4409 #7466) +#8206 := [quant-inst]: #8212 +#8263 := [unit-resolution #8206 #7305 #8253]: false +#8264 := [lemma #8263]: #7466 +#17999 := (or #9672 #8252 #5939) +#4050 := (<= #108 0::int) +#7308 := (or #1854 #4050) +#7309 := [th-lemma]: #7308 +#7310 := [unit-resolution #7309 #7307]: #4050 +#5512 := (not #4050) +#17980 := (or #9672 #5512 #8252 #5939) +#17982 := [th-lemma]: #17980 +#17995 := [unit-resolution #17982 #7310]: #17999 +#18007 := [unit-resolution #17995 #8264]: #18009 +#23469 := [unit-resolution #18007 #23464]: #5939 +#7005 := (not #5939) +#7006 := (or #5963 #7005) +#7007 := [def-axiom]: #7006 +#23506 := [unit-resolution #7007 #23469]: #5963 +#5968 := (not #5963) +#18000 := (or #5947 #5968) +#4208 := (or #4595 #1657) +#4210 := [def-axiom]: #4208 +#16348 := [unit-resolution #4210 #7389]: #1657 +#6992 := (or #4442 #1656 #5947 #5968) +#5934 := (+ #5933 #5931) +#5935 := (+ #144 #5934) +#5936 := (= #5935 0::int) +#5937 := (not #5936) +#5943 := (or #5942 #5939 #5937) +#5944 := (not #5943) +#5946 := (or #5945 #1656 #5944) +#6993 := (or #4442 #5946) +#7000 := (iff #6993 #6992) +#5974 := (or #1656 #5947 #5968) +#6995 := (or #4442 #5974) +#6998 := (iff #6995 #6992) +#6999 := [rewrite]: #6998 +#6996 := (iff #6993 #6995) +#5977 := (iff #5946 #5974) +#5971 := (or #5947 #1656 #5968) +#5975 := (iff #5971 #5974) +#5976 := [rewrite]: #5975 +#5972 := (iff #5946 #5971) +#5969 := (iff #5944 #5968) +#5966 := (iff #5943 #5963) +#5960 := (or #5942 #5939 #5957) +#5964 := (iff #5960 #5963) +#5965 := [rewrite]: #5964 +#5961 := (iff #5943 #5960) +#5958 := (iff #5937 #5957) +#5955 := (iff #5936 #5954) +#5952 := (= #5935 #5951) +#5953 := [rewrite]: #5952 +#5956 := [monotonicity #5953]: #5955 +#5959 := [monotonicity #5956]: #5958 +#5962 := [monotonicity #5959]: #5961 +#5967 := [trans #5962 #5965]: #5966 +#5970 := [monotonicity #5967]: #5969 +#5948 := (iff #5945 #5947) +#5949 := [rewrite]: #5948 +#5973 := [monotonicity #5949 #5970]: #5972 +#5978 := [trans #5973 #5976]: #5977 +#6997 := [monotonicity #5978]: #6996 +#7001 := [trans #6997 #6999]: #7000 +#6994 := [quant-inst]: #6993 +#7002 := [mp #6994 #7001]: #6992 +#18030 := [unit-resolution #7002 #8030 #16348]: #18000 +#23510 := [unit-resolution #18030 #23506]: #5947 +#12009 := [symm #23510]: #5945 +#12163 := [monotonicity #12009]: #12197 +#12023 := [symm #12163]: #23377 +#12215 := (not #23377) +#12216 := (or #12215 #7864) +#11464 := [th-lemma]: #12216 +#12217 := [unit-resolution #11464 #23470 #12023]: false +#12311 := [lemma #12217]: #7864 +#9540 := (or #2234 #9539) +#6554 := (uf_1 uf_16 ?x47!10) +#6555 := (uf_10 #6554) +#6435 := (* -1::int #2233) +#6597 := (+ #6435 #6555) +#6598 := (+ #144 #6597) +#8322 := (<= #6598 0::int) +#6601 := (= #6598 0::int) +#6538 := (* -1::int #6555) +#6539 := (+ uf_9 #6538) +#6540 := (<= #6539 0::int) +#8336 := (not #6540) +#6327 := (uf_4 uf_14 ?x47!10) +#6471 := (* -1::int #6327) +#6590 := (+ #6471 #6555) +#6591 := (+ #144 #6590) +#6592 := (>= #6591 0::int) +#6650 := (or #6540 #6592) +#6660 := (not #6650) +#6344 := (= #2233 #6327) +#9327 := (not #6344) +#6472 := (+ #2233 #6471) +#8301 := (>= #6472 0::int) +#9317 := (not #8301) +#9321 := [hypothesis]: #2235 +#6814 := (>= #6327 0::int) +#7656 := (or #4409 #6814) +#7657 := [quant-inst]: #7656 +#9322 := [unit-resolution #7657 #7305]: #6814 +#9323 := (not #6814) +#9324 := (or #9317 #2234 #9323) +#9325 := [th-lemma]: #9324 +#9326 := [unit-resolution #9325 #9322 #9321]: #9317 +#9347 := (or #9327 #8301) +#9348 := [th-lemma]: #9347 +#9349 := [unit-resolution #9348 #9326]: #9327 +#6662 := (or #6344 #6660) +#8339 := (or #4461 #6344 #6660) +#6541 := (+ #1449 #6538) +#6536 := (+ #6327 #6541) +#6542 := (<= #6536 0::int) +#6641 := (or #6542 #6540) +#6642 := (not #6641) +#6328 := (= #6327 #2233) +#6643 := (or #6328 #6642) +#8314 := (or #4461 #6643) +#8333 := (iff #8314 #8339) +#8281 := (or #4461 #6662) +#8343 := (iff #8281 #8339) +#8332 := [rewrite]: #8343 +#8341 := (iff #8314 #8281) +#6663 := (iff #6643 #6662) +#6661 := (iff #6642 #6660) +#6655 := (iff #6641 #6650) +#6649 := (or #6592 #6540) +#6653 := (iff #6649 #6650) +#6654 := [rewrite]: #6653 +#6651 := (iff #6641 #6649) +#6595 := (iff #6542 #6592) +#6544 := (+ #6327 #6538) +#6545 := (+ #1449 #6544) +#6562 := (<= #6545 0::int) +#6593 := (iff #6562 #6592) +#6594 := [rewrite]: #6593 +#6587 := (iff #6542 #6562) +#6546 := (= #6536 #6545) +#6561 := [rewrite]: #6546 +#6589 := [monotonicity #6561]: #6587 +#6596 := [trans #6589 #6594]: #6595 +#6652 := [monotonicity #6596]: #6651 +#6658 := [trans #6652 #6654]: #6655 +#6659 := [monotonicity #6658]: #6661 +#6383 := (iff #6328 #6344) +#6384 := [rewrite]: #6383 +#6664 := [monotonicity #6384 #6659]: #6663 +#8342 := [monotonicity #6664]: #8341 +#8334 := [trans #8342 #8332]: #8333 +#8340 := [quant-inst]: #8314 +#8335 := [mp #8340 #8334]: #8339 +#9350 := [unit-resolution #8335 #8574]: #6662 +#9351 := [unit-resolution #9350 #9349]: #6660 +#8327 := (or #6650 #8336) +#8352 := [def-axiom]: #8327 +#9346 := [unit-resolution #8352 #9351]: #8336 +#8360 := (not #6592) +#8348 := (or #6650 #8360) +#8353 := [def-axiom]: #8348 +#9352 := [unit-resolution #8353 #9351]: #8360 +#6607 := (or #6540 #6592 #6601) +#8318 := (or #4469 #6540 #6592 #6601) +#6556 := (+ #6555 #6435) +#6557 := (+ #144 #6556) +#6537 := (= #6557 0::int) +#6543 := (or #6542 #6540 #6537) +#8344 := (or #4469 #6543) +#8331 := (iff #8344 #8318) +#8349 := (or #4469 #6607) +#8330 := (iff #8349 #8318) +#8325 := [rewrite]: #8330 +#8328 := (iff #8344 #8349) +#6578 := (iff #6543 #6607) +#6604 := (or #6592 #6540 #6601) +#6579 := (iff #6604 #6607) +#6580 := [rewrite]: #6579 +#6605 := (iff #6543 #6604) +#6602 := (iff #6537 #6601) +#6599 := (= #6557 #6598) +#6600 := [rewrite]: #6599 +#6603 := [monotonicity #6600]: #6602 +#6606 := [monotonicity #6596 #6603]: #6605 +#6581 := [trans #6606 #6580]: #6578 +#8329 := [monotonicity #6581]: #8328 +#8320 := [trans #8329 #8325]: #8331 +#8345 := [quant-inst]: #8344 +#8321 := [mp #8345 #8320]: #8318 +#9353 := [unit-resolution #8321 #7390]: #6607 +#9354 := [unit-resolution #9353 #9352 #9346]: #6601 +#9355 := (not #6601) +#9356 := (or #9355 #8322) +#9367 := [th-lemma]: #9356 +#9368 := [unit-resolution #9367 #9354]: #8322 +#9369 := [hypothesis]: #7864 +#4041 := (>= #108 0::int) +#9370 := (or #1854 #4041) +#9371 := [th-lemma]: #9370 +#9366 := [unit-resolution #9371 #7307]: #4041 +#8900 := (uf_1 #7128 ?x47!10) +#8901 := (uf_10 #8900) +#8908 := (* -1::int #8901) +#9307 := (+ #6555 #8908) +#9320 := (>= #9307 0::int) +#9304 := (= #6555 #8901) +#9374 := (= #8901 #6555) +#9372 := (= #8900 #6554) +#9373 := [monotonicity #8036]: #9372 +#9375 := [monotonicity #9373]: #9374 +#9376 := [symm #9375]: #9304 +#9382 := (not #9304) +#9383 := (or #9382 #9320) +#9432 := [th-lemma]: #9383 +#9433 := [unit-resolution #9432 #9376]: #9320 +#9287 := (>= #8901 0::int) +#9080 := (<= #8901 0::int) +#9207 := (not #9080) +#8373 := (= ?x47!10 #7128) +#8730 := (not #8373) +#6710 := (uf_6 uf_15 ?x47!10) +#6711 := (= uf_8 #6710) +#8599 := (ite #8373 #5314 #6711) +#8729 := (not #8599) +#7658 := (uf_6 #7203 ?x47!10) +#8338 := (= uf_8 #7658) +#8700 := (iff #8338 #8599) +#8684 := (or #7026 #8700) +#7640 := (ite #8373 #6089 #6711) +#7653 := (= #7658 uf_8) +#8337 := (iff #7653 #7640) +#8723 := (or #7026 #8337) +#8725 := (iff #8723 #8684) +#8726 := (iff #8684 #8684) +#8722 := [rewrite]: #8726 +#8682 := (iff #8337 #8700) +#8326 := (iff #7640 #8599) +#8600 := [monotonicity #6102]: #8326 +#8410 := (iff #7653 #8338) +#8521 := [rewrite]: #8410 +#8683 := [monotonicity #8521 #8600]: #8682 +#8720 := [monotonicity #8683]: #8725 +#8727 := [trans #8720 #8722]: #8725 +#8724 := [quant-inst]: #8723 +#8728 := [mp #8724 #8727]: #8684 +#9434 := [unit-resolution #8728 #4320]: #8700 +#8895 := (not #8338) +#6323 := (uf_6 uf_17 ?x47!10) +#6325 := (= uf_8 #6323) +#6326 := (not #6325) +#9431 := (iff #6326 #8895) +#9429 := (iff #6325 #8338) +#9427 := (iff #8338 #6325) +#9426 := (= #7658 #6323) +#9421 := [monotonicity #8591]: #9426 +#9428 := [monotonicity #9421]: #9427 +#9430 := [symm #9428]: #9429 +#9460 := [monotonicity #9430]: #9431 +#6382 := (or #6326 #6344) +#8268 := (or #4486 #6326 #6344) +#6343 := (or #6328 #6326) +#8269 := (or #4486 #6343) +#8297 := (iff #8269 #8268) +#8294 := (or #4486 #6382) +#8296 := (iff #8294 #8268) +#8283 := [rewrite]: #8296 +#8284 := (iff #8269 #8294) +#6390 := (iff #6343 #6382) +#6385 := (or #6344 #6326) +#6388 := (iff #6385 #6382) +#6389 := [rewrite]: #6388 +#6386 := (iff #6343 #6385) +#6387 := [monotonicity #6384]: #6386 +#6391 := [trans #6387 #6389]: #6390 +#8295 := [monotonicity #6391]: #8284 +#8298 := [trans #8295 #8283]: #8297 +#8293 := [quant-inst]: #8269 +#8299 := [mp #8293 #8298]: #8268 +#9424 := [unit-resolution #8299 #9423]: #6382 +#9425 := [unit-resolution #9424 #9349]: #6326 +#9461 := [mp #9425 #9460]: #8895 +#8917 := (not #8700) +#8920 := (or #8917 #8338 #8729) +#8894 := [def-axiom]: #8920 +#9462 := [unit-resolution #8894 #9461 #9434]: #8729 +#9463 := (or #8599 #8730) +#7040 := (not #5314) +#8907 := (or #8599 #8730 #7040) +#8910 := [def-axiom]: #8907 +#9464 := [unit-resolution #8910 #8579]: #9463 +#9459 := [unit-resolution #9464 #9462]: #8730 +#9206 := (or #8373 #9207) +#9214 := (or #7093 #8373 #9207) +#9208 := (= #7128 ?x47!10) +#9209 := (or #9208 #9207) +#9215 := (or #7093 #9209) +#9241 := (iff #9215 #9214) +#9242 := (or #7093 #9206) +#9245 := (iff #9242 #9214) +#9246 := [rewrite]: #9245 +#9243 := (iff #9215 #9242) +#9212 := (iff #9209 #9206) +#9210 := (iff #9208 #8373) +#9211 := [rewrite]: #9210 +#9213 := [monotonicity #9211]: #9212 +#9244 := [monotonicity #9213]: #9243 +#9247 := [trans #9244 #9246]: #9241 +#9216 := [quant-inst]: #9215 +#9248 := [mp #9216 #9247]: #9214 +#9465 := [unit-resolution #9248 #4347]: #9206 +#9466 := [unit-resolution #9465 #9459]: #9207 +#9467 := (or #9287 #9080) +#9468 := [th-lemma]: #9467 +#9469 := [unit-resolution #9468 #9466]: #9287 +#9538 := [th-lemma #9321 #9469 #9433 #9366 #9369 #9368]: false +#9541 := [lemma #9538]: #9540 +#29775 := [unit-resolution #9541 #12311]: #2234 +#4222 := (or #4571 #4565) +#4223 := [def-axiom]: #4222 +#25326 := [unit-resolution #4223 #9422]: #4565 +#25357 := (or #4568 #4562) +#6056 := (= #108 #172) +#25354 := (iff #6056 #173) +#25353 := [commutativity]: #1490 +#25351 := (iff #6056 #645) +#25352 := [monotonicity #7307]: #25351 +#25355 := [trans #25352 #25353]: #25354 +#6015 := (uf_10 #6008) +#6019 := (* -1::int #6015) +#6022 := (+ #1449 #6019) +#6023 := (+ #108 #6022) +#6024 := (<= #6023 0::int) +#6020 := (+ uf_9 #6019) +#6021 := (<= #6020 0::int) +#6058 := (or #6021 #6024) +#7125 := (>= #6015 0::int) +#7105 := (= #6015 0::int) +#7087 := (<= #6015 0::int) +#4062 := (not #6024) +#7293 := [hypothesis]: #4062 +#7312 := (or #7087 #6024) +#7088 := (not #7087) +#7292 := [hypothesis]: #7088 +#6001 := (>= #144 0::int) +#7016 := (or #4409 #6001) +#7017 := [quant-inst]: #7016 +#7306 := [unit-resolution #7017 #7305]: #6001 +#7311 := [th-lemma #7310 #7306 #7293 #7292]: false +#7313 := [lemma #7311]: #7312 +#7186 := [unit-resolution #7313 #7293]: #7087 +#7090 := (or #5947 #7088) +#7094 := (or #7093 #5947 #7088) +#7089 := (or #5945 #7088) +#7095 := (or #7093 #7089) +#7102 := (iff #7095 #7094) +#7097 := (or #7093 #7090) +#7100 := (iff #7097 #7094) +#7101 := [rewrite]: #7100 +#7098 := (iff #7095 #7097) +#7091 := (iff #7089 #7090) +#7092 := [monotonicity #5949]: #7091 +#7099 := [monotonicity #7092]: #7098 +#7103 := [trans #7099 #7101]: #7102 +#7096 := [quant-inst]: #7095 +#7104 := [mp #7096 #7103]: #7094 +#7187 := [unit-resolution #7104 #4347]: #7090 +#7182 := [unit-resolution #7187 #7186]: #5947 +#7108 := (not #5947) +#7111 := (or #7108 #7105) +#7114 := (or #6863 #7108 #7105) +#7106 := (not #5945) +#7107 := (or #7106 #7105) +#7115 := (or #6863 #7107) +#7122 := (iff #7115 #7114) +#7117 := (or #6863 #7111) +#7120 := (iff #7117 #7114) +#7121 := [rewrite]: #7120 +#7118 := (iff #7115 #7117) +#7112 := (iff #7107 #7111) +#7109 := (iff #7106 #7108) +#7110 := [monotonicity #5949]: #7109 +#7113 := [monotonicity #7110]: #7112 +#7119 := [monotonicity #7113]: #7118 +#7123 := [trans #7119 #7121]: #7122 +#7116 := [quant-inst]: #7115 +#7124 := [mp #7116 #7123]: #7114 +#7188 := [unit-resolution #7124 #4341]: #7111 +#7189 := [unit-resolution #7188 #7182]: #7105 +#7190 := (not #7105) +#7191 := (or #7190 #7125) +#7192 := [th-lemma]: #7191 +#7196 := [unit-resolution #7192 #7189]: #7125 +#7197 := [th-lemma #7310 #7306 #7293 #7196]: false +#7195 := [lemma #7197]: #6024 +#5757 := (or #6058 #4062) +#5573 := [def-axiom]: #5757 +#25327 := [unit-resolution #5573 #7195]: #6058 +#6061 := (not #6058) +#6064 := (or #6056 #6061) +#6681 := (or #4461 #6056 #6061) +#6054 := (or #6024 #6021) +#6055 := (not #6054) +#6057 := (or #6056 #6055) +#6682 := (or #4461 #6057) +#6844 := (iff #6682 #6681) +#6138 := (or #4461 #6064) +#6392 := (iff #6138 #6681) +#6683 := [rewrite]: #6392 +#6118 := (iff #6682 #6138) +#6065 := (iff #6057 #6064) +#6062 := (iff #6055 #6061) +#6059 := (iff #6054 #6058) +#6060 := [rewrite]: #6059 +#6063 := [monotonicity #6060]: #6062 +#6066 := [monotonicity #6063]: #6065 +#6735 := [monotonicity #6066]: #6118 +#6845 := [trans #6735 #6683]: #6844 +#6137 := [quant-inst]: #6682 +#6878 := [mp #6137 #6845]: #6681 +#25328 := [unit-resolution #6878 #8574]: #6064 +#25329 := [unit-resolution #25328 #25327]: #6056 +#25356 := [mp #25329 #25355]: #173 +#4237 := (or #4568 #1492 #4562) +#4066 := [def-axiom]: #4237 +#25358 := [unit-resolution #4066 #25356]: #25357 +#25359 := [unit-resolution #25358 #25326]: #4562 +#4232 := (or #4559 #4553) +#4233 := [def-axiom]: #4232 +#25339 := [unit-resolution #4233 #25359]: #4553 +#4087 := (or #4556 #2235 #4550) +#4088 := [def-axiom]: #4087 +#25340 := [unit-resolution #4088 #25339]: #4553 +#29776 := [unit-resolution #25340 #29775]: #4550 +#4242 := (or #4547 #4541) +#4243 := [def-axiom]: #4242 +#29777 := [unit-resolution #4243 #29776]: #4541 +#25343 := (or #4544 #4538) +#12812 := (= #2249 #5856) +#12998 := (= ?x48!12 uf_16) +#10849 := (= ?x48!12 #7128) +#10847 := (uf_6 uf_15 ?x48!12) +#10848 := (= uf_8 #10847) +#10857 := (ite #10849 #5314 #10848) +#10851 := (uf_6 #7203 ?x48!12) +#10854 := (= uf_8 #10851) +#10860 := (iff #10854 #10857) +#12152 := (or #7026 #10860) +#10850 := (ite #10849 #6089 #10848) +#10852 := (= #10851 uf_8) +#10853 := (iff #10852 #10850) +#12155 := (or #7026 #10853) +#10823 := (iff #12155 #12152) +#10879 := (iff #12152 #12152) +#10880 := [rewrite]: #10879 +#10861 := (iff #10853 #10860) +#10858 := (iff #10850 #10857) +#10859 := [monotonicity #6102]: #10858 +#10855 := (iff #10852 #10854) +#10856 := [rewrite]: #10855 +#10862 := [monotonicity #10856 #10859]: #10861 +#10824 := [monotonicity #10862]: #10823 +#11111 := [trans #10824 #10880]: #10823 +#12156 := [quant-inst]: #12155 +#11091 := [mp #12156 #11111]: #12152 +#13286 := [unit-resolution #11091 #4320]: #10860 +#12615 := (= #2254 #10851) +#12608 := (= #10851 #2254) +#12613 := [monotonicity #8591]: #12608 +#12730 := [symm #12613]: #12615 +#12965 := [hypothesis]: #3403 +#3920 := (or #3398 #2255) +#4261 := [def-axiom]: #3920 +#12611 := [unit-resolution #4261 #12965]: #2255 +#13209 := [trans #12611 #12730]: #10854 +#11918 := (not #10854) +#10920 := (not #10860) +#11919 := (or #10920 #11918 #10857) +#12057 := [def-axiom]: #11919 +#13217 := [unit-resolution #12057 #13209 #13286]: #10857 +#10236 := (not #10848) +#11183 := (uf_4 uf_14 ?x48!12) +#11200 := (* -1::int #11183) +#13667 := (+ #7168 #11200) +#13668 := (>= #13667 0::int) +#13764 := (not #13668) +#12813 := (+ #2249 #5902) +#12814 := (<= #12813 0::int) +#13407 := (not #12814) +#11572 := (uf_4 uf_14 ?x49!11) +#11589 := (* -1::int #11572) +#11709 := (+ #144 #11589) +#11710 := (<= #11709 0::int) +#11467 := (uf_6 uf_15 ?x49!11) +#11468 := (= uf_8 #11467) +#12026 := (not #11468) +#11469 := (= ?x49!11 #7128) +#11477 := (ite #11469 #5314 #11468) +#12038 := (not #11477) +#11471 := (uf_6 #7203 ?x49!11) +#11474 := (= uf_8 #11471) +#11480 := (iff #11474 #11477) +#12030 := (or #7026 #11480) +#11470 := (ite #11469 #6089 #11468) +#11472 := (= #11471 uf_8) +#11473 := (iff #11472 #11470) +#12028 := (or #7026 #11473) +#12024 := (iff #12028 #12030) +#12033 := (iff #12030 #12030) +#12035 := [rewrite]: #12033 +#11481 := (iff #11473 #11480) +#11478 := (iff #11470 #11477) +#11479 := [monotonicity #6102]: #11478 +#11475 := (iff #11472 #11474) +#11476 := [rewrite]: #11475 +#11482 := [monotonicity #11476 #11479]: #11481 +#12032 := [monotonicity #11482]: #12024 +#12036 := [trans #12032 #12035]: #12024 +#12031 := [quant-inst]: #12028 +#12034 := [mp #12031 #12036]: #12030 +#13262 := [unit-resolution #12034 #4320]: #11480 +#12051 := (not #11474) +#13387 := (iff #2258 #12051) +#13353 := (iff #2257 #11474) +#12939 := (iff #11474 #2257) +#13219 := (= #11471 #2256) +#13243 := [monotonicity #8591]: #13219 +#13039 := [monotonicity #13243]: #12939 +#13377 := [symm #13039]: #13353 +#13388 := [monotonicity #13377]: #13387 +#3924 := (or #3398 #2258) +#3925 := [def-axiom]: #3924 +#12937 := [unit-resolution #3925 #12965]: #2258 +#12543 := [mp #12937 #13388]: #12051 +#12047 := (not #11480) +#12048 := (or #12047 #11474 #12038) +#12050 := [def-axiom]: #12048 +#12544 := [unit-resolution #12050 #12543 #13262]: #12038 +#12039 := (not #11469) +#12539 := (or #11477 #12039) +#12042 := (or #11477 #12039 #7040) +#12043 := [def-axiom]: #12042 +#12545 := [unit-resolution #12043 #8579]: #12539 +#12546 := [unit-resolution #12545 #12544]: #12039 +#12044 := (or #11477 #11469 #12026) +#12045 := [def-axiom]: #12044 +#12548 := [unit-resolution #12045 #12546 #12544]: #12026 +#11715 := (or #11468 #11710) +#4217 := (or #4595 #4446) +#4221 := [def-axiom]: #4217 +#12574 := [unit-resolution #4221 #7389]: #4446 +#12438 := (or #4451 #11468 #11710) +#11700 := (+ #11572 #1449) +#11701 := (>= #11700 0::int) +#11702 := (or #11468 #11701) +#12444 := (or #4451 #11702) +#12454 := (iff #12444 #12438) +#12448 := (or #4451 #11715) +#12452 := (iff #12448 #12438) +#12453 := [rewrite]: #12452 +#12450 := (iff #12444 #12448) +#11716 := (iff #11702 #11715) +#11713 := (iff #11701 #11710) +#11703 := (+ #1449 #11572) +#11706 := (>= #11703 0::int) +#11711 := (iff #11706 #11710) +#11712 := [rewrite]: #11711 +#11707 := (iff #11701 #11706) +#11704 := (= #11700 #11703) +#11705 := [rewrite]: #11704 +#11708 := [monotonicity #11705]: #11707 +#11714 := [trans #11708 #11712]: #11713 +#11717 := [monotonicity #11714]: #11716 +#12451 := [monotonicity #11717]: #12450 +#12449 := [trans #12451 #12453]: #12454 +#12447 := [quant-inst]: #12444 +#12455 := [mp #12447 #12449]: #12438 +#12575 := [unit-resolution #12455 #12574]: #11715 +#12576 := [unit-resolution #12575 #12548]: #11710 +#3926 := (not #2855) +#3927 := (or #3398 #3926) +#4263 := [def-axiom]: #3927 +#12577 := [unit-resolution #4263 #12965]: #3926 +#13397 := (not #11710) +#12612 := (or #13407 #2855 #11469 #13397) +#11605 := (uf_1 uf_16 ?x49!11) +#11606 := (uf_10 #11605) +#11631 := (+ #2853 #11606) +#11632 := (+ #144 #11631) +#12233 := (<= #11632 0::int) +#11635 := (= #11632 0::int) +#11610 := (* -1::int #11606) +#11611 := (+ uf_9 #11610) +#11612 := (<= #11611 0::int) +#12253 := (not #11612) +#11624 := (+ #11589 #11606) +#11625 := (+ #144 #11624) +#11626 := (>= #11625 0::int) +#11669 := (or #11612 #11626) +#11674 := (not #11669) +#11663 := (= #2251 #11572) +#13437 := (not #11663) +#11590 := (+ #2251 #11589) +#12252 := (>= #11590 0::int) +#13367 := (not #12252) +#13284 := [hypothesis]: #11710 +#13478 := [hypothesis]: #3926 +#13215 := [hypothesis]: #12814 +#13389 := (or #13367 #13397 #2855 #13407 #13409) +#13410 := [th-lemma]: #13389 +#13436 := [unit-resolution #13410 #13215 #13478 #13284 #12934]: #13367 +#13434 := (or #13437 #12252) +#12573 := [th-lemma]: #13434 +#13419 := [unit-resolution #12573 #13436]: #13437 +#11677 := (or #11663 #11674) +#12241 := (or #4461 #11663 #11674) +#11613 := (+ #1449 #11610) +#11614 := (+ #11572 #11613) +#11615 := (<= #11614 0::int) +#11659 := (or #11615 #11612) +#11660 := (not #11659) +#11661 := (= #11572 #2251) +#11662 := (or #11661 #11660) +#12242 := (or #4461 #11662) +#12249 := (iff #12242 #12241) +#12245 := (or #4461 #11677) +#12247 := (iff #12245 #12241) +#12248 := [rewrite]: #12247 +#12239 := (iff #12242 #12245) +#11678 := (iff #11662 #11677) +#11675 := (iff #11660 #11674) +#11672 := (iff #11659 #11669) +#11666 := (or #11626 #11612) +#11670 := (iff #11666 #11669) +#11671 := [rewrite]: #11670 +#11667 := (iff #11659 #11666) +#11629 := (iff #11615 #11626) +#11617 := (+ #11572 #11610) +#11618 := (+ #1449 #11617) +#11621 := (<= #11618 0::int) +#11627 := (iff #11621 #11626) +#11628 := [rewrite]: #11627 +#11622 := (iff #11615 #11621) +#11619 := (= #11614 #11618) +#11620 := [rewrite]: #11619 +#11623 := [monotonicity #11620]: #11622 +#11630 := [trans #11623 #11628]: #11629 +#11668 := [monotonicity #11630]: #11667 +#11673 := [trans #11668 #11671]: #11672 +#11676 := [monotonicity #11673]: #11675 +#11664 := (iff #11661 #11663) +#11665 := [rewrite]: #11664 +#11679 := [monotonicity #11665 #11676]: #11678 +#12246 := [monotonicity #11679]: #12239 +#12244 := [trans #12246 #12248]: #12249 +#12243 := [quant-inst]: #12242 +#12251 := [mp #12243 #12244]: #12241 +#13417 := [unit-resolution #12251 #8574]: #11677 +#13423 := [unit-resolution #13417 #13419]: #11674 +#12254 := (or #11669 #12253) +#12258 := [def-axiom]: #12254 +#13426 := [unit-resolution #12258 #13423]: #12253 +#12250 := (not #11626) +#12259 := (or #11669 #12250) +#12257 := [def-axiom]: #12259 +#13412 := [unit-resolution #12257 #13423]: #12250 +#11641 := (or #11612 #11626 #11635) +#12229 := (or #4469 #11612 #11626 #11635) +#11607 := (+ #11606 #2853) +#11608 := (+ #144 #11607) +#11609 := (= #11608 0::int) +#11616 := (or #11615 #11612 #11609) +#12222 := (or #4469 #11616) +#12236 := (iff #12222 #12229) +#12231 := (or #4469 #11641) +#12227 := (iff #12231 #12229) +#12235 := [rewrite]: #12227 +#12232 := (iff #12222 #12231) +#11644 := (iff #11616 #11641) +#11638 := (or #11626 #11612 #11635) +#11642 := (iff #11638 #11641) +#11643 := [rewrite]: #11642 +#11639 := (iff #11616 #11638) +#11636 := (iff #11609 #11635) +#11633 := (= #11608 #11632) +#11634 := [rewrite]: #11633 +#11637 := [monotonicity #11634]: #11636 +#11640 := [monotonicity #11630 #11637]: #11639 +#11645 := [trans #11640 #11643]: #11644 +#12234 := [monotonicity #11645]: #12232 +#12237 := [trans #12234 #12235]: #12236 +#12230 := [quant-inst]: #12222 +#12238 := [mp #12230 #12237]: #12229 +#13413 := [unit-resolution #12238 #7390]: #11641 +#13433 := [unit-resolution #13413 #13412 #13426]: #11635 +#13370 := (not #11635) +#13390 := (or #13370 #12233) +#13391 := [th-lemma]: #13390 +#13385 := [unit-resolution #13391 #13433]: #12233 +#12787 := (uf_1 #7128 ?x49!11) +#12788 := (uf_10 #12787) +#12790 := (* -1::int #12788) +#13285 := (+ #11606 #12790) +#13280 := (>= #13285 0::int) +#13216 := (= #11606 #12788) +#12644 := (= #12788 #11606) +#12645 := (= #12787 #11605) +#13418 := [monotonicity #8036]: #12645 +#12646 := [monotonicity #13418]: #12644 +#12578 := [symm #12646]: #13216 +#12641 := (not #13216) +#12647 := (or #12641 #13280) +#12643 := [th-lemma]: #12647 +#12582 := [unit-resolution #12643 #12578]: #13280 +#13068 := (<= #12788 0::int) +#13063 := (not #13068) +#12581 := [hypothesis]: #12039 +#13211 := (or #7093 #11469 #13063) +#12936 := (= #7128 ?x49!11) +#13162 := (or #12936 #13063) +#13263 := (or #7093 #13162) +#13354 := (iff #13263 #13211) +#13255 := (or #11469 #13063) +#13161 := (or #7093 #13255) +#13351 := (iff #13161 #13211) +#13352 := [rewrite]: #13351 +#13071 := (iff #13263 #13161) +#13069 := (iff #13162 #13255) +#13204 := (iff #12936 #11469) +#13265 := [rewrite]: #13204 +#13210 := [monotonicity #13265]: #13069 +#13163 := [monotonicity #13210]: #13071 +#13067 := [trans #13163 #13352]: #13354 +#13160 := [quant-inst]: #13263 +#13355 := [mp #13160 #13067]: #13211 +#12609 := [unit-resolution #13355 #4347 #12581]: #13063 +#12610 := [th-lemma #13478 #13215 #12934 #12609 #12582 #13385]: false +#12583 := [lemma #12610]: #12612 +#12780 := [unit-resolution #12583 #12577 #12546 #12576]: #13407 +#11201 := (+ #2249 #11200) +#11202 := (<= #11201 0::int) +#12087 := (or #4477 #11202) +#11190 := (+ #11183 #2250) +#11193 := (>= #11190 0::int) +#12088 := (or #4477 #11193) +#12090 := (iff #12088 #12087) +#12092 := (iff #12087 #12087) +#12093 := [rewrite]: #12092 +#11205 := (iff #11193 #11202) +#11194 := (+ #2250 #11183) +#11197 := (>= #11194 0::int) +#11203 := (iff #11197 #11202) +#11204 := [rewrite]: #11203 +#11198 := (iff #11193 #11197) +#11195 := (= #11190 #11194) +#11196 := [rewrite]: #11195 +#11199 := [monotonicity #11196]: #11198 +#11206 := [trans #11199 #11204]: #11205 +#12091 := [monotonicity #11206]: #12090 +#12095 := [trans #12091 #12093]: #12090 +#12085 := [quant-inst]: #12088 +#12097 := [mp #12085 #12095]: #12087 +#13218 := [unit-resolution #12097 #14110]: #11202 +#12617 := (not #11202) +#12729 := (not #7572) +#12728 := (or #13764 #12729 #12814 #12617 #12642) +#12733 := [th-lemma]: #12728 +#12616 := [unit-resolution #12733 #13218 #9185 #14101 #12780]: #13764 +#13844 := (or #10236 #13668) +#13842 := [hypothesis]: #13764 +#13843 := [hypothesis]: #10848 +#13801 := (or #4426 #7247 #10236 #13668) +#13669 := (or #10236 #7247 #13668) +#13802 := (or #4426 #13669) +#13788 := (iff #13802 #13801) +#13670 := (or #7247 #10236 #13668) +#13804 := (or #4426 #13670) +#13786 := (iff #13804 #13801) +#13787 := [rewrite]: #13786 +#13805 := (iff #13802 #13804) +#13665 := (iff #13669 #13670) +#13671 := [rewrite]: #13665 +#13806 := [monotonicity #13671]: #13805 +#13789 := [trans #13806 #13787]: #13788 +#13803 := [quant-inst]: #13802 +#13790 := [mp #13803 #13789]: #13801 +#13838 := [unit-resolution #13790 #9153 #9152 #13843 #13842]: false +#13845 := [lemma #13838]: #13844 +#12734 := [unit-resolution #13845 #12616]: #10236 +#11030 := (not #10857) +#11307 := (or #11030 #10849 #10848) +#11294 := [def-axiom]: #11307 +#12742 := [unit-resolution #11294 #12734 #13217]: #10849 +#12732 := [trans #12742 #8036]: #12998 +#12743 := [monotonicity #12732]: #12812 +#12969 := (not #12812) +#12967 := (or #12969 #12814) +#12973 := [th-lemma]: #12967 +#12786 := [unit-resolution #12973 #12780]: #12969 +#12771 := [unit-resolution #12786 #12743]: false +#12735 := [lemma #12771]: #3398 +#4251 := (or #4544 #3403 #4538) +#4248 := [def-axiom]: #4251 +#25338 := [unit-resolution #4248 #12735]: #25343 +#29778 := [unit-resolution #25338 #29777]: #4538 +#3968 := (or #4535 #4529) +#3969 := [def-axiom]: #3968 +#29779 := [unit-resolution #3969 #29778]: #4529 +#25346 := (or #4532 #4526) +#17148 := [hypothesis]: #3449 +#4266 := (or #3444 #2287) +#4267 := [def-axiom]: #4266 +#17149 := [unit-resolution #4267 #17148]: #2287 +#9931 := (uf_1 uf_16 ?x50!14) +#9932 := (uf_10 #9931) +#9936 := (* -1::int #9932) +#17081 := (+ #2281 #9936) +#17083 := (>= #17081 0::int) +#17080 := (= #2281 #9932) +#17180 := (= #2280 #9931) +#17179 := (= ?x51!13 uf_16) +#10336 := (= ?x51!13 #7128) +#10334 := (uf_6 uf_15 ?x51!13) +#10335 := (= uf_8 #10334) +#10367 := (not #10335) +#10193 := (uf_4 uf_14 ?x51!13) +#9874 := (uf_4 uf_14 ?x50!14) +#9915 := (* -1::int #9874) +#10384 := (+ #9915 #10193) +#10385 := (+ #2281 #10384) +#10388 := (>= #10385 0::int) +#17156 := (not #10388) +#9916 := (+ #2276 #9915) +#9917 := (<= #9916 0::int) +#16708 := (or #4477 #9917) +#9907 := (+ #9874 #2277) +#9908 := (>= #9907 0::int) +#16709 := (or #4477 #9908) +#16711 := (iff #16709 #16708) +#16713 := (iff #16708 #16708) +#16714 := [rewrite]: #16713 +#9920 := (iff #9908 #9917) +#9909 := (+ #2277 #9874) +#9912 := (>= #9909 0::int) +#9918 := (iff #9912 #9917) +#9919 := [rewrite]: #9918 +#9913 := (iff #9908 #9912) +#9910 := (= #9907 #9909) +#9911 := [rewrite]: #9910 +#9914 := [monotonicity #9911]: #9913 +#9921 := [trans #9914 #9919]: #9920 +#16712 := [monotonicity #9921]: #16711 +#16715 := [trans #16712 #16714]: #16711 +#16710 := [quant-inst]: #16709 +#16716 := [mp #16710 #16715]: #16708 +#17308 := [unit-resolution #16716 #14110]: #9917 +#3906 := (not #2882) +#4269 := (or #3444 #3906) +#4271 := [def-axiom]: #4269 +#17151 := [unit-resolution #4271 #17148]: #3906 +#10228 := (* -1::int #10193) +#10229 := (+ #2278 #10228) +#15878 := (>= #10229 0::int) +#10198 := (= #2278 #10193) +#4262 := (or #3444 #2289) +#4268 := [def-axiom]: #4262 +#17152 := [unit-resolution #4268 #17148]: #2289 +#16493 := (or #4486 #3429 #10198) +#10194 := (= #10193 #2278) +#10197 := (or #10194 #3429) +#16494 := (or #4486 #10197) +#16503 := (iff #16494 #16493) +#10204 := (or #3429 #10198) +#16498 := (or #4486 #10204) +#16501 := (iff #16498 #16493) +#16502 := [rewrite]: #16501 +#16499 := (iff #16494 #16498) +#10207 := (iff #10197 #10204) +#10201 := (or #10198 #3429) +#10205 := (iff #10201 #10204) +#10206 := [rewrite]: #10205 +#10202 := (iff #10197 #10201) +#10199 := (iff #10194 #10198) +#10200 := [rewrite]: #10199 +#10203 := [monotonicity #10200]: #10202 +#10208 := [trans #10203 #10206]: #10207 +#16500 := [monotonicity #10208]: #16499 +#16504 := [trans #16500 #16502]: #16503 +#16497 := [quant-inst]: #16494 +#16505 := [mp #16497 #16504]: #16493 +#17150 := [unit-resolution #16505 #9423 #17152]: #10198 +#17153 := (not #10198) +#17154 := (or #17153 #15878) +#17155 := [th-lemma]: #17154 +#17147 := [unit-resolution #17155 #17150]: #15878 +#17313 := (not #9917) +#17315 := (not #15878) +#17157 := (or #17156 #17315 #17313 #2882) +#17158 := [th-lemma]: #17157 +#17159 := [unit-resolution #17158 #17147 #17151 #17308]: #17156 +#17146 := (or #10367 #10388) +#16749 := (or #4417 #2286 #10367 #10388) +#10380 := (+ #10193 #9915) +#10381 := (+ #2281 #10380) +#10382 := (>= #10381 0::int) +#10383 := (or #10367 #2286 #10382) +#16750 := (or #4417 #10383) +#16757 := (iff #16750 #16749) +#10394 := (or #2286 #10367 #10388) +#16752 := (or #4417 #10394) +#16755 := (iff #16752 #16749) +#16756 := [rewrite]: #16755 +#16753 := (iff #16750 #16752) +#10397 := (iff #10383 #10394) +#10391 := (or #10367 #2286 #10388) +#10395 := (iff #10391 #10394) +#10396 := [rewrite]: #10395 +#10392 := (iff #10383 #10391) +#10389 := (iff #10382 #10388) +#10386 := (= #10381 #10385) +#10387 := [rewrite]: #10386 +#10390 := [monotonicity #10387]: #10389 +#10393 := [monotonicity #10390]: #10392 +#10398 := [trans #10393 #10396]: #10397 +#16754 := [monotonicity #10398]: #16753 +#16758 := [trans #16754 #16756]: #16757 +#16751 := [quant-inst]: #16750 +#16759 := [mp #16751 #16758]: #16749 +#17160 := [unit-resolution #16759 #8027 #17149]: #17146 +#17161 := [unit-resolution #17160 #17159]: #10367 +#10344 := (ite #10336 #5314 #10335) +#10338 := (uf_6 #7203 ?x51!13) +#10341 := (= uf_8 #10338) +#10347 := (iff #10341 #10344) +#16506 := (or #7026 #10347) +#10337 := (ite #10336 #6089 #10335) +#10339 := (= #10338 uf_8) +#10340 := (iff #10339 #10337) +#16507 := (or #7026 #10340) +#16509 := (iff #16507 #16506) +#16511 := (iff #16506 #16506) +#16512 := [rewrite]: #16511 +#10348 := (iff #10340 #10347) +#10345 := (iff #10337 #10344) +#10346 := [monotonicity #6102]: #10345 +#10342 := (iff #10339 #10341) +#10343 := [rewrite]: #10342 +#10349 := [monotonicity #10343 #10346]: #10348 +#16510 := [monotonicity #10349]: #16509 +#16513 := [trans #16510 #16512]: #16509 +#16508 := [quant-inst]: #16507 +#16514 := [mp #16508 #16513]: #16506 +#17162 := [unit-resolution #16514 #4320]: #10347 +#17171 := (= #2288 #10338) +#17163 := (= #10338 #2288) +#17164 := [monotonicity #8591]: #17163 +#17174 := [symm #17164]: #17171 +#17175 := [trans #17152 #17174]: #10341 +#16528 := (not #10341) +#16525 := (not #10347) +#16529 := (or #16525 #16528 #10344) +#16530 := [def-axiom]: #16529 +#17176 := [unit-resolution #16530 #17175 #17162]: #10344 +#16515 := (not #10344) +#16519 := (or #16515 #10336 #10335) +#16520 := [def-axiom]: #16519 +#17178 := [unit-resolution #16520 #17176 #17161]: #10336 +#17177 := [trans #17178 #8036]: #17179 +#17181 := [monotonicity #17177]: #17180 +#17182 := [monotonicity #17181]: #17080 +#17187 := (not #17080) +#17188 := (or #17187 #17083) +#17186 := [th-lemma]: #17188 +#17189 := [unit-resolution #17186 #17182]: #17083 +#9937 := (+ uf_9 #9936) +#9938 := (<= #9937 0::int) +#9950 := (+ #9915 #9932) +#9951 := (+ #144 #9950) +#9952 := (>= #9951 0::int) +#16744 := (not #9952) +#10475 := (uf_2 #2280) +#11002 := (uf_4 uf_14 #10475) +#11016 := (* -1::int #11002) +#16798 := (+ #10193 #11016) +#16800 := (>= #16798 0::int) +#16797 := (= #10193 #11002) +#10476 := (= ?x51!13 #10475) +#16793 := (or #7136 #10476) +#16794 := [quant-inst]: #16793 +#17292 := [unit-resolution #16794 #4306]: #10476 +#17295 := [monotonicity #17292]: #16797 +#17296 := (not #16797) +#17297 := (or #17296 #16800) +#17298 := [th-lemma]: #17297 +#17299 := [unit-resolution #17298 #17295]: #16800 +#11017 := (+ #144 #11016) +#11018 := (<= #11017 0::int) +#17065 := (= #144 #11002) +#17202 := (= #11002 #144) +#17194 := (= #10475 uf_16) +#17192 := (= #10475 #7128) +#17190 := (= #10475 ?x51!13) +#17191 := [symm #17292]: #17190 +#17193 := [trans #17191 #17178]: #17192 +#17195 := [trans #17193 #8036]: #17194 +#17203 := [monotonicity #17195]: #17202 +#17204 := [symm #17203]: #17065 +#17205 := (not #17065) +#17206 := (or #17205 #11018) +#17201 := [th-lemma]: #17206 +#17207 := [unit-resolution #17201 #17204]: #11018 +#17316 := (not #11018) +#17314 := (not #16800) +#17208 := (not #17083) +#17209 := (or #16744 #17208 #17313 #2882 #17314 #17315 #17316) +#17210 := [th-lemma]: #17209 +#17211 := [unit-resolution #17210 #17207 #17308 #17147 #17151 #17299 #17189]: #16744 +#9957 := (+ #2277 #9932) +#9958 := (+ #144 #9957) +#9961 := (= #9958 0::int) +#17223 := (not #9961) +#16729 := (>= #9958 0::int) +#17219 := (not #16729) +#17220 := (or #17219 #17208 #2882 #17314 #17315 #17316) +#17221 := [th-lemma]: #17220 +#17222 := [unit-resolution #17221 #17207 #17147 #17151 #17299 #17189]: #17219 +#17218 := (or #17223 #16729) +#17224 := [th-lemma]: #17218 +#17225 := [unit-resolution #17224 #17222]: #17223 +#9967 := (or #9938 #9952 #9961) +#16717 := (or #4469 #9938 #9952 #9961) +#9933 := (+ #9932 #2277) +#9934 := (+ #144 #9933) +#9935 := (= #9934 0::int) +#9939 := (+ #1449 #9936) +#9940 := (+ #9874 #9939) +#9941 := (<= #9940 0::int) +#9942 := (or #9941 #9938 #9935) +#16718 := (or #4469 #9942) +#16725 := (iff #16718 #16717) +#16720 := (or #4469 #9967) +#16723 := (iff #16720 #16717) +#16724 := [rewrite]: #16723 +#16721 := (iff #16718 #16720) +#9970 := (iff #9942 #9967) +#9964 := (or #9952 #9938 #9961) +#9968 := (iff #9964 #9967) +#9969 := [rewrite]: #9968 +#9965 := (iff #9942 #9964) +#9962 := (iff #9935 #9961) +#9959 := (= #9934 #9958) +#9960 := [rewrite]: #9959 +#9963 := [monotonicity #9960]: #9962 +#9955 := (iff #9941 #9952) +#9943 := (+ #9874 #9936) +#9944 := (+ #1449 #9943) +#9947 := (<= #9944 0::int) +#9953 := (iff #9947 #9952) +#9954 := [rewrite]: #9953 +#9948 := (iff #9941 #9947) +#9945 := (= #9940 #9944) +#9946 := [rewrite]: #9945 +#9949 := [monotonicity #9946]: #9948 +#9956 := [trans #9949 #9954]: #9955 +#9966 := [monotonicity #9956 #9963]: #9965 +#9971 := [trans #9966 #9969]: #9970 +#16722 := [monotonicity #9971]: #16721 +#16726 := [trans #16722 #16724]: #16725 +#16719 := [quant-inst]: #16718 +#16727 := [mp #16719 #16726]: #16717 +#17226 := [unit-resolution #16727 #7390]: #9967 +#17227 := [unit-resolution #17226 #17225 #17211]: #9938 +#17228 := [th-lemma #17227 #17189 #17149]: false +#17253 := [lemma #17228]: #3444 +#4253 := (or #4532 #3449 #4526) +#4257 := [def-axiom]: #4253 +#25347 := [unit-resolution #4257 #17253]: #25346 +#29780 := [unit-resolution #25347 #29779]: #4526 +#3983 := (or #4523 #4515) +#3984 := [def-axiom]: #3983 +#29781 := [unit-resolution #3984 #29780]: #4515 +#20547 := (or #4520 #7166 #15518 #15538) +#15505 := (+ #2307 #15504) +#15506 := (+ #7167 #15505) +#15507 := (= #15506 0::int) +#15508 := (not #15507) +#15509 := (+ #7167 #2307) +#15510 := (>= #15509 0::int) +#15511 := (or #7166 #15510 #15508) +#17626 := (or #4520 #15511) +#20578 := (iff #17626 #20547) +#20549 := (or #4520 #15541) +#20516 := (iff #20549 #20547) +#20517 := [rewrite]: #20516 +#20386 := (iff #17626 #20549) +#15542 := (iff #15511 #15541) +#15539 := (iff #15508 #15538) +#15536 := (iff #15507 #15533) +#15523 := (+ #7167 #15504) +#15524 := (+ #2307 #15523) +#15527 := (= #15524 0::int) +#15534 := (iff #15527 #15533) +#15535 := [rewrite]: #15534 +#15528 := (iff #15507 #15527) +#15525 := (= #15506 #15524) +#15526 := [rewrite]: #15525 +#15529 := [monotonicity #15526]: #15528 +#15537 := [trans #15529 #15535]: #15536 +#15540 := [monotonicity #15537]: #15539 +#15521 := (iff #15510 #15518) +#15512 := (+ #2307 #7167) +#15515 := (>= #15512 0::int) +#15519 := (iff #15515 #15518) +#15520 := [rewrite]: #15519 +#15516 := (iff #15510 #15515) +#15513 := (= #15509 #15512) +#15514 := [rewrite]: #15513 +#15517 := [monotonicity #15514]: #15516 +#15522 := [trans #15517 #15520]: #15521 +#15543 := [monotonicity #15522 #15540]: #15542 +#20388 := [monotonicity #15543]: #20386 +#20485 := [trans #20388 #20517]: #20578 +#20612 := [quant-inst]: #17626 +#20795 := [mp #20612 #20485]: #20547 +#29782 := [unit-resolution #20795 #29781]: #15541 +#29783 := [unit-resolution #29782 #29774 #25367]: #15518 +#28771 := (<= #28711 0::int) +#28772 := (not #28771) +#28773 := (= #7128 #27533) +#29792 := (not #28773) +#28056 := (not #18779) +#29793 := (iff #28056 #29792) +#29790 := (iff #18779 #28773) +#29788 := (iff #28773 #18779) +#29785 := (iff #28773 #25219) +#29786 := [monotonicity #29764]: #29785 +#29789 := [trans #29786 #29787]: #29788 +#29791 := [symm #29789]: #29790 +#29794 := [monotonicity #29791]: #29793 +#29784 := [hypothesis]: #28056 +#29795 := [mp #29784 #29794]: #29792 +#28775 := (or #28772 #28773) +#29743 := (or #7093 #28772 #28773) +#28774 := (or #28773 #28772) +#29744 := (or #7093 #28774) +#29751 := (iff #29744 #29743) +#29746 := (or #7093 #28775) +#29749 := (iff #29746 #29743) +#29750 := [rewrite]: #29749 +#29747 := (iff #29744 #29746) +#28776 := (iff #28774 #28775) +#28777 := [rewrite]: #28776 +#29748 := [monotonicity #28777]: #29747 +#29752 := [trans #29748 #29750]: #29751 +#29745 := [quant-inst]: #29744 +#29753 := [mp #29745 #29752]: #29743 +#29796 := [unit-resolution #29753 #4347]: #28775 +#29797 := [unit-resolution #29796 #29795]: #28772 +#29798 := [th-lemma #29797 #29783 #25304 #25235 #12934 #29773]: false +#29800 := [lemma #29798]: #29799 +#32973 := [unit-resolution #29800 #32972]: #18779 +#32992 := [mp #32973 #32991]: #32602 +#32632 := (not #32602) +#32993 := (or #32591 #32632) +#29702 := (or #32591 #32632 #7040) +#30315 := [def-axiom]: #29702 +#32995 := [unit-resolution #30315 #8579]: #32993 +#32996 := [unit-resolution #32995 #32992]: #32591 +#32599 := (not #32591) +#28920 := (not #32564) +#30916 := (or #28920 #32538 #32599) +#31002 := [def-axiom]: #30916 +#32987 := [unit-resolution #31002 #32996 #32981 #32954]: false +#32997 := [lemma #32987]: #15367 +#39375 := (or #32956 #14218) +#39376 := [th-lemma]: #39375 +#39377 := [unit-resolution #39376 #32997]: #14218 +#34351 := (not #14218) +#34357 := (or #34324 #34351) +#4272 := (or #4523 #2319) +#4270 := [def-axiom]: #4272 +#34292 := [unit-resolution #4270 #29780]: #2319 +#34293 := [hypothesis]: #14218 +#34291 := [hypothesis]: #16010 +#34294 := [th-lemma #34291 #34293 #34292]: false +#34641 := [lemma #34294]: #34357 +#39378 := [unit-resolution #34641 #39377]: #34324 +#39380 := (or #16010 #16030) +#4273 := (or #4523 #2896) +#4259 := [def-axiom]: #4273 +#39379 := [unit-resolution #4259 #29780]: #2896 +#17785 := (or #4442 #2893 #16010 #16030) +#15998 := (+ #15997 #15995) +#15999 := (+ #15362 #15998) +#16000 := (= #15999 0::int) +#16001 := (not #16000) +#16007 := (or #16006 #16003 #16001) +#16008 := (not #16007) +#16011 := (or #2320 #16010 #16008) +#20897 := (or #4442 #16011) +#21453 := (iff #20897 #17785) +#16033 := (or #2893 #16010 #16030) +#21128 := (or #4442 #16033) +#18084 := (iff #21128 #17785) +#22003 := [rewrite]: #18084 +#20788 := (iff #20897 #21128) +#16034 := (iff #16011 #16033) +#16031 := (iff #16008 #16030) +#16028 := (iff #16007 #16025) +#16022 := (or #16006 #16003 #16019) +#16026 := (iff #16022 #16025) +#16027 := [rewrite]: #16026 +#16023 := (iff #16007 #16022) +#16020 := (iff #16001 #16019) +#16017 := (iff #16000 #16016) +#16014 := (= #15999 #16013) +#16015 := [rewrite]: #16014 +#16018 := [monotonicity #16015]: #16017 +#16021 := [monotonicity #16018]: #16020 +#16024 := [monotonicity #16021]: #16023 +#16029 := [trans #16024 #16027]: #16028 +#16032 := [monotonicity #16029]: #16031 +#16035 := [monotonicity #2895 #16032]: #16034 +#21237 := [monotonicity #16035]: #20788 +#21330 := [trans #21237 #22003]: #21453 +#21234 := [quant-inst]: #20897 +#21506 := [mp #21234 #21330]: #17785 +#39381 := [unit-resolution #21506 #8030 #39379]: #39380 +#39382 := [unit-resolution #39381 #39378]: #16030 +#22030 := (or #16025 #16016) +#18058 := [def-axiom]: #22030 +#32365 := [unit-resolution #18058 #39382]: #16016 +#32368 := (or #16019 #22006) +#29568 := [th-lemma]: #32368 +#29612 := [unit-resolution #29568 #32365]: #22006 +#20411 := (+ uf_9 #15995) +#20412 := (<= #20411 0::int) +#20215 := (uf_6 uf_17 #15992) +#20216 := (= uf_8 #20215) +#20606 := (uf_2 #15993) +#39327 := (uf_6 #7203 #20606) +#38901 := (= #39327 #20215) +#38905 := (= #20215 #39327) +#20607 := (= #15992 #20606) +#25402 := (or #7136 #20607) +#25393 := [quant-inst]: #25402 +#39384 := [unit-resolution #25393 #4306]: #20607 +#8247 := (= uf_17 #7203) +#8291 := (= #150 #7203) +#8151 := [symm #8593]: #8291 +#8423 := [trans #8578 #8151]: #8247 +#38906 := [monotonicity #8423 #39384]: #38905 +#38907 := [symm #38906]: #38901 +#39330 := (= uf_8 #39327) +#21345 := (uf_6 uf_15 #20606) +#21346 := (= uf_8 #21345) +#39333 := (= #7128 #20606) +#39336 := (ite #39333 #5314 #21346) +#39339 := (iff #39330 #39336) +#38867 := (or #7026 #39339) +#39325 := (= #20606 #7128) +#39326 := (ite #39325 #6089 #21346) +#39328 := (= #39327 uf_8) +#39329 := (iff #39328 #39326) +#38877 := (or #7026 #39329) +#38879 := (iff #38877 #38867) +#38890 := (iff #38867 #38867) +#38888 := [rewrite]: #38890 +#39340 := (iff #39329 #39339) +#39337 := (iff #39326 #39336) +#39334 := (iff #39325 #39333) +#39335 := [rewrite]: #39334 +#39338 := [monotonicity #39335 #6102]: #39337 +#39331 := (iff #39328 #39330) +#39332 := [rewrite]: #39331 +#39341 := [monotonicity #39332 #39338]: #39340 +#38889 := [monotonicity #39341]: #38879 +#38894 := [trans #38889 #38888]: #38879 +#38878 := [quant-inst]: #38877 +#38893 := [mp #38878 #38894]: #38867 +#38919 := [unit-resolution #38893 #4320]: #39339 +#38895 := (not #39339) +#38922 := (or #38895 #39330) +#39351 := (not #39336) +#39371 := [hypothesis]: #39351 +#39352 := (not #39333) +#39372 := (or #39336 #39352) +#39357 := (or #39336 #39352 #7040) +#39358 := [def-axiom]: #39357 +#39373 := [unit-resolution #39358 #8579]: #39372 +#39374 := [unit-resolution #39373 #39371]: #39352 +#39394 := (or #39336 #39333) +#39391 := (= #16004 #21345) +#39387 := (= #21345 #16004) +#39385 := (= #20606 #15992) +#39386 := [symm #39384]: #39385 +#39388 := [monotonicity #39386]: #39387 +#39392 := [symm #39388]: #39391 +#21698 := (or #16025 #16005) +#21997 := [def-axiom]: #21698 +#39383 := [unit-resolution #21997 #39382]: #16005 +#39393 := [trans #39383 #39392]: #21346 +#21347 := (not #21346) +#39359 := (or #39336 #39333 #21347) +#39360 := [def-axiom]: #39359 +#39395 := [unit-resolution #39360 #39393]: #39394 +#39396 := [unit-resolution #39395 #39374 #39371]: false +#39397 := [lemma #39396]: #39336 +#38896 := (or #38895 #39330 #39351) +#38897 := [def-axiom]: #38896 +#38902 := [unit-resolution #38897 #39397]: #38922 +#38903 := [unit-resolution #38902 #38919]: #39330 +#38908 := [trans #38903 #38907]: #20216 +#20217 := (not #20216) +#38921 := [hypothesis]: #20217 +#38924 := [unit-resolution #38921 #38908]: false +#38927 := [lemma #38924]: #20216 +#20218 := (uf_18 #15992) +#20235 := (* -1::int #20218) +#20421 := (+ #15995 #20235) +#20422 := (+ #2306 #20421) +#20423 := (<= #20422 0::int) +#31764 := (not #20423) +#26473 := (>= #20422 0::int) +#20236 := (+ #15996 #20235) +#20237 := (>= #20236 0::int) +#25927 := (or #4477 #20237) +#26030 := [quant-inst]: #25927 +#27294 := [unit-resolution #26030 #14110]: #20237 +#32338 := (not #20237) +#29920 := (not #22006) +#29919 := (or #26473 #29920 #34351 #32338) +#32347 := [th-lemma]: #29919 +#32336 := [unit-resolution #32347 #29612 #27294 #39377]: #26473 +#20465 := (= #20422 0::int) +#20470 := (not #20465) +#20454 := (+ #2306 #20235) +#20455 := (<= #20454 0::int) +#32331 := (not #20455) +#22027 := (not #16003) +#21587 := (or #16025 #22027) +#21112 := [def-axiom]: #21587 +#32344 := [unit-resolution #21112 #39382]: #22027 +#32343 := (or #32331 #16003 #34351 #32338) +#32335 := [th-lemma]: #32343 +#32375 := [unit-resolution #32335 #32344 #27294 #39377]: #32331 +#20473 := (or #20217 #20455 #20470) +#25436 := (or #4520 #20217 #20455 #20470) +#20442 := (+ #2307 #15994) +#20443 := (+ #20218 #20442) +#20444 := (= #20443 0::int) +#20445 := (not #20444) +#20406 := (+ #20218 #2307) +#20446 := (>= #20406 0::int) +#20447 := (or #20217 #20446 #20445) +#26505 := (or #4520 #20447) +#26705 := (iff #26505 #25436) +#11899 := (or #4520 #20473) +#26697 := (iff #11899 #25436) +#26704 := [rewrite]: #26697 +#25435 := (iff #26505 #11899) +#20474 := (iff #20447 #20473) +#20471 := (iff #20445 #20470) +#20468 := (iff #20444 #20465) +#20414 := (+ #15994 #20218) +#20415 := (+ #2307 #20414) +#20462 := (= #20415 0::int) +#20466 := (iff #20462 #20465) +#20467 := [rewrite]: #20466 +#20463 := (iff #20444 #20462) +#20460 := (= #20443 #20415) +#20461 := [rewrite]: #20460 +#20464 := [monotonicity #20461]: #20463 +#20469 := [trans #20464 #20467]: #20468 +#20472 := [monotonicity #20469]: #20471 +#20458 := (iff #20446 #20455) +#20448 := (+ #2307 #20218) +#20451 := (>= #20448 0::int) +#20456 := (iff #20451 #20455) +#20457 := [rewrite]: #20456 +#20452 := (iff #20446 #20451) +#20449 := (= #20406 #20448) +#20450 := [rewrite]: #20449 +#20453 := [monotonicity #20450]: #20452 +#20459 := [trans #20453 #20457]: #20458 +#20475 := [monotonicity #20459 #20472]: #20474 +#26696 := [monotonicity #20475]: #25435 +#26698 := [trans #26696 #26704]: #26705 +#26504 := [quant-inst]: #26505 +#26615 := [mp #26504 #26698]: #25436 +#29644 := [unit-resolution #26615 #29781]: #20473 +#29611 := [unit-resolution #29644 #38927 #32375]: #20470 +#32349 := (not #26473) +#31762 := (or #20465 #31764 #32349) +#29679 := [th-lemma]: #31762 +#32366 := [unit-resolution #29679 #29611 #32336]: #31764 +#20428 := (or #20217 #20412 #20423) +#4260 := (or #4523 #4506) +#3982 := [def-axiom]: #4260 +#29894 := [unit-resolution #3982 #29780]: #4506 +#26338 := (or #4511 #20217 #20412 #20423) +#20407 := (+ #15994 #20406) +#20410 := (>= #20407 0::int) +#20413 := (or #20217 #20412 #20410) +#26340 := (or #4511 #20413) +#25469 := (iff #26340 #26338) +#10685 := (or #4511 #20428) +#25478 := (iff #10685 #26338) +#25474 := [rewrite]: #25478 +#25437 := (iff #26340 #10685) +#20429 := (iff #20413 #20428) +#20426 := (iff #20410 #20423) +#20418 := (>= #20415 0::int) +#20424 := (iff #20418 #20423) +#20425 := [rewrite]: #20424 +#20419 := (iff #20410 #20418) +#20416 := (= #20407 #20415) +#20417 := [rewrite]: #20416 +#20420 := [monotonicity #20417]: #20419 +#20427 := [trans #20420 #20425]: #20426 +#20430 := [monotonicity #20427]: #20429 +#25388 := [monotonicity #20430]: #25437 +#25466 := [trans #25388 #25474]: #25469 +#25477 := [quant-inst]: #26340 +#25432 := [mp #25477 #25466]: #26338 +#31826 := [unit-resolution #25432 #29894]: #20428 +#29616 := [unit-resolution #31826 #32366 #38927]: #20412 +[th-lemma #34292 #39377 #29616 #29612 #29563]: false +unsat diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/Boogie/Examples/cert/Boogie_b_max --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Boogie/Examples/cert/Boogie_b_max Wed Nov 04 11:40:59 2009 +0100 @@ -0,0 +1,19 @@ +(benchmark Isabelle +:extrafuns ( + (uf_3 Int Int) + (uf_5 Int) + (uf_7 Int) + (uf_11 Int) + (uf_4 Int) + (uf_9 Int) + (uf_13 Int) + (uf_1 Int) + (uf_2 Int) + (uf_6 Int) + (uf_10 Int) + (uf_8 Int) + (uf_12 Int) + ) +:assumption (not (implies true (implies (< 0 uf_1) (implies true (implies (= uf_2 (uf_3 0)) (implies (and (<= 1 1) (and (<= 1 1) (and (<= 0 0) (<= 0 0)))) (and (implies (forall (?x1 Int) (implies (and (< ?x1 1) (<= 0 ?x1)) (<= (uf_3 ?x1) uf_2))) (and (implies (= (uf_3 0) uf_2) (implies true (implies (and (<= 1 uf_4) (<= 0 uf_5)) (implies (forall (?x2 Int) (implies (and (< ?x2 uf_4) (<= 0 ?x2)) (<= (uf_3 ?x2) uf_6))) (implies (= (uf_3 uf_5) uf_6) (implies (and (<= 1 uf_4) (<= 0 uf_5)) (and (implies true (implies (and (<= 1 uf_4) (<= 0 uf_5)) (implies (< uf_4 uf_1) (implies (and (<= 1 uf_4) (<= 0 uf_5)) (and (implies true (implies (and (<= 1 uf_4) (<= 0 uf_5)) (implies (<= (uf_3 uf_4) uf_6) (implies (and (<= 1 uf_4) (<= 0 uf_5)) (implies true (implies (= uf_7 uf_5) (implies (= uf_8 uf_6) (implies true (implies (and (<= 1 uf_4) (<= 0 uf_7)) (implies (= uf_9 (+ uf_4 1)) (implies (and (<= 2 uf_9) (<= 0 uf_7)) (implies true (and (implies (forall (?x3 Int) (implies (and (< ?x3 uf_9) (<= 0 ?x3)) (<= (uf_3 ?x3) uf_8))) (and (implies (= (uf_3 uf_7) uf_8) (implies false true)) (= (uf_3 uf_7) uf_8))) (forall (?x4 Int) (implies (and (< ?x4 uf_9) (<= 0 ?x4)) (<= (uf_3 ?x4) uf_8)))))))))))))))) (implies true (implies (and (<= 1 uf_4) (<= 0 uf_5)) (implies (< uf_6 (uf_3 uf_4)) (implies (= uf_10 (uf_3 uf_4)) (implies (and (<= 1 uf_4) (<= 1 uf_4)) (implies true (implies (= uf_7 uf_4) (implies (= uf_8 uf_10) (implies true (implies (and (<= 1 uf_4) (<= 0 uf_7)) (implies (= uf_9 (+ uf_4 1)) (implies (and (<= 2 uf_9) (<= 0 uf_7)) (implies true (and (implies (forall (?x5 Int) (implies (and (< ?x5 uf_9) (<= 0 ?x5)) (<= (uf_3 ?x5) uf_8))) (and (implies (= (uf_3 uf_7) uf_8) (implies false true)) (= (uf_3 uf_7) uf_8))) (forall (?x6 Int) (implies (and (< ?x6 uf_9) (<= 0 ?x6)) (<= (uf_3 ?x6) uf_8)))))))))))))))))))))) (implies true (implies (and (<= 1 uf_4) (<= 0 uf_5)) (implies (<= uf_1 uf_4) (implies (and (<= 1 uf_4) (<= 0 uf_5)) (implies true (implies (= uf_11 uf_5) (implies (= uf_12 uf_6) (implies (= uf_13 uf_4) (implies true (and (implies (exists (?x7 Int) (implies (and (< ?x7 uf_1) (<= 0 ?x7)) (= (uf_3 ?x7) uf_12))) (and (implies (forall (?x8 Int) (implies (and (< ?x8 uf_1) (<= 0 ?x8)) (<= (uf_3 ?x8) uf_12))) true) (forall (?x9 Int) (implies (and (< ?x9 uf_1) (<= 0 ?x9)) (<= (uf_3 ?x9) uf_12))))) (exists (?x10 Int) (implies (and (< ?x10 uf_1) (<= 0 ?x10)) (= (uf_3 ?x10) uf_12)))))))))))))))))))) (= (uf_3 0) uf_2))) (forall (?x11 Int) (implies (and (< ?x11 1) (<= 0 ?x11)) (<= (uf_3 ?x11) uf_2)))))))))) +:formula true +) diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/Boogie/Examples/cert/Boogie_b_max.proof --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Boogie/Examples/cert/Boogie_b_max.proof Wed Nov 04 11:40:59 2009 +0100 @@ -0,0 +1,2329 @@ +#2 := false +#4 := 0::int +decl uf_3 :: (-> int int) +decl ?x3!1 :: int +#1188 := ?x3!1 +#1195 := (uf_3 ?x3!1) +#760 := -1::int +#1381 := (* -1::int #1195) +decl uf_4 :: int +#25 := uf_4 +#39 := (uf_3 uf_4) +#2763 := (+ #39 #1381) +#2765 := (>= #2763 0::int) +#2762 := (= #39 #1195) +#2631 := (= uf_4 ?x3!1) +#1394 := (* -1::int ?x3!1) +#2564 := (+ uf_4 #1394) +#2628 := (>= #2564 0::int) +decl uf_9 :: int +#47 := uf_9 +#806 := (* -1::int uf_9) +#838 := (+ uf_4 #806) +#1977 := (>= #838 -1::int) +#837 := (= #838 -1::int) +#1395 := (+ uf_9 #1394) +#1396 := (<= #1395 0::int) +decl uf_8 :: int +#43 := uf_8 +#1382 := (+ uf_8 #1381) +#1383 := (>= #1382 0::int) +#1192 := (>= ?x3!1 0::int) +#1625 := (not #1192) +#1640 := (or #1625 #1383 #1396) +#1645 := (not #1640) +#16 := (:var 0 int) +#20 := (uf_3 #16) +#2303 := (pattern #20) +#807 := (+ #16 #806) +#805 := (>= #807 0::int) +#799 := (* -1::int uf_8) +#800 := (+ #20 #799) +#801 := (<= #800 0::int) +#753 := (>= #16 0::int) +#1548 := (not #753) +#1607 := (or #1548 #801 #805) +#2320 := (forall (vars (?x3 int)) (:pat #2303) #1607) +#2325 := (not #2320) +decl uf_7 :: int +#41 := uf_7 +#58 := (uf_3 uf_7) +#221 := (= uf_8 #58) +#2328 := (or #221 #2325) +#2331 := (not #2328) +#2334 := (or #2331 #1645) +#2337 := (not #2334) +#851 := (* -1::int #39) +decl uf_6 :: int +#32 := uf_6 +#852 := (+ uf_6 #851) +#850 := (>= #852 0::int) +#840 := (not #837) +#50 := 2::int +#791 := (>= uf_9 2::int) +#1656 := (not #791) +#788 := (>= uf_7 0::int) +#1655 := (not #788) +decl uf_5 :: int +#27 := uf_5 +#779 := (>= uf_5 0::int) +#1654 := (not #779) +#10 := 1::int +#776 := (>= uf_4 1::int) +#886 := (not #776) +#361 := (= uf_4 uf_7) +#376 := (not #361) +decl uf_10 :: int +#78 := uf_10 +#356 := (= #39 uf_10) +#401 := (not #356) +#82 := (= uf_8 uf_10) +#367 := (not #82) +#2346 := (or #367 #401 #376 #886 #1654 #1655 #1656 #840 #850 #2337) +#2349 := (not #2346) +#854 := (not #850) +#194 := (= uf_6 uf_8) +#301 := (not #194) +#191 := (= uf_5 uf_7) +#310 := (not #191) +#2340 := (or #310 #301 #886 #1654 #1655 #1656 #840 #854 #2337) +#2343 := (not #2340) +#2352 := (or #2343 #2349) +#2355 := (not #2352) +#924 := (* -1::int uf_4) +decl uf_1 :: int +#5 := uf_1 +#925 := (+ uf_1 #924) +#926 := (<= #925 0::int) +#2358 := (or #886 #1654 #926 #2355) +#2361 := (not #2358) +decl ?x7!2 :: int +#1279 := ?x7!2 +#1287 := (uf_3 ?x7!2) +decl uf_12 :: int +#99 := uf_12 +#1469 := (= uf_12 #1287) +#1284 := (>= ?x7!2 0::int) +#1711 := (not #1284) +#1280 := (* -1::int ?x7!2) +#1281 := (+ uf_1 #1280) +#1282 := (<= #1281 0::int) +#1726 := (or #1282 #1711 #1469) +#1757 := (not #1726) +decl ?x8!3 :: int +#1297 := ?x8!3 +#1298 := (uf_3 ?x8!3) +#1493 := (* -1::int #1298) +#1494 := (+ uf_12 #1493) +#1495 := (>= #1494 0::int) +#1305 := (>= ?x8!3 0::int) +#1731 := (not #1305) +#1301 := (* -1::int ?x8!3) +#1302 := (+ uf_1 #1301) +#1303 := (<= #1302 0::int) +#1795 := (or #1303 #1731 #1495 #1757) +#1798 := (not #1795) +#951 := (* -1::int #16) +#952 := (+ uf_1 #951) +#953 := (<= #952 0::int) +#105 := (= #20 uf_12) +#1700 := (or #105 #1548 #953) +#1705 := (not #1700) +#2364 := (forall (vars (?x7 int)) (:pat #2303) #1705) +#2369 := (or #2364 #1798) +#2372 := (not #2369) +#927 := (not #926) +decl uf_13 :: int +#101 := uf_13 +#472 := (= uf_4 uf_13) +#542 := (not #472) +#469 := (= uf_6 uf_12) +#551 := (not #469) +decl uf_11 :: int +#97 := uf_11 +#466 := (= uf_5 uf_11) +#560 := (not #466) +#2375 := (or #560 #551 #542 #886 #1654 #927 #2372) +#2378 := (not #2375) +#2381 := (or #2361 #2378) +#2384 := (not #2381) +#1030 := (+ #16 #924) +#1029 := (>= #1030 0::int) +#1024 := (* -1::int uf_6) +#1025 := (+ #20 #1024) +#1026 := (<= #1025 0::int) +#1585 := (or #1548 #1026 #1029) +#2312 := (forall (vars (?x2 int)) (:pat #2303) #1585) +#2317 := (not #2312) +#763 := (* -1::int #20) +decl uf_2 :: int +#7 := uf_2 +#764 := (+ uf_2 #763) +#762 := (>= #764 0::int) +#749 := (>= #16 1::int) +#1563 := (or #749 #1548 #762) +#2304 := (forall (vars (?x1 int)) (:pat #2303) #1563) +#2309 := (not #2304) +#36 := (uf_3 uf_5) +#188 := (= uf_6 #36) +#619 := (not #188) +#2387 := (or #619 #886 #1654 #2309 #2317 #2384) +#2390 := (not #2387) +decl ?x1!0 :: int +#1152 := ?x1!0 +#1156 := (>= ?x1!0 1::int) +#1155 := (>= ?x1!0 0::int) +#1164 := (not #1155) +#1153 := (uf_3 ?x1!0) +#1150 := (* -1::int #1153) +#1151 := (+ uf_2 #1150) +#1154 := (>= #1151 0::int) +#1540 := (or #1154 #1164 #1156) +#2202 := (= uf_2 #1153) +#8 := (uf_3 0::int) +#2191 := (= #8 #1153) +#2188 := (= #1153 #8) +#2207 := (= ?x1!0 0::int) +#1157 := (not #1156) +#1545 := (not #1540) +#2205 := [hypothesis]: #1545 +#1976 := (or #1540 #1157) +#1967 := [def-axiom]: #1976 +#2206 := [unit-resolution #1967 #2205]: #1157 +#1975 := (or #1540 #1155) +#1890 := [def-axiom]: #1975 +#2203 := [unit-resolution #1890 #2205]: #1155 +#2187 := [th-lemma #2203 #2206]: #2207 +#2190 := [monotonicity #2187]: #2188 +#2192 := [symm #2190]: #2191 +#9 := (= uf_2 #8) +#1078 := (<= uf_1 0::int) +#1031 := (not #1029) +#1034 := (and #753 #1031) +#1037 := (not #1034) +#1040 := (or #1026 #1037) +#1043 := (forall (vars (?x2 int)) #1040) +#1046 := (not #1043) +#972 := (* -1::int uf_12) +#973 := (+ #20 #972) +#974 := (<= #973 0::int) +#954 := (not #953) +#957 := (and #753 #954) +#960 := (not #957) +#980 := (or #960 #974) +#985 := (forall (vars (?x8 int)) #980) +#963 := (or #105 #960) +#966 := (exists (vars (?x7 int)) #963) +#969 := (not #966) +#988 := (or #969 #985) +#991 := (and #966 #988) +#781 := (and #776 #779) +#784 := (not #781) +#1016 := (or #560 #551 #542 #784 #927 #991) +#843 := (and #776 #788) +#846 := (not #843) +#804 := (not #805) +#810 := (and #753 #804) +#813 := (not #810) +#816 := (or #801 #813) +#819 := (forall (vars (?x3 int)) #816) +#822 := (not #819) +#828 := (or #221 #822) +#833 := (and #819 #828) +#793 := (and #788 #791) +#796 := (not #793) +#916 := (or #367 #401 #376 #886 #784 #796 #833 #840 #846 #850) +#881 := (or #310 #301 #784 #796 #833 #840 #846 #854) +#921 := (and #881 #916) +#946 := (or #784 #921 #926) +#1021 := (and #946 #1016) +#652 := (not #9) +#1064 := (or #652 #619 #784 #1021 #1046) +#1069 := (and #9 #1064) +#747 := (not #749) +#754 := (and #747 #753) +#757 := (not #754) +#766 := (or #757 #762) +#769 := (forall (vars (?x1 int)) #766) +#772 := (not #769) +#1072 := (or #772 #1069) +#1075 := (and #769 #1072) +#1098 := (or #652 #1075 #1078) +#1103 := (not #1098) +#21 := (<= #20 uf_2) +#18 := (<= 0::int #16) +#17 := (< #16 1::int) +#19 := (and #17 #18) +#22 := (implies #19 #21) +#23 := (forall (vars (?x1 int)) #22) +#24 := (= #8 uf_2) +#103 := (< #16 uf_1) +#104 := (and #103 #18) +#106 := (implies #104 #105) +#107 := (exists (vars (?x7 int)) #106) +#108 := (<= #20 uf_12) +#109 := (implies #104 #108) +#110 := (forall (vars (?x8 int)) #109) +#1 := true +#111 := (implies #110 true) +#112 := (and #111 #110) +#113 := (implies #107 #112) +#114 := (and #113 #107) +#115 := (implies true #114) +#102 := (= uf_13 uf_4) +#116 := (implies #102 #115) +#100 := (= uf_12 uf_6) +#117 := (implies #100 #116) +#98 := (= uf_11 uf_5) +#118 := (implies #98 #117) +#119 := (implies true #118) +#28 := (<= 0::int uf_5) +#26 := (<= 1::int uf_4) +#29 := (and #26 #28) +#120 := (implies #29 #119) +#96 := (<= uf_1 uf_4) +#121 := (implies #96 #120) +#122 := (implies #29 #121) +#123 := (implies true #122) +#55 := (<= #20 uf_8) +#53 := (< #16 uf_9) +#54 := (and #53 #18) +#56 := (implies #54 #55) +#57 := (forall (vars (?x3 int)) #56) +#59 := (= #58 uf_8) +#60 := (implies false true) +#61 := (implies #59 #60) +#62 := (and #61 #59) +#63 := (implies #57 #62) +#64 := (and #63 #57) +#65 := (implies true #64) +#45 := (<= 0::int uf_7) +#51 := (<= 2::int uf_9) +#52 := (and #51 #45) +#66 := (implies #52 #65) +#48 := (+ uf_4 1::int) +#49 := (= uf_9 #48) +#67 := (implies #49 #66) +#46 := (and #26 #45) +#68 := (implies #46 #67) +#69 := (implies true #68) +#83 := (implies #82 #69) +#81 := (= uf_7 uf_4) +#84 := (implies #81 #83) +#85 := (implies true #84) +#80 := (and #26 #26) +#86 := (implies #80 #85) +#79 := (= uf_10 #39) +#87 := (implies #79 #86) +#77 := (< uf_6 #39) +#88 := (implies #77 #87) +#89 := (implies #29 #88) +#90 := (implies true #89) +#44 := (= uf_8 uf_6) +#70 := (implies #44 #69) +#42 := (= uf_7 uf_5) +#71 := (implies #42 #70) +#72 := (implies true #71) +#73 := (implies #29 #72) +#40 := (<= #39 uf_6) +#74 := (implies #40 #73) +#75 := (implies #29 #74) +#76 := (implies true #75) +#91 := (and #76 #90) +#92 := (implies #29 #91) +#38 := (< uf_4 uf_1) +#93 := (implies #38 #92) +#94 := (implies #29 #93) +#95 := (implies true #94) +#124 := (and #95 #123) +#125 := (implies #29 #124) +#37 := (= #36 uf_6) +#126 := (implies #37 #125) +#33 := (<= #20 uf_6) +#30 := (< #16 uf_4) +#31 := (and #30 #18) +#34 := (implies #31 #33) +#35 := (forall (vars (?x2 int)) #34) +#127 := (implies #35 #126) +#128 := (implies #29 #127) +#129 := (implies true #128) +#130 := (implies #24 #129) +#131 := (and #130 #24) +#132 := (implies #23 #131) +#133 := (and #132 #23) +#12 := (<= 0::int 0::int) +#13 := (and #12 #12) +#11 := (<= 1::int 1::int) +#14 := (and #11 #13) +#15 := (and #11 #14) +#134 := (implies #15 #133) +#135 := (implies #9 #134) +#136 := (implies true #135) +#6 := (< 0::int uf_1) +#137 := (implies #6 #136) +#138 := (implies true #137) +#139 := (not #138) +#1106 := (iff #139 #1103) +#475 := (and #18 #103) +#481 := (not #475) +#493 := (or #108 #481) +#498 := (forall (vars (?x8 int)) #493) +#482 := (or #105 #481) +#487 := (exists (vars (?x7 int)) #482) +#518 := (not #487) +#519 := (or #518 #498) +#527 := (and #487 #519) +#543 := (or #542 #527) +#552 := (or #551 #543) +#561 := (or #560 #552) +#326 := (not #29) +#576 := (or #326 #561) +#584 := (not #96) +#585 := (or #584 #576) +#593 := (or #326 #585) +#206 := (and #18 #53) +#212 := (not #206) +#213 := (or #55 #212) +#218 := (forall (vars (?x3 int)) #213) +#243 := (not #218) +#244 := (or #243 #221) +#252 := (and #218 #244) +#203 := (and #45 #51) +#267 := (not #203) +#268 := (or #267 #252) +#197 := (+ 1::int uf_4) +#200 := (= uf_9 #197) +#276 := (not #200) +#277 := (or #276 #268) +#285 := (not #46) +#286 := (or #285 #277) +#368 := (or #367 #286) +#377 := (or #376 #368) +#392 := (not #26) +#393 := (or #392 #377) +#402 := (or #401 #393) +#410 := (not #77) +#411 := (or #410 #402) +#419 := (or #326 #411) +#302 := (or #301 #286) +#311 := (or #310 #302) +#327 := (or #326 #311) +#335 := (not #40) +#336 := (or #335 #327) +#344 := (or #326 #336) +#431 := (and #344 #419) +#437 := (or #326 #431) +#445 := (not #38) +#446 := (or #445 #437) +#454 := (or #326 #446) +#605 := (and #454 #593) +#611 := (or #326 #605) +#620 := (or #619 #611) +#173 := (and #18 #30) +#179 := (not #173) +#180 := (or #33 #179) +#185 := (forall (vars (?x2 int)) #180) +#628 := (not #185) +#629 := (or #628 #620) +#637 := (or #326 #629) +#653 := (or #652 #637) +#661 := (and #9 #653) +#164 := (not #19) +#165 := (or #164 #21) +#168 := (forall (vars (?x1 int)) #165) +#669 := (not #168) +#670 := (or #669 #661) +#678 := (and #168 #670) +#158 := (and #11 #12) +#161 := (and #11 #158) +#686 := (not #161) +#687 := (or #686 #678) +#695 := (or #652 #687) +#710 := (not #6) +#711 := (or #710 #695) +#723 := (not #711) +#1104 := (iff #723 #1103) +#1101 := (iff #711 #1098) +#1089 := (or false #1075) +#1092 := (or #652 #1089) +#1095 := (or #1078 #1092) +#1099 := (iff #1095 #1098) +#1100 := [rewrite]: #1099 +#1096 := (iff #711 #1095) +#1093 := (iff #695 #1092) +#1090 := (iff #687 #1089) +#1076 := (iff #678 #1075) +#1073 := (iff #670 #1072) +#1070 := (iff #661 #1069) +#1067 := (iff #653 #1064) +#1049 := (or #784 #1021) +#1052 := (or #619 #1049) +#1055 := (or #1046 #1052) +#1058 := (or #784 #1055) +#1061 := (or #652 #1058) +#1065 := (iff #1061 #1064) +#1066 := [rewrite]: #1065 +#1062 := (iff #653 #1061) +#1059 := (iff #637 #1058) +#1056 := (iff #629 #1055) +#1053 := (iff #620 #1052) +#1050 := (iff #611 #1049) +#1022 := (iff #605 #1021) +#1019 := (iff #593 #1016) +#998 := (or #542 #991) +#1001 := (or #551 #998) +#1004 := (or #560 #1001) +#1007 := (or #784 #1004) +#1010 := (or #927 #1007) +#1013 := (or #784 #1010) +#1017 := (iff #1013 #1016) +#1018 := [rewrite]: #1017 +#1014 := (iff #593 #1013) +#1011 := (iff #585 #1010) +#1008 := (iff #576 #1007) +#1005 := (iff #561 #1004) +#1002 := (iff #552 #1001) +#999 := (iff #543 #998) +#992 := (iff #527 #991) +#989 := (iff #519 #988) +#986 := (iff #498 #985) +#983 := (iff #493 #980) +#977 := (or #974 #960) +#981 := (iff #977 #980) +#982 := [rewrite]: #981 +#978 := (iff #493 #977) +#961 := (iff #481 #960) +#958 := (iff #475 #957) +#955 := (iff #103 #954) +#956 := [rewrite]: #955 +#751 := (iff #18 #753) +#752 := [rewrite]: #751 +#959 := [monotonicity #752 #956]: #958 +#962 := [monotonicity #959]: #961 +#975 := (iff #108 #974) +#976 := [rewrite]: #975 +#979 := [monotonicity #976 #962]: #978 +#984 := [trans #979 #982]: #983 +#987 := [quant-intro #984]: #986 +#970 := (iff #518 #969) +#967 := (iff #487 #966) +#964 := (iff #482 #963) +#965 := [monotonicity #962]: #964 +#968 := [quant-intro #965]: #967 +#971 := [monotonicity #968]: #970 +#990 := [monotonicity #971 #987]: #989 +#993 := [monotonicity #968 #990]: #992 +#1000 := [monotonicity #993]: #999 +#1003 := [monotonicity #1000]: #1002 +#1006 := [monotonicity #1003]: #1005 +#785 := (iff #326 #784) +#782 := (iff #29 #781) +#778 := (iff #28 #779) +#780 := [rewrite]: #778 +#775 := (iff #26 #776) +#777 := [rewrite]: #775 +#783 := [monotonicity #777 #780]: #782 +#786 := [monotonicity #783]: #785 +#1009 := [monotonicity #786 #1006]: #1008 +#996 := (iff #584 #927) +#994 := (iff #96 #926) +#995 := [rewrite]: #994 +#997 := [monotonicity #995]: #996 +#1012 := [monotonicity #997 #1009]: #1011 +#1015 := [monotonicity #786 #1012]: #1014 +#1020 := [trans #1015 #1018]: #1019 +#949 := (iff #454 #946) +#937 := (or #784 #921) +#940 := (or #926 #937) +#943 := (or #784 #940) +#947 := (iff #943 #946) +#948 := [rewrite]: #947 +#944 := (iff #454 #943) +#941 := (iff #446 #940) +#938 := (iff #437 #937) +#922 := (iff #431 #921) +#919 := (iff #419 #916) +#857 := (or #796 #833) +#860 := (or #840 #857) +#863 := (or #846 #860) +#898 := (or #367 #863) +#901 := (or #376 #898) +#904 := (or #886 #901) +#907 := (or #401 #904) +#910 := (or #850 #907) +#913 := (or #784 #910) +#917 := (iff #913 #916) +#918 := [rewrite]: #917 +#914 := (iff #419 #913) +#911 := (iff #411 #910) +#908 := (iff #402 #907) +#905 := (iff #393 #904) +#902 := (iff #377 #901) +#899 := (iff #368 #898) +#864 := (iff #286 #863) +#861 := (iff #277 #860) +#858 := (iff #268 #857) +#834 := (iff #252 #833) +#831 := (iff #244 #828) +#825 := (or #822 #221) +#829 := (iff #825 #828) +#830 := [rewrite]: #829 +#826 := (iff #244 #825) +#823 := (iff #243 #822) +#820 := (iff #218 #819) +#817 := (iff #213 #816) +#814 := (iff #212 #813) +#811 := (iff #206 #810) +#808 := (iff #53 #804) +#809 := [rewrite]: #808 +#812 := [monotonicity #752 #809]: #811 +#815 := [monotonicity #812]: #814 +#802 := (iff #55 #801) +#803 := [rewrite]: #802 +#818 := [monotonicity #803 #815]: #817 +#821 := [quant-intro #818]: #820 +#824 := [monotonicity #821]: #823 +#827 := [monotonicity #824]: #826 +#832 := [trans #827 #830]: #831 +#835 := [monotonicity #821 #832]: #834 +#797 := (iff #267 #796) +#794 := (iff #203 #793) +#790 := (iff #51 #791) +#792 := [rewrite]: #790 +#787 := (iff #45 #788) +#789 := [rewrite]: #787 +#795 := [monotonicity #789 #792]: #794 +#798 := [monotonicity #795]: #797 +#859 := [monotonicity #798 #835]: #858 +#841 := (iff #276 #840) +#836 := (iff #200 #837) +#839 := [rewrite]: #836 +#842 := [monotonicity #839]: #841 +#862 := [monotonicity #842 #859]: #861 +#847 := (iff #285 #846) +#844 := (iff #46 #843) +#845 := [monotonicity #777 #789]: #844 +#848 := [monotonicity #845]: #847 +#865 := [monotonicity #848 #862]: #864 +#900 := [monotonicity #865]: #899 +#903 := [monotonicity #900]: #902 +#887 := (iff #392 #886) +#888 := [monotonicity #777]: #887 +#906 := [monotonicity #888 #903]: #905 +#909 := [monotonicity #906]: #908 +#896 := (iff #410 #850) +#891 := (not #854) +#894 := (iff #891 #850) +#895 := [rewrite]: #894 +#892 := (iff #410 #891) +#889 := (iff #77 #854) +#890 := [rewrite]: #889 +#893 := [monotonicity #890]: #892 +#897 := [trans #893 #895]: #896 +#912 := [monotonicity #897 #909]: #911 +#915 := [monotonicity #786 #912]: #914 +#920 := [trans #915 #918]: #919 +#884 := (iff #344 #881) +#866 := (or #301 #863) +#869 := (or #310 #866) +#872 := (or #784 #869) +#875 := (or #854 #872) +#878 := (or #784 #875) +#882 := (iff #878 #881) +#883 := [rewrite]: #882 +#879 := (iff #344 #878) +#876 := (iff #336 #875) +#873 := (iff #327 #872) +#870 := (iff #311 #869) +#867 := (iff #302 #866) +#868 := [monotonicity #865]: #867 +#871 := [monotonicity #868]: #870 +#874 := [monotonicity #786 #871]: #873 +#855 := (iff #335 #854) +#849 := (iff #40 #850) +#853 := [rewrite]: #849 +#856 := [monotonicity #853]: #855 +#877 := [monotonicity #856 #874]: #876 +#880 := [monotonicity #786 #877]: #879 +#885 := [trans #880 #883]: #884 +#923 := [monotonicity #885 #920]: #922 +#939 := [monotonicity #786 #923]: #938 +#935 := (iff #445 #926) +#930 := (not #927) +#933 := (iff #930 #926) +#934 := [rewrite]: #933 +#931 := (iff #445 #930) +#928 := (iff #38 #927) +#929 := [rewrite]: #928 +#932 := [monotonicity #929]: #931 +#936 := [trans #932 #934]: #935 +#942 := [monotonicity #936 #939]: #941 +#945 := [monotonicity #786 #942]: #944 +#950 := [trans #945 #948]: #949 +#1023 := [monotonicity #950 #1020]: #1022 +#1051 := [monotonicity #786 #1023]: #1050 +#1054 := [monotonicity #1051]: #1053 +#1047 := (iff #628 #1046) +#1044 := (iff #185 #1043) +#1041 := (iff #180 #1040) +#1038 := (iff #179 #1037) +#1035 := (iff #173 #1034) +#1032 := (iff #30 #1031) +#1033 := [rewrite]: #1032 +#1036 := [monotonicity #752 #1033]: #1035 +#1039 := [monotonicity #1036]: #1038 +#1027 := (iff #33 #1026) +#1028 := [rewrite]: #1027 +#1042 := [monotonicity #1028 #1039]: #1041 +#1045 := [quant-intro #1042]: #1044 +#1048 := [monotonicity #1045]: #1047 +#1057 := [monotonicity #1048 #1054]: #1056 +#1060 := [monotonicity #786 #1057]: #1059 +#1063 := [monotonicity #1060]: #1062 +#1068 := [trans #1063 #1066]: #1067 +#1071 := [monotonicity #1068]: #1070 +#773 := (iff #669 #772) +#770 := (iff #168 #769) +#767 := (iff #165 #766) +#761 := (iff #21 #762) +#765 := [rewrite]: #761 +#758 := (iff #164 #757) +#755 := (iff #19 #754) +#748 := (iff #17 #747) +#750 := [rewrite]: #748 +#756 := [monotonicity #750 #752]: #755 +#759 := [monotonicity #756]: #758 +#768 := [monotonicity #759 #765]: #767 +#771 := [quant-intro #768]: #770 +#774 := [monotonicity #771]: #773 +#1074 := [monotonicity #774 #1071]: #1073 +#1077 := [monotonicity #771 #1074]: #1076 +#745 := (iff #686 false) +#740 := (not true) +#743 := (iff #740 false) +#744 := [rewrite]: #743 +#741 := (iff #686 #740) +#738 := (iff #161 true) +#730 := (and true true) +#733 := (and true #730) +#736 := (iff #733 true) +#737 := [rewrite]: #736 +#734 := (iff #161 #733) +#731 := (iff #158 #730) +#728 := (iff #12 true) +#729 := [rewrite]: #728 +#726 := (iff #11 true) +#727 := [rewrite]: #726 +#732 := [monotonicity #727 #729]: #731 +#735 := [monotonicity #727 #732]: #734 +#739 := [trans #735 #737]: #738 +#742 := [monotonicity #739]: #741 +#746 := [trans #742 #744]: #745 +#1091 := [monotonicity #746 #1077]: #1090 +#1094 := [monotonicity #1091]: #1093 +#1087 := (iff #710 #1078) +#1079 := (not #1078) +#1082 := (not #1079) +#1085 := (iff #1082 #1078) +#1086 := [rewrite]: #1085 +#1083 := (iff #710 #1082) +#1080 := (iff #6 #1079) +#1081 := [rewrite]: #1080 +#1084 := [monotonicity #1081]: #1083 +#1088 := [trans #1084 #1086]: #1087 +#1097 := [monotonicity #1088 #1094]: #1096 +#1102 := [trans #1097 #1100]: #1101 +#1105 := [monotonicity #1102]: #1104 +#724 := (iff #139 #723) +#721 := (iff #138 #711) +#716 := (implies true #711) +#719 := (iff #716 #711) +#720 := [rewrite]: #719 +#717 := (iff #138 #716) +#714 := (iff #137 #711) +#707 := (implies #6 #695) +#712 := (iff #707 #711) +#713 := [rewrite]: #712 +#708 := (iff #137 #707) +#705 := (iff #136 #695) +#700 := (implies true #695) +#703 := (iff #700 #695) +#704 := [rewrite]: #703 +#701 := (iff #136 #700) +#698 := (iff #135 #695) +#692 := (implies #9 #687) +#696 := (iff #692 #695) +#697 := [rewrite]: #696 +#693 := (iff #135 #692) +#690 := (iff #134 #687) +#683 := (implies #161 #678) +#688 := (iff #683 #687) +#689 := [rewrite]: #688 +#684 := (iff #134 #683) +#681 := (iff #133 #678) +#675 := (and #670 #168) +#679 := (iff #675 #678) +#680 := [rewrite]: #679 +#676 := (iff #133 #675) +#169 := (iff #23 #168) +#166 := (iff #22 #165) +#167 := [rewrite]: #166 +#170 := [quant-intro #167]: #169 +#673 := (iff #132 #670) +#666 := (implies #168 #661) +#671 := (iff #666 #670) +#672 := [rewrite]: #671 +#667 := (iff #132 #666) +#664 := (iff #131 #661) +#658 := (and #653 #9) +#662 := (iff #658 #661) +#663 := [rewrite]: #662 +#659 := (iff #131 #658) +#171 := (iff #24 #9) +#172 := [rewrite]: #171 +#656 := (iff #130 #653) +#649 := (implies #9 #637) +#654 := (iff #649 #653) +#655 := [rewrite]: #654 +#650 := (iff #130 #649) +#647 := (iff #129 #637) +#642 := (implies true #637) +#645 := (iff #642 #637) +#646 := [rewrite]: #645 +#643 := (iff #129 #642) +#640 := (iff #128 #637) +#634 := (implies #29 #629) +#638 := (iff #634 #637) +#639 := [rewrite]: #638 +#635 := (iff #128 #634) +#632 := (iff #127 #629) +#625 := (implies #185 #620) +#630 := (iff #625 #629) +#631 := [rewrite]: #630 +#626 := (iff #127 #625) +#623 := (iff #126 #620) +#616 := (implies #188 #611) +#621 := (iff #616 #620) +#622 := [rewrite]: #621 +#617 := (iff #126 #616) +#614 := (iff #125 #611) +#608 := (implies #29 #605) +#612 := (iff #608 #611) +#613 := [rewrite]: #612 +#609 := (iff #125 #608) +#606 := (iff #124 #605) +#603 := (iff #123 #593) +#598 := (implies true #593) +#601 := (iff #598 #593) +#602 := [rewrite]: #601 +#599 := (iff #123 #598) +#596 := (iff #122 #593) +#590 := (implies #29 #585) +#594 := (iff #590 #593) +#595 := [rewrite]: #594 +#591 := (iff #122 #590) +#588 := (iff #121 #585) +#581 := (implies #96 #576) +#586 := (iff #581 #585) +#587 := [rewrite]: #586 +#582 := (iff #121 #581) +#579 := (iff #120 #576) +#573 := (implies #29 #561) +#577 := (iff #573 #576) +#578 := [rewrite]: #577 +#574 := (iff #120 #573) +#571 := (iff #119 #561) +#566 := (implies true #561) +#569 := (iff #566 #561) +#570 := [rewrite]: #569 +#567 := (iff #119 #566) +#564 := (iff #118 #561) +#557 := (implies #466 #552) +#562 := (iff #557 #561) +#563 := [rewrite]: #562 +#558 := (iff #118 #557) +#555 := (iff #117 #552) +#548 := (implies #469 #543) +#553 := (iff #548 #552) +#554 := [rewrite]: #553 +#549 := (iff #117 #548) +#546 := (iff #116 #543) +#539 := (implies #472 #527) +#544 := (iff #539 #543) +#545 := [rewrite]: #544 +#540 := (iff #116 #539) +#537 := (iff #115 #527) +#532 := (implies true #527) +#535 := (iff #532 #527) +#536 := [rewrite]: #535 +#533 := (iff #115 #532) +#530 := (iff #114 #527) +#524 := (and #519 #487) +#528 := (iff #524 #527) +#529 := [rewrite]: #528 +#525 := (iff #114 #524) +#488 := (iff #107 #487) +#485 := (iff #106 #482) +#478 := (implies #475 #105) +#483 := (iff #478 #482) +#484 := [rewrite]: #483 +#479 := (iff #106 #478) +#476 := (iff #104 #475) +#477 := [rewrite]: #476 +#480 := [monotonicity #477]: #479 +#486 := [trans #480 #484]: #485 +#489 := [quant-intro #486]: #488 +#522 := (iff #113 #519) +#515 := (implies #487 #498) +#520 := (iff #515 #519) +#521 := [rewrite]: #520 +#516 := (iff #113 #515) +#513 := (iff #112 #498) +#508 := (and true #498) +#511 := (iff #508 #498) +#512 := [rewrite]: #511 +#509 := (iff #112 #508) +#499 := (iff #110 #498) +#496 := (iff #109 #493) +#490 := (implies #475 #108) +#494 := (iff #490 #493) +#495 := [rewrite]: #494 +#491 := (iff #109 #490) +#492 := [monotonicity #477]: #491 +#497 := [trans #492 #495]: #496 +#500 := [quant-intro #497]: #499 +#506 := (iff #111 true) +#501 := (implies #498 true) +#504 := (iff #501 true) +#505 := [rewrite]: #504 +#502 := (iff #111 #501) +#503 := [monotonicity #500]: #502 +#507 := [trans #503 #505]: #506 +#510 := [monotonicity #507 #500]: #509 +#514 := [trans #510 #512]: #513 +#517 := [monotonicity #489 #514]: #516 +#523 := [trans #517 #521]: #522 +#526 := [monotonicity #523 #489]: #525 +#531 := [trans #526 #529]: #530 +#534 := [monotonicity #531]: #533 +#538 := [trans #534 #536]: #537 +#473 := (iff #102 #472) +#474 := [rewrite]: #473 +#541 := [monotonicity #474 #538]: #540 +#547 := [trans #541 #545]: #546 +#470 := (iff #100 #469) +#471 := [rewrite]: #470 +#550 := [monotonicity #471 #547]: #549 +#556 := [trans #550 #554]: #555 +#467 := (iff #98 #466) +#468 := [rewrite]: #467 +#559 := [monotonicity #468 #556]: #558 +#565 := [trans #559 #563]: #564 +#568 := [monotonicity #565]: #567 +#572 := [trans #568 #570]: #571 +#575 := [monotonicity #572]: #574 +#580 := [trans #575 #578]: #579 +#583 := [monotonicity #580]: #582 +#589 := [trans #583 #587]: #588 +#592 := [monotonicity #589]: #591 +#597 := [trans #592 #595]: #596 +#600 := [monotonicity #597]: #599 +#604 := [trans #600 #602]: #603 +#464 := (iff #95 #454) +#459 := (implies true #454) +#462 := (iff #459 #454) +#463 := [rewrite]: #462 +#460 := (iff #95 #459) +#457 := (iff #94 #454) +#451 := (implies #29 #446) +#455 := (iff #451 #454) +#456 := [rewrite]: #455 +#452 := (iff #94 #451) +#449 := (iff #93 #446) +#442 := (implies #38 #437) +#447 := (iff #442 #446) +#448 := [rewrite]: #447 +#443 := (iff #93 #442) +#440 := (iff #92 #437) +#434 := (implies #29 #431) +#438 := (iff #434 #437) +#439 := [rewrite]: #438 +#435 := (iff #92 #434) +#432 := (iff #91 #431) +#429 := (iff #90 #419) +#424 := (implies true #419) +#427 := (iff #424 #419) +#428 := [rewrite]: #427 +#425 := (iff #90 #424) +#422 := (iff #89 #419) +#416 := (implies #29 #411) +#420 := (iff #416 #419) +#421 := [rewrite]: #420 +#417 := (iff #89 #416) +#414 := (iff #88 #411) +#407 := (implies #77 #402) +#412 := (iff #407 #411) +#413 := [rewrite]: #412 +#408 := (iff #88 #407) +#405 := (iff #87 #402) +#398 := (implies #356 #393) +#403 := (iff #398 #402) +#404 := [rewrite]: #403 +#399 := (iff #87 #398) +#396 := (iff #86 #393) +#389 := (implies #26 #377) +#394 := (iff #389 #393) +#395 := [rewrite]: #394 +#390 := (iff #86 #389) +#387 := (iff #85 #377) +#382 := (implies true #377) +#385 := (iff #382 #377) +#386 := [rewrite]: #385 +#383 := (iff #85 #382) +#380 := (iff #84 #377) +#373 := (implies #361 #368) +#378 := (iff #373 #377) +#379 := [rewrite]: #378 +#374 := (iff #84 #373) +#371 := (iff #83 #368) +#364 := (implies #82 #286) +#369 := (iff #364 #368) +#370 := [rewrite]: #369 +#365 := (iff #83 #364) +#296 := (iff #69 #286) +#291 := (implies true #286) +#294 := (iff #291 #286) +#295 := [rewrite]: #294 +#292 := (iff #69 #291) +#289 := (iff #68 #286) +#282 := (implies #46 #277) +#287 := (iff #282 #286) +#288 := [rewrite]: #287 +#283 := (iff #68 #282) +#280 := (iff #67 #277) +#273 := (implies #200 #268) +#278 := (iff #273 #277) +#279 := [rewrite]: #278 +#274 := (iff #67 #273) +#271 := (iff #66 #268) +#264 := (implies #203 #252) +#269 := (iff #264 #268) +#270 := [rewrite]: #269 +#265 := (iff #66 #264) +#262 := (iff #65 #252) +#257 := (implies true #252) +#260 := (iff #257 #252) +#261 := [rewrite]: #260 +#258 := (iff #65 #257) +#255 := (iff #64 #252) +#249 := (and #244 #218) +#253 := (iff #249 #252) +#254 := [rewrite]: #253 +#250 := (iff #64 #249) +#219 := (iff #57 #218) +#216 := (iff #56 #213) +#209 := (implies #206 #55) +#214 := (iff #209 #213) +#215 := [rewrite]: #214 +#210 := (iff #56 #209) +#207 := (iff #54 #206) +#208 := [rewrite]: #207 +#211 := [monotonicity #208]: #210 +#217 := [trans #211 #215]: #216 +#220 := [quant-intro #217]: #219 +#247 := (iff #63 #244) +#240 := (implies #218 #221) +#245 := (iff #240 #244) +#246 := [rewrite]: #245 +#241 := (iff #63 #240) +#238 := (iff #62 #221) +#233 := (and true #221) +#236 := (iff #233 #221) +#237 := [rewrite]: #236 +#234 := (iff #62 #233) +#222 := (iff #59 #221) +#223 := [rewrite]: #222 +#231 := (iff #61 true) +#226 := (implies #221 true) +#229 := (iff #226 true) +#230 := [rewrite]: #229 +#227 := (iff #61 #226) +#224 := (iff #60 true) +#225 := [rewrite]: #224 +#228 := [monotonicity #223 #225]: #227 +#232 := [trans #228 #230]: #231 +#235 := [monotonicity #232 #223]: #234 +#239 := [trans #235 #237]: #238 +#242 := [monotonicity #220 #239]: #241 +#248 := [trans #242 #246]: #247 +#251 := [monotonicity #248 #220]: #250 +#256 := [trans #251 #254]: #255 +#259 := [monotonicity #256]: #258 +#263 := [trans #259 #261]: #262 +#204 := (iff #52 #203) +#205 := [rewrite]: #204 +#266 := [monotonicity #205 #263]: #265 +#272 := [trans #266 #270]: #271 +#201 := (iff #49 #200) +#198 := (= #48 #197) +#199 := [rewrite]: #198 +#202 := [monotonicity #199]: #201 +#275 := [monotonicity #202 #272]: #274 +#281 := [trans #275 #279]: #280 +#284 := [monotonicity #281]: #283 +#290 := [trans #284 #288]: #289 +#293 := [monotonicity #290]: #292 +#297 := [trans #293 #295]: #296 +#366 := [monotonicity #297]: #365 +#372 := [trans #366 #370]: #371 +#362 := (iff #81 #361) +#363 := [rewrite]: #362 +#375 := [monotonicity #363 #372]: #374 +#381 := [trans #375 #379]: #380 +#384 := [monotonicity #381]: #383 +#388 := [trans #384 #386]: #387 +#359 := (iff #80 #26) +#360 := [rewrite]: #359 +#391 := [monotonicity #360 #388]: #390 +#397 := [trans #391 #395]: #396 +#357 := (iff #79 #356) +#358 := [rewrite]: #357 +#400 := [monotonicity #358 #397]: #399 +#406 := [trans #400 #404]: #405 +#409 := [monotonicity #406]: #408 +#415 := [trans #409 #413]: #414 +#418 := [monotonicity #415]: #417 +#423 := [trans #418 #421]: #422 +#426 := [monotonicity #423]: #425 +#430 := [trans #426 #428]: #429 +#354 := (iff #76 #344) +#349 := (implies true #344) +#352 := (iff #349 #344) +#353 := [rewrite]: #352 +#350 := (iff #76 #349) +#347 := (iff #75 #344) +#341 := (implies #29 #336) +#345 := (iff #341 #344) +#346 := [rewrite]: #345 +#342 := (iff #75 #341) +#339 := (iff #74 #336) +#332 := (implies #40 #327) +#337 := (iff #332 #336) +#338 := [rewrite]: #337 +#333 := (iff #74 #332) +#330 := (iff #73 #327) +#323 := (implies #29 #311) +#328 := (iff #323 #327) +#329 := [rewrite]: #328 +#324 := (iff #73 #323) +#321 := (iff #72 #311) +#316 := (implies true #311) +#319 := (iff #316 #311) +#320 := [rewrite]: #319 +#317 := (iff #72 #316) +#314 := (iff #71 #311) +#307 := (implies #191 #302) +#312 := (iff #307 #311) +#313 := [rewrite]: #312 +#308 := (iff #71 #307) +#305 := (iff #70 #302) +#298 := (implies #194 #286) +#303 := (iff #298 #302) +#304 := [rewrite]: #303 +#299 := (iff #70 #298) +#195 := (iff #44 #194) +#196 := [rewrite]: #195 +#300 := [monotonicity #196 #297]: #299 +#306 := [trans #300 #304]: #305 +#192 := (iff #42 #191) +#193 := [rewrite]: #192 +#309 := [monotonicity #193 #306]: #308 +#315 := [trans #309 #313]: #314 +#318 := [monotonicity #315]: #317 +#322 := [trans #318 #320]: #321 +#325 := [monotonicity #322]: #324 +#331 := [trans #325 #329]: #330 +#334 := [monotonicity #331]: #333 +#340 := [trans #334 #338]: #339 +#343 := [monotonicity #340]: #342 +#348 := [trans #343 #346]: #347 +#351 := [monotonicity #348]: #350 +#355 := [trans #351 #353]: #354 +#433 := [monotonicity #355 #430]: #432 +#436 := [monotonicity #433]: #435 +#441 := [trans #436 #439]: #440 +#444 := [monotonicity #441]: #443 +#450 := [trans #444 #448]: #449 +#453 := [monotonicity #450]: #452 +#458 := [trans #453 #456]: #457 +#461 := [monotonicity #458]: #460 +#465 := [trans #461 #463]: #464 +#607 := [monotonicity #465 #604]: #606 +#610 := [monotonicity #607]: #609 +#615 := [trans #610 #613]: #614 +#189 := (iff #37 #188) +#190 := [rewrite]: #189 +#618 := [monotonicity #190 #615]: #617 +#624 := [trans #618 #622]: #623 +#186 := (iff #35 #185) +#183 := (iff #34 #180) +#176 := (implies #173 #33) +#181 := (iff #176 #180) +#182 := [rewrite]: #181 +#177 := (iff #34 #176) +#174 := (iff #31 #173) +#175 := [rewrite]: #174 +#178 := [monotonicity #175]: #177 +#184 := [trans #178 #182]: #183 +#187 := [quant-intro #184]: #186 +#627 := [monotonicity #187 #624]: #626 +#633 := [trans #627 #631]: #632 +#636 := [monotonicity #633]: #635 +#641 := [trans #636 #639]: #640 +#644 := [monotonicity #641]: #643 +#648 := [trans #644 #646]: #647 +#651 := [monotonicity #172 #648]: #650 +#657 := [trans #651 #655]: #656 +#660 := [monotonicity #657 #172]: #659 +#665 := [trans #660 #663]: #664 +#668 := [monotonicity #170 #665]: #667 +#674 := [trans #668 #672]: #673 +#677 := [monotonicity #674 #170]: #676 +#682 := [trans #677 #680]: #681 +#162 := (iff #15 #161) +#159 := (iff #14 #158) +#156 := (iff #13 #12) +#157 := [rewrite]: #156 +#160 := [monotonicity #157]: #159 +#163 := [monotonicity #160]: #162 +#685 := [monotonicity #163 #682]: #684 +#691 := [trans #685 #689]: #690 +#694 := [monotonicity #691]: #693 +#699 := [trans #694 #697]: #698 +#702 := [monotonicity #699]: #701 +#706 := [trans #702 #704]: #705 +#709 := [monotonicity #706]: #708 +#715 := [trans #709 #713]: #714 +#718 := [monotonicity #715]: #717 +#722 := [trans #718 #720]: #721 +#725 := [monotonicity #722]: #724 +#1107 := [trans #725 #1105]: #1106 +#155 := [asserted]: #139 +#1108 := [mp #155 #1107]: #1103 +#1109 := [not-or-elim #1108]: #9 +#2193 := [trans #1109 #2192]: #2202 +#1888 := (not #1154) +#1974 := (or #1540 #1888) +#1889 := [def-axiom]: #1974 +#2194 := [unit-resolution #1889 #2205]: #1888 +#2195 := (not #2202) +#2196 := (or #2195 #1154) +#2197 := [th-lemma]: #2196 +#2198 := [unit-resolution #2197 #2194 #2193]: false +#2199 := [lemma #2198]: #1540 +#2393 := (or #1545 #2390) +#1708 := (forall (vars (?x7 int)) #1705) +#1801 := (or #1708 #1798) +#1804 := (not #1801) +#1807 := (or #560 #551 #542 #886 #1654 #927 #1804) +#1810 := (not #1807) +#1612 := (forall (vars (?x3 int)) #1607) +#1618 := (not #1612) +#1619 := (or #221 #1618) +#1620 := (not #1619) +#1648 := (or #1620 #1645) +#1657 := (not #1648) +#1667 := (or #367 #401 #376 #886 #1654 #1655 #1656 #840 #850 #1657) +#1668 := (not #1667) +#1658 := (or #310 #301 #886 #1654 #1655 #1656 #840 #854 #1657) +#1659 := (not #1658) +#1673 := (or #1659 #1668) +#1679 := (not #1673) +#1680 := (or #886 #1654 #926 #1679) +#1681 := (not #1680) +#1813 := (or #1681 #1810) +#1816 := (not #1813) +#1590 := (forall (vars (?x2 int)) #1585) +#1784 := (not #1590) +#1568 := (forall (vars (?x1 int)) #1563) +#1783 := (not #1568) +#1819 := (or #619 #886 #1654 #1783 #1784 #1816) +#1822 := (not #1819) +#1825 := (or #1545 #1822) +#2394 := (iff #1825 #2393) +#2391 := (iff #1822 #2390) +#2388 := (iff #1819 #2387) +#2385 := (iff #1816 #2384) +#2382 := (iff #1813 #2381) +#2379 := (iff #1810 #2378) +#2376 := (iff #1807 #2375) +#2373 := (iff #1804 #2372) +#2370 := (iff #1801 #2369) +#2367 := (iff #1708 #2364) +#2365 := (iff #1705 #1705) +#2366 := [refl]: #2365 +#2368 := [quant-intro #2366]: #2367 +#2371 := [monotonicity #2368]: #2370 +#2374 := [monotonicity #2371]: #2373 +#2377 := [monotonicity #2374]: #2376 +#2380 := [monotonicity #2377]: #2379 +#2362 := (iff #1681 #2361) +#2359 := (iff #1680 #2358) +#2356 := (iff #1679 #2355) +#2353 := (iff #1673 #2352) +#2350 := (iff #1668 #2349) +#2347 := (iff #1667 #2346) +#2338 := (iff #1657 #2337) +#2335 := (iff #1648 #2334) +#2332 := (iff #1620 #2331) +#2329 := (iff #1619 #2328) +#2326 := (iff #1618 #2325) +#2323 := (iff #1612 #2320) +#2321 := (iff #1607 #1607) +#2322 := [refl]: #2321 +#2324 := [quant-intro #2322]: #2323 +#2327 := [monotonicity #2324]: #2326 +#2330 := [monotonicity #2327]: #2329 +#2333 := [monotonicity #2330]: #2332 +#2336 := [monotonicity #2333]: #2335 +#2339 := [monotonicity #2336]: #2338 +#2348 := [monotonicity #2339]: #2347 +#2351 := [monotonicity #2348]: #2350 +#2344 := (iff #1659 #2343) +#2341 := (iff #1658 #2340) +#2342 := [monotonicity #2339]: #2341 +#2345 := [monotonicity #2342]: #2344 +#2354 := [monotonicity #2345 #2351]: #2353 +#2357 := [monotonicity #2354]: #2356 +#2360 := [monotonicity #2357]: #2359 +#2363 := [monotonicity #2360]: #2362 +#2383 := [monotonicity #2363 #2380]: #2382 +#2386 := [monotonicity #2383]: #2385 +#2318 := (iff #1784 #2317) +#2315 := (iff #1590 #2312) +#2313 := (iff #1585 #1585) +#2314 := [refl]: #2313 +#2316 := [quant-intro #2314]: #2315 +#2319 := [monotonicity #2316]: #2318 +#2310 := (iff #1783 #2309) +#2307 := (iff #1568 #2304) +#2305 := (iff #1563 #1563) +#2306 := [refl]: #2305 +#2308 := [quant-intro #2306]: #2307 +#2311 := [monotonicity #2308]: #2310 +#2389 := [monotonicity #2311 #2319 #2386]: #2388 +#2392 := [monotonicity #2389]: #2391 +#2395 := [monotonicity #2392]: #2394 +#1304 := (not #1303) +#1481 := (and #1304 #1305) +#1484 := (not #1481) +#1500 := (or #1484 #1495) +#1503 := (not #1500) +#1283 := (not #1282) +#1472 := (and #1283 #1284) +#1475 := (not #1472) +#1478 := (or #1469 #1475) +#1506 := (and #1478 #1503) +#1273 := (not #963) +#1276 := (forall (vars (?x7 int)) #1273) +#1509 := (or #1276 #1506) +#1515 := (and #466 #469 #472 #776 #779 #926 #1509) +#1401 := (not #1396) +#1404 := (and #1192 #1401) +#1407 := (not #1404) +#1410 := (or #1383 #1407) +#1413 := (not #1410) +#1204 := (not #221) +#1214 := (and #1204 #819) +#1419 := (or #1214 #1413) +#1447 := (and #82 #356 #361 #776 #779 #788 #791 #837 #854 #1419) +#1431 := (and #191 #194 #776 #779 #788 #791 #837 #850 #1419) +#1452 := (or #1431 #1447) +#1458 := (and #776 #779 #927 #1452) +#1520 := (or #1458 #1515) +#1526 := (and #188 #769 #776 #779 #1043 #1520) +#1348 := (and #1155 #1157) +#1351 := (not #1348) +#1357 := (or #1154 #1351) +#1362 := (not #1357) +#1531 := (or #1362 #1526) +#1828 := (iff #1531 #1825) +#1746 := (or #1303 #1731 #1495) +#1758 := (or #1757 #1746) +#1759 := (not #1758) +#1764 := (or #1708 #1759) +#1770 := (not #1764) +#1771 := (or #560 #551 #542 #886 #1654 #927 #1770) +#1772 := (not #1771) +#1777 := (or #1681 #1772) +#1785 := (not #1777) +#1786 := (or #619 #886 #1654 #1783 #1784 #1785) +#1787 := (not #1786) +#1792 := (or #1545 #1787) +#1826 := (iff #1792 #1825) +#1823 := (iff #1787 #1822) +#1820 := (iff #1786 #1819) +#1817 := (iff #1785 #1816) +#1814 := (iff #1777 #1813) +#1811 := (iff #1772 #1810) +#1808 := (iff #1771 #1807) +#1805 := (iff #1770 #1804) +#1802 := (iff #1764 #1801) +#1799 := (iff #1759 #1798) +#1796 := (iff #1758 #1795) +#1797 := [rewrite]: #1796 +#1800 := [monotonicity #1797]: #1799 +#1803 := [monotonicity #1800]: #1802 +#1806 := [monotonicity #1803]: #1805 +#1809 := [monotonicity #1806]: #1808 +#1812 := [monotonicity #1809]: #1811 +#1815 := [monotonicity #1812]: #1814 +#1818 := [monotonicity #1815]: #1817 +#1821 := [monotonicity #1818]: #1820 +#1824 := [monotonicity #1821]: #1823 +#1827 := [monotonicity #1824]: #1826 +#1793 := (iff #1531 #1792) +#1790 := (iff #1526 #1787) +#1780 := (and #188 #1568 #776 #779 #1590 #1777) +#1788 := (iff #1780 #1787) +#1789 := [rewrite]: #1788 +#1781 := (iff #1526 #1780) +#1778 := (iff #1520 #1777) +#1775 := (iff #1515 #1772) +#1767 := (and #466 #469 #472 #776 #779 #926 #1764) +#1773 := (iff #1767 #1772) +#1774 := [rewrite]: #1773 +#1768 := (iff #1515 #1767) +#1765 := (iff #1509 #1764) +#1762 := (iff #1506 #1759) +#1751 := (not #1746) +#1754 := (and #1726 #1751) +#1760 := (iff #1754 #1759) +#1761 := [rewrite]: #1760 +#1755 := (iff #1506 #1754) +#1752 := (iff #1503 #1751) +#1749 := (iff #1500 #1746) +#1732 := (or #1303 #1731) +#1743 := (or #1732 #1495) +#1747 := (iff #1743 #1746) +#1748 := [rewrite]: #1747 +#1744 := (iff #1500 #1743) +#1741 := (iff #1484 #1732) +#1733 := (not #1732) +#1736 := (not #1733) +#1739 := (iff #1736 #1732) +#1740 := [rewrite]: #1739 +#1737 := (iff #1484 #1736) +#1734 := (iff #1481 #1733) +#1735 := [rewrite]: #1734 +#1738 := [monotonicity #1735]: #1737 +#1742 := [trans #1738 #1740]: #1741 +#1745 := [monotonicity #1742]: #1744 +#1750 := [trans #1745 #1748]: #1749 +#1753 := [monotonicity #1750]: #1752 +#1729 := (iff #1478 #1726) +#1712 := (or #1282 #1711) +#1723 := (or #1469 #1712) +#1727 := (iff #1723 #1726) +#1728 := [rewrite]: #1727 +#1724 := (iff #1478 #1723) +#1721 := (iff #1475 #1712) +#1713 := (not #1712) +#1716 := (not #1713) +#1719 := (iff #1716 #1712) +#1720 := [rewrite]: #1719 +#1717 := (iff #1475 #1716) +#1714 := (iff #1472 #1713) +#1715 := [rewrite]: #1714 +#1718 := [monotonicity #1715]: #1717 +#1722 := [trans #1718 #1720]: #1721 +#1725 := [monotonicity #1722]: #1724 +#1730 := [trans #1725 #1728]: #1729 +#1756 := [monotonicity #1730 #1753]: #1755 +#1763 := [trans #1756 #1761]: #1762 +#1709 := (iff #1276 #1708) +#1706 := (iff #1273 #1705) +#1703 := (iff #963 #1700) +#1686 := (or #1548 #953) +#1697 := (or #105 #1686) +#1701 := (iff #1697 #1700) +#1702 := [rewrite]: #1701 +#1698 := (iff #963 #1697) +#1695 := (iff #960 #1686) +#1687 := (not #1686) +#1690 := (not #1687) +#1693 := (iff #1690 #1686) +#1694 := [rewrite]: #1693 +#1691 := (iff #960 #1690) +#1688 := (iff #957 #1687) +#1689 := [rewrite]: #1688 +#1692 := [monotonicity #1689]: #1691 +#1696 := [trans #1692 #1694]: #1695 +#1699 := [monotonicity #1696]: #1698 +#1704 := [trans #1699 #1702]: #1703 +#1707 := [monotonicity #1704]: #1706 +#1710 := [quant-intro #1707]: #1709 +#1766 := [monotonicity #1710 #1763]: #1765 +#1769 := [monotonicity #1766]: #1768 +#1776 := [trans #1769 #1774]: #1775 +#1684 := (iff #1458 #1681) +#1676 := (and #776 #779 #927 #1673) +#1682 := (iff #1676 #1681) +#1683 := [rewrite]: #1682 +#1677 := (iff #1458 #1676) +#1674 := (iff #1452 #1673) +#1671 := (iff #1447 #1668) +#1664 := (and #82 #356 #361 #776 #779 #788 #791 #837 #854 #1648) +#1669 := (iff #1664 #1668) +#1670 := [rewrite]: #1669 +#1665 := (iff #1447 #1664) +#1649 := (iff #1419 #1648) +#1646 := (iff #1413 #1645) +#1643 := (iff #1410 #1640) +#1626 := (or #1625 #1396) +#1637 := (or #1383 #1626) +#1641 := (iff #1637 #1640) +#1642 := [rewrite]: #1641 +#1638 := (iff #1410 #1637) +#1635 := (iff #1407 #1626) +#1627 := (not #1626) +#1630 := (not #1627) +#1633 := (iff #1630 #1626) +#1634 := [rewrite]: #1633 +#1631 := (iff #1407 #1630) +#1628 := (iff #1404 #1627) +#1629 := [rewrite]: #1628 +#1632 := [monotonicity #1629]: #1631 +#1636 := [trans #1632 #1634]: #1635 +#1639 := [monotonicity #1636]: #1638 +#1644 := [trans #1639 #1642]: #1643 +#1647 := [monotonicity #1644]: #1646 +#1623 := (iff #1214 #1620) +#1615 := (and #1204 #1612) +#1621 := (iff #1615 #1620) +#1622 := [rewrite]: #1621 +#1616 := (iff #1214 #1615) +#1613 := (iff #819 #1612) +#1610 := (iff #816 #1607) +#1593 := (or #1548 #805) +#1604 := (or #801 #1593) +#1608 := (iff #1604 #1607) +#1609 := [rewrite]: #1608 +#1605 := (iff #816 #1604) +#1602 := (iff #813 #1593) +#1594 := (not #1593) +#1597 := (not #1594) +#1600 := (iff #1597 #1593) +#1601 := [rewrite]: #1600 +#1598 := (iff #813 #1597) +#1595 := (iff #810 #1594) +#1596 := [rewrite]: #1595 +#1599 := [monotonicity #1596]: #1598 +#1603 := [trans #1599 #1601]: #1602 +#1606 := [monotonicity #1603]: #1605 +#1611 := [trans #1606 #1609]: #1610 +#1614 := [quant-intro #1611]: #1613 +#1617 := [monotonicity #1614]: #1616 +#1624 := [trans #1617 #1622]: #1623 +#1650 := [monotonicity #1624 #1647]: #1649 +#1666 := [monotonicity #1650]: #1665 +#1672 := [trans #1666 #1670]: #1671 +#1662 := (iff #1431 #1659) +#1651 := (and #191 #194 #776 #779 #788 #791 #837 #850 #1648) +#1660 := (iff #1651 #1659) +#1661 := [rewrite]: #1660 +#1652 := (iff #1431 #1651) +#1653 := [monotonicity #1650]: #1652 +#1663 := [trans #1653 #1661]: #1662 +#1675 := [monotonicity #1663 #1672]: #1674 +#1678 := [monotonicity #1675]: #1677 +#1685 := [trans #1678 #1683]: #1684 +#1779 := [monotonicity #1685 #1776]: #1778 +#1591 := (iff #1043 #1590) +#1588 := (iff #1040 #1585) +#1571 := (or #1548 #1029) +#1582 := (or #1026 #1571) +#1586 := (iff #1582 #1585) +#1587 := [rewrite]: #1586 +#1583 := (iff #1040 #1582) +#1580 := (iff #1037 #1571) +#1572 := (not #1571) +#1575 := (not #1572) +#1578 := (iff #1575 #1571) +#1579 := [rewrite]: #1578 +#1576 := (iff #1037 #1575) +#1573 := (iff #1034 #1572) +#1574 := [rewrite]: #1573 +#1577 := [monotonicity #1574]: #1576 +#1581 := [trans #1577 #1579]: #1580 +#1584 := [monotonicity #1581]: #1583 +#1589 := [trans #1584 #1587]: #1588 +#1592 := [quant-intro #1589]: #1591 +#1569 := (iff #769 #1568) +#1566 := (iff #766 #1563) +#1549 := (or #749 #1548) +#1560 := (or #1549 #762) +#1564 := (iff #1560 #1563) +#1565 := [rewrite]: #1564 +#1561 := (iff #766 #1560) +#1558 := (iff #757 #1549) +#1550 := (not #1549) +#1553 := (not #1550) +#1556 := (iff #1553 #1549) +#1557 := [rewrite]: #1556 +#1554 := (iff #757 #1553) +#1551 := (iff #754 #1550) +#1552 := [rewrite]: #1551 +#1555 := [monotonicity #1552]: #1554 +#1559 := [trans #1555 #1557]: #1558 +#1562 := [monotonicity #1559]: #1561 +#1567 := [trans #1562 #1565]: #1566 +#1570 := [quant-intro #1567]: #1569 +#1782 := [monotonicity #1570 #1592 #1779]: #1781 +#1791 := [trans #1782 #1789]: #1790 +#1546 := (iff #1362 #1545) +#1543 := (iff #1357 #1540) +#1165 := (or #1164 #1156) +#1537 := (or #1154 #1165) +#1541 := (iff #1537 #1540) +#1542 := [rewrite]: #1541 +#1538 := (iff #1357 #1537) +#1535 := (iff #1351 #1165) +#1292 := (not #1165) +#1314 := (not #1292) +#1347 := (iff #1314 #1165) +#1534 := [rewrite]: #1347 +#1202 := (iff #1351 #1314) +#1293 := (iff #1348 #1292) +#1313 := [rewrite]: #1293 +#1203 := [monotonicity #1313]: #1202 +#1536 := [trans #1203 #1534]: #1535 +#1539 := [monotonicity #1536]: #1538 +#1544 := [trans #1539 #1542]: #1543 +#1547 := [monotonicity #1544]: #1546 +#1794 := [monotonicity #1547 #1791]: #1793 +#1829 := [trans #1794 #1827]: #1828 +#1299 := (+ #1298 #972) +#1300 := (<= #1299 0::int) +#1306 := (and #1305 #1304) +#1307 := (not #1306) +#1308 := (or #1307 #1300) +#1309 := (not #1308) +#1285 := (and #1284 #1283) +#1286 := (not #1285) +#1288 := (= #1287 uf_12) +#1289 := (or #1288 #1286) +#1315 := (and #1289 #1309) +#1319 := (or #1276 #1315) +#1176 := (not #784) +#1268 := (not #542) +#1265 := (not #551) +#1262 := (not #560) +#1323 := (and #1262 #1265 #1268 #1176 #930 #1319) +#1225 := (not #846) +#1222 := (not #840) +#1189 := (+ ?x3!1 #806) +#1190 := (>= #1189 0::int) +#1191 := (not #1190) +#1193 := (and #1192 #1191) +#1194 := (not #1193) +#1196 := (+ #1195 #799) +#1197 := (<= #1196 0::int) +#1198 := (or #1197 #1194) +#1199 := (not #1198) +#1218 := (or #1199 #1214) +#1185 := (not #796) +#1243 := (not #886) +#1240 := (not #376) +#1237 := (not #401) +#1234 := (not #367) +#1248 := (and #1234 #1237 #1240 #1243 #1176 #1185 #1218 #1222 #1225 #854) +#1182 := (not #301) +#1179 := (not #310) +#1230 := (and #1179 #1182 #1176 #1185 #1218 #1222 #1225 #891) +#1252 := (or #1230 #1248) +#1258 := (and #1176 #1252 #927) +#1327 := (or #1258 #1323) +#1166 := (not #619) +#1338 := (and #1166 #769 #1176 #1327 #1043) +#1158 := (and #1157 #1155) +#1159 := (not #1158) +#1160 := (or #1159 #1154) +#1161 := (not #1160) +#1342 := (or #1161 #1338) +#1532 := (iff #1342 #1531) +#1529 := (iff #1338 #1526) +#1523 := (and #188 #769 #781 #1520 #1043) +#1527 := (iff #1523 #1526) +#1528 := [rewrite]: #1527 +#1524 := (iff #1338 #1523) +#1521 := (iff #1327 #1520) +#1518 := (iff #1323 #1515) +#1512 := (and #466 #469 #472 #781 #926 #1509) +#1516 := (iff #1512 #1515) +#1517 := [rewrite]: #1516 +#1513 := (iff #1323 #1512) +#1510 := (iff #1319 #1509) +#1507 := (iff #1315 #1506) +#1504 := (iff #1309 #1503) +#1501 := (iff #1308 #1500) +#1498 := (iff #1300 #1495) +#1487 := (+ #972 #1298) +#1490 := (<= #1487 0::int) +#1496 := (iff #1490 #1495) +#1497 := [rewrite]: #1496 +#1491 := (iff #1300 #1490) +#1488 := (= #1299 #1487) +#1489 := [rewrite]: #1488 +#1492 := [monotonicity #1489]: #1491 +#1499 := [trans #1492 #1497]: #1498 +#1485 := (iff #1307 #1484) +#1482 := (iff #1306 #1481) +#1483 := [rewrite]: #1482 +#1486 := [monotonicity #1483]: #1485 +#1502 := [monotonicity #1486 #1499]: #1501 +#1505 := [monotonicity #1502]: #1504 +#1479 := (iff #1289 #1478) +#1476 := (iff #1286 #1475) +#1473 := (iff #1285 #1472) +#1474 := [rewrite]: #1473 +#1477 := [monotonicity #1474]: #1476 +#1470 := (iff #1288 #1469) +#1471 := [rewrite]: #1470 +#1480 := [monotonicity #1471 #1477]: #1479 +#1508 := [monotonicity #1480 #1505]: #1507 +#1511 := [monotonicity #1508]: #1510 +#1367 := (iff #1176 #781) +#1368 := [rewrite]: #1367 +#1467 := (iff #1268 #472) +#1468 := [rewrite]: #1467 +#1465 := (iff #1265 #469) +#1466 := [rewrite]: #1465 +#1463 := (iff #1262 #466) +#1464 := [rewrite]: #1463 +#1514 := [monotonicity #1464 #1466 #1468 #1368 #934 #1511]: #1513 +#1519 := [trans #1514 #1517]: #1518 +#1461 := (iff #1258 #1458) +#1455 := (and #781 #1452 #927) +#1459 := (iff #1455 #1458) +#1460 := [rewrite]: #1459 +#1456 := (iff #1258 #1455) +#1453 := (iff #1252 #1452) +#1450 := (iff #1248 #1447) +#1444 := (and #82 #356 #361 #776 #781 #793 #1419 #837 #843 #854) +#1448 := (iff #1444 #1447) +#1449 := [rewrite]: #1448 +#1445 := (iff #1248 #1444) +#1426 := (iff #1225 #843) +#1427 := [rewrite]: #1426 +#1424 := (iff #1222 #837) +#1425 := [rewrite]: #1424 +#1422 := (iff #1218 #1419) +#1416 := (or #1413 #1214) +#1420 := (iff #1416 #1419) +#1421 := [rewrite]: #1420 +#1417 := (iff #1218 #1416) +#1414 := (iff #1199 #1413) +#1411 := (iff #1198 #1410) +#1408 := (iff #1194 #1407) +#1405 := (iff #1193 #1404) +#1402 := (iff #1191 #1401) +#1399 := (iff #1190 #1396) +#1388 := (+ #806 ?x3!1) +#1391 := (>= #1388 0::int) +#1397 := (iff #1391 #1396) +#1398 := [rewrite]: #1397 +#1392 := (iff #1190 #1391) +#1389 := (= #1189 #1388) +#1390 := [rewrite]: #1389 +#1393 := [monotonicity #1390]: #1392 +#1400 := [trans #1393 #1398]: #1399 +#1403 := [monotonicity #1400]: #1402 +#1406 := [monotonicity #1403]: #1405 +#1409 := [monotonicity #1406]: #1408 +#1386 := (iff #1197 #1383) +#1375 := (+ #799 #1195) +#1378 := (<= #1375 0::int) +#1384 := (iff #1378 #1383) +#1385 := [rewrite]: #1384 +#1379 := (iff #1197 #1378) +#1376 := (= #1196 #1375) +#1377 := [rewrite]: #1376 +#1380 := [monotonicity #1377]: #1379 +#1387 := [trans #1380 #1385]: #1386 +#1412 := [monotonicity #1387 #1409]: #1411 +#1415 := [monotonicity #1412]: #1414 +#1418 := [monotonicity #1415]: #1417 +#1423 := [trans #1418 #1421]: #1422 +#1373 := (iff #1185 #793) +#1374 := [rewrite]: #1373 +#1442 := (iff #1243 #776) +#1443 := [rewrite]: #1442 +#1440 := (iff #1240 #361) +#1441 := [rewrite]: #1440 +#1438 := (iff #1237 #356) +#1439 := [rewrite]: #1438 +#1436 := (iff #1234 #82) +#1437 := [rewrite]: #1436 +#1446 := [monotonicity #1437 #1439 #1441 #1443 #1368 #1374 #1423 #1425 #1427]: #1445 +#1451 := [trans #1446 #1449]: #1450 +#1434 := (iff #1230 #1431) +#1428 := (and #191 #194 #781 #793 #1419 #837 #843 #850) +#1432 := (iff #1428 #1431) +#1433 := [rewrite]: #1432 +#1429 := (iff #1230 #1428) +#1371 := (iff #1182 #194) +#1372 := [rewrite]: #1371 +#1369 := (iff #1179 #191) +#1370 := [rewrite]: #1369 +#1430 := [monotonicity #1370 #1372 #1368 #1374 #1423 #1425 #1427 #895]: #1429 +#1435 := [trans #1430 #1433]: #1434 +#1454 := [monotonicity #1435 #1451]: #1453 +#1457 := [monotonicity #1368 #1454]: #1456 +#1462 := [trans #1457 #1460]: #1461 +#1522 := [monotonicity #1462 #1519]: #1521 +#1365 := (iff #1166 #188) +#1366 := [rewrite]: #1365 +#1525 := [monotonicity #1366 #1368 #1522]: #1524 +#1530 := [trans #1525 #1528]: #1529 +#1363 := (iff #1161 #1362) +#1360 := (iff #1160 #1357) +#1354 := (or #1351 #1154) +#1358 := (iff #1354 #1357) +#1359 := [rewrite]: #1358 +#1355 := (iff #1160 #1354) +#1352 := (iff #1159 #1351) +#1349 := (iff #1158 #1348) +#1350 := [rewrite]: #1349 +#1353 := [monotonicity #1350]: #1352 +#1356 := [monotonicity #1353]: #1355 +#1361 := [trans #1356 #1359]: #1360 +#1364 := [monotonicity #1361]: #1363 +#1533 := [monotonicity #1364 #1530]: #1532 +#1138 := (or #619 #772 #784 #1021 #1046) +#1143 := (and #769 #1138) +#1146 := (not #1143) +#1343 := (~ #1146 #1342) +#1339 := (not #1138) +#1340 := (~ #1339 #1338) +#1335 := (not #1046) +#1336 := (~ #1335 #1043) +#1333 := (~ #1043 #1043) +#1331 := (~ #1040 #1040) +#1332 := [refl]: #1331 +#1334 := [nnf-pos #1332]: #1333 +#1337 := [nnf-neg #1334]: #1336 +#1328 := (not #1021) +#1329 := (~ #1328 #1327) +#1324 := (not #1016) +#1325 := (~ #1324 #1323) +#1320 := (not #991) +#1321 := (~ #1320 #1319) +#1316 := (not #988) +#1317 := (~ #1316 #1315) +#1310 := (not #985) +#1311 := (~ #1310 #1309) +#1312 := [sk]: #1311 +#1294 := (not #969) +#1295 := (~ #1294 #1289) +#1290 := (~ #966 #1289) +#1291 := [sk]: #1290 +#1296 := [nnf-neg #1291]: #1295 +#1318 := [nnf-neg #1296 #1312]: #1317 +#1277 := (~ #969 #1276) +#1274 := (~ #1273 #1273) +#1275 := [refl]: #1274 +#1278 := [nnf-neg #1275]: #1277 +#1322 := [nnf-neg #1278 #1318]: #1321 +#1271 := (~ #930 #930) +#1272 := [refl]: #1271 +#1177 := (~ #1176 #1176) +#1178 := [refl]: #1177 +#1269 := (~ #1268 #1268) +#1270 := [refl]: #1269 +#1266 := (~ #1265 #1265) +#1267 := [refl]: #1266 +#1263 := (~ #1262 #1262) +#1264 := [refl]: #1263 +#1326 := [nnf-neg #1264 #1267 #1270 #1178 #1272 #1322]: #1325 +#1259 := (not #946) +#1260 := (~ #1259 #1258) +#1256 := (~ #927 #927) +#1257 := [refl]: #1256 +#1253 := (not #921) +#1254 := (~ #1253 #1252) +#1249 := (not #916) +#1250 := (~ #1249 #1248) +#1246 := (~ #854 #854) +#1247 := [refl]: #1246 +#1226 := (~ #1225 #1225) +#1227 := [refl]: #1226 +#1223 := (~ #1222 #1222) +#1224 := [refl]: #1223 +#1219 := (not #833) +#1220 := (~ #1219 #1218) +#1215 := (not #828) +#1216 := (~ #1215 #1214) +#1211 := (not #822) +#1212 := (~ #1211 #819) +#1209 := (~ #819 #819) +#1207 := (~ #816 #816) +#1208 := [refl]: #1207 +#1210 := [nnf-pos #1208]: #1209 +#1213 := [nnf-neg #1210]: #1212 +#1205 := (~ #1204 #1204) +#1206 := [refl]: #1205 +#1217 := [nnf-neg #1206 #1213]: #1216 +#1200 := (~ #822 #1199) +#1201 := [sk]: #1200 +#1221 := [nnf-neg #1201 #1217]: #1220 +#1186 := (~ #1185 #1185) +#1187 := [refl]: #1186 +#1244 := (~ #1243 #1243) +#1245 := [refl]: #1244 +#1241 := (~ #1240 #1240) +#1242 := [refl]: #1241 +#1238 := (~ #1237 #1237) +#1239 := [refl]: #1238 +#1235 := (~ #1234 #1234) +#1236 := [refl]: #1235 +#1251 := [nnf-neg #1236 #1239 #1242 #1245 #1178 #1187 #1221 #1224 #1227 #1247]: #1250 +#1231 := (not #881) +#1232 := (~ #1231 #1230) +#1228 := (~ #891 #891) +#1229 := [refl]: #1228 +#1183 := (~ #1182 #1182) +#1184 := [refl]: #1183 +#1180 := (~ #1179 #1179) +#1181 := [refl]: #1180 +#1233 := [nnf-neg #1181 #1184 #1178 #1187 #1221 #1224 #1227 #1229]: #1232 +#1255 := [nnf-neg #1233 #1251]: #1254 +#1261 := [nnf-neg #1178 #1255 #1257]: #1260 +#1330 := [nnf-neg #1261 #1326]: #1329 +#1173 := (not #772) +#1174 := (~ #1173 #769) +#1171 := (~ #769 #769) +#1169 := (~ #766 #766) +#1170 := [refl]: #1169 +#1172 := [nnf-pos #1170]: #1171 +#1175 := [nnf-neg #1172]: #1174 +#1167 := (~ #1166 #1166) +#1168 := [refl]: #1167 +#1341 := [nnf-neg #1168 #1175 #1178 #1330 #1337]: #1340 +#1162 := (~ #772 #1161) +#1163 := [sk]: #1162 +#1344 := [nnf-neg #1163 #1341]: #1343 +#1110 := (not #1075) +#1147 := (iff #1110 #1146) +#1144 := (iff #1075 #1143) +#1141 := (iff #1072 #1138) +#1123 := (or #619 #784 #1021 #1046) +#1135 := (or #772 #1123) +#1139 := (iff #1135 #1138) +#1140 := [rewrite]: #1139 +#1136 := (iff #1072 #1135) +#1133 := (iff #1069 #1123) +#1128 := (and true #1123) +#1131 := (iff #1128 #1123) +#1132 := [rewrite]: #1131 +#1129 := (iff #1069 #1128) +#1126 := (iff #1064 #1123) +#1120 := (or false #619 #784 #1021 #1046) +#1124 := (iff #1120 #1123) +#1125 := [rewrite]: #1124 +#1121 := (iff #1064 #1120) +#1118 := (iff #652 false) +#1116 := (iff #652 #740) +#1115 := (iff #9 true) +#1113 := [iff-true #1109]: #1115 +#1117 := [monotonicity #1113]: #1116 +#1119 := [trans #1117 #744]: #1118 +#1122 := [monotonicity #1119]: #1121 +#1127 := [trans #1122 #1125]: #1126 +#1130 := [monotonicity #1113 #1127]: #1129 +#1134 := [trans #1130 #1132]: #1133 +#1137 := [monotonicity #1134]: #1136 +#1142 := [trans #1137 #1140]: #1141 +#1145 := [monotonicity #1142]: #1144 +#1148 := [monotonicity #1145]: #1147 +#1111 := [not-or-elim #1108]: #1110 +#1149 := [mp #1111 #1148]: #1146 +#1345 := [mp~ #1149 #1344]: #1342 +#1346 := [mp #1345 #1533]: #1531 +#1830 := [mp #1346 #1829]: #1825 +#2396 := [mp #1830 #2395]: #2393 +#1909 := [unit-resolution #2396 #2199]: #2390 +#2214 := (or #2387 #2381) +#2210 := [def-axiom]: #2214 +#2414 := [unit-resolution #2210 #1909]: #2381 +#2426 := (uf_3 uf_11) +#2430 := (= uf_12 #2426) +#2480 := (= #36 #2426) +#2478 := (= #2426 #36) +#2463 := [hypothesis]: #2378 +#2138 := (or #2375 #466) +#2139 := [def-axiom]: #2138 +#2474 := [unit-resolution #2139 #2463]: #466 +#2475 := [symm #2474]: #98 +#2479 := [monotonicity #2475]: #2478 +#2481 := [symm #2479]: #2480 +#2482 := (= uf_12 #36) +#2221 := (or #2387 #188) +#2222 := [def-axiom]: #2221 +#2476 := [unit-resolution #2222 #1909]: #188 +#2132 := (or #2375 #469) +#2140 := [def-axiom]: #2132 +#2466 := [unit-resolution #2140 #2463]: #469 +#2477 := [symm #2466]: #100 +#2483 := [trans #2477 #2476]: #2482 +#2484 := [trans #2483 #2481]: #2430 +#2458 := (not #2430) +#2424 := (>= uf_11 0::int) +#2425 := (not #2424) +#2421 := (* -1::int uf_11) +#2422 := (+ uf_1 #2421) +#2423 := (<= #2422 0::int) +#2436 := (or #2423 #2425 #2430) +#2441 := (not #2436) +#2227 := (or #2375 #2369) +#2219 := [def-axiom]: #2227 +#2464 := [unit-resolution #2219 #2463]: #2369 +#2238 := (or #2375 #926) +#2225 := [def-axiom]: #2238 +#2465 := [unit-resolution #2225 #2463]: #926 +#2031 := (+ uf_6 #972) +#2032 := (<= #2031 0::int) +#2467 := (or #551 #2032) +#2468 := [th-lemma]: #2467 +#2469 := [unit-resolution #2468 #2466]: #2032 +#1925 := (not #2032) +#1915 := (or #1795 #1925 #927) +#2004 := (+ uf_4 #1301) +#2005 := (<= #2004 0::int) +#1918 := (not #2005) +#1910 := [hypothesis]: #926 +#1911 := [hypothesis]: #1798 +#2086 := (or #1795 #1304) +#2239 := [def-axiom]: #2086 +#1916 := [unit-resolution #2239 #1911]: #1304 +#1919 := (or #1918 #927 #1303) +#1921 := [th-lemma]: #1919 +#1917 := [unit-resolution #1921 #1916 #1910]: #1918 +#2021 := (+ uf_6 #1493) +#2022 := (>= #2021 0::int) +#1930 := (not #2022) +#1936 := [hypothesis]: #2032 +#2243 := (not #1495) +#2241 := (or #1795 #2243) +#2244 := [def-axiom]: #2241 +#1922 := [unit-resolution #2244 #1911]: #2243 +#1931 := (or #1930 #1495 #1925) +#1924 := [hypothesis]: #2243 +#1926 := [hypothesis]: #2022 +#1927 := [th-lemma #1926 #1924 #1936]: false +#1906 := [lemma #1927]: #1931 +#1905 := [unit-resolution #1906 #1922 #1936]: #1930 +#1913 := (or #2005 #2022) +#2240 := (or #1795 #1305) +#2242 := [def-axiom]: #2240 +#1908 := [unit-resolution #2242 #1911]: #1305 +#2212 := (or #2387 #2312) +#2213 := [def-axiom]: #2212 +#1912 := [unit-resolution #2213 #1909]: #2312 +#1994 := (or #2317 #1731 #2005 #2022) +#2034 := (+ ?x8!3 #924) +#2035 := (>= #2034 0::int) +#2036 := (+ #1298 #1024) +#2037 := (<= #2036 0::int) +#2026 := (or #1731 #2037 #2035) +#1961 := (or #2317 #2026) +#1971 := (iff #1961 #1994) +#1991 := (or #1731 #2005 #2022) +#1964 := (or #2317 #1991) +#1969 := (iff #1964 #1994) +#1970 := [rewrite]: #1969 +#1955 := (iff #1961 #1964) +#1993 := (iff #2026 #1991) +#2008 := (or #1731 #2022 #2005) +#1983 := (iff #2008 #1991) +#1992 := [rewrite]: #1983 +#1989 := (iff #2026 #2008) +#2007 := (iff #2035 #2005) +#2010 := (+ #924 ?x8!3) +#2012 := (>= #2010 0::int) +#1997 := (iff #2012 #2005) +#2006 := [rewrite]: #1997 +#2014 := (iff #2035 #2012) +#2011 := (= #2034 #2010) +#2013 := [rewrite]: #2011 +#2003 := [monotonicity #2013]: #2014 +#1998 := [trans #2003 #2006]: #2007 +#2024 := (iff #2037 #2022) +#2038 := (+ #1024 #1298) +#2018 := (<= #2038 0::int) +#2023 := (iff #2018 #2022) +#2016 := [rewrite]: #2023 +#2019 := (iff #2037 #2018) +#2015 := (= #2036 #2038) +#2017 := [rewrite]: #2015 +#2020 := [monotonicity #2017]: #2019 +#2009 := [trans #2020 #2016]: #2024 +#1990 := [monotonicity #2009 #1998]: #1989 +#1984 := [trans #1990 #1992]: #1993 +#1968 := [monotonicity #1984]: #1955 +#1972 := [trans #1968 #1970]: #1971 +#1963 := [quant-inst]: #1961 +#1962 := [mp #1963 #1972]: #1994 +#1914 := [unit-resolution #1962 #1912 #1908]: #1913 +#1907 := [unit-resolution #1914 #1905 #1917]: false +#1900 := [lemma #1907]: #1915 +#2470 := [unit-resolution #1900 #2469 #2465]: #1795 +#2121 := (or #2372 #2364 #1798) +#2136 := [def-axiom]: #2121 +#2471 := [unit-resolution #2136 #2470 #2464]: #2364 +#2235 := (not #2364) +#2444 := (or #2235 #2441) +#2427 := (= #2426 uf_12) +#2428 := (or #2427 #2425 #2423) +#2429 := (not #2428) +#2445 := (or #2235 #2429) +#2447 := (iff #2445 #2444) +#2449 := (iff #2444 #2444) +#2450 := [rewrite]: #2449 +#2442 := (iff #2429 #2441) +#2439 := (iff #2428 #2436) +#2433 := (or #2430 #2425 #2423) +#2437 := (iff #2433 #2436) +#2438 := [rewrite]: #2437 +#2434 := (iff #2428 #2433) +#2431 := (iff #2427 #2430) +#2432 := [rewrite]: #2431 +#2435 := [monotonicity #2432]: #2434 +#2440 := [trans #2435 #2438]: #2439 +#2443 := [monotonicity #2440]: #2442 +#2448 := [monotonicity #2443]: #2447 +#2451 := [trans #2448 #2450]: #2447 +#2446 := [quant-inst]: #2445 +#2452 := [mp #2446 #2451]: #2444 +#2472 := [unit-resolution #2452 #2471]: #2441 +#2459 := (or #2436 #2458) +#2460 := [def-axiom]: #2459 +#2473 := [unit-resolution #2460 #2472]: #2458 +#2485 := [unit-resolution #2473 #2484]: false +#2486 := [lemma #2485]: #2375 +#2231 := (or #2384 #2361 #2378) +#2220 := [def-axiom]: #2231 +#2415 := [unit-resolution #2220 #2486 #2414]: #2361 +#2106 := (or #2358 #2352) +#2248 := [def-axiom]: #2106 +#2416 := [unit-resolution #2248 #2415]: #2352 +#2417 := [hypothesis]: #840 +#2285 := (or #2340 #837) +#1923 := [def-axiom]: #2285 +#2418 := [unit-resolution #1923 #2417]: #2340 +#1987 := (or #2346 #837) +#1988 := [def-axiom]: #1987 +#2419 := [unit-resolution #1988 #2417]: #2346 +#2255 := (or #2355 #2343 #2349) +#2256 := [def-axiom]: #2255 +#2420 := [unit-resolution #2256 #2419 #2418 #2416]: false +#2403 := [lemma #2420]: #837 +#2690 := (or #840 #1977) +#2691 := [th-lemma]: #2690 +#2692 := [unit-resolution #2691 #2403]: #1977 +#2661 := [hypothesis]: #2349 +#2272 := (or #2346 #361) +#2273 := [def-axiom]: #2272 +#2662 := [unit-resolution #2273 #2661]: #361 +#2629 := (= #58 #1195) +#2642 := (not #2629) +#2630 := (+ #58 #1381) +#2632 := (>= #2630 0::int) +#2636 := (not #2632) +#2402 := (+ #39 #799) +#2405 := (<= #2402 0::int) +#2404 := (= #39 uf_8) +#2665 := (= uf_10 uf_8) +#2000 := (or #2346 #82) +#2001 := [def-axiom]: #2000 +#2663 := [unit-resolution #2001 #2661]: #82 +#2666 := [symm #2663]: #2665 +#2002 := (or #2346 #356) +#1896 := [def-axiom]: #2002 +#2664 := [unit-resolution #1896 #2661]: #356 +#2667 := [trans #2664 #2666]: #2404 +#2668 := (not #2404) +#2669 := (or #2668 #2405) +#2670 := [th-lemma]: #2669 +#2671 := [unit-resolution #2670 #2667]: #2405 +#1966 := (not #1383) +#1982 := (or #2346 #2334) +#2264 := [def-axiom]: #1982 +#2672 := [unit-resolution #2264 #2661]: #2334 +#2537 := (= #39 #58) +#2674 := (= #58 #39) +#2673 := [symm #2662]: #81 +#2675 := [monotonicity #2673]: #2674 +#2677 := [symm #2675]: #2537 +#2678 := (= uf_8 #39) +#2676 := [symm #2664]: #79 +#2679 := [trans #2663 #2676]: #2678 +#2680 := [trans #2679 #2677]: #221 +#1981 := (or #2328 #1204) +#1960 := [def-axiom]: #1981 +#2681 := [unit-resolution #1960 #2680]: #2328 +#1953 := (or #2337 #2331 #1645) +#2294 := [def-axiom]: #1953 +#2682 := [unit-resolution #2294 #2681 #2672]: #1645 +#2298 := (or #1640 #1966) +#2299 := [def-axiom]: #2298 +#2683 := [unit-resolution #2299 #2682]: #1966 +#2033 := (* -1::int #58) +#2538 := (+ #39 #2033) +#2540 := (>= #2538 0::int) +#2684 := (not #2537) +#2685 := (or #2684 #2540) +#2686 := [th-lemma]: #2685 +#2687 := [unit-resolution #2686 #2677]: #2540 +#2617 := (not #2405) +#2637 := (not #2540) +#2638 := (or #2636 #2637 #1383 #2617) +#2633 := [hypothesis]: #2632 +#2609 := [hypothesis]: #2405 +#2610 := [hypothesis]: #1966 +#2634 := [hypothesis]: #2540 +#2635 := [th-lemma #2634 #2610 #2609 #2633]: false +#2639 := [lemma #2635]: #2638 +#2688 := [unit-resolution #2639 #2687 #2683 #2671]: #2636 +#2643 := (or #2642 #2632) +#2644 := [th-lemma]: #2643 +#2689 := [unit-resolution #2644 #2688]: #2642 +#2300 := (or #1640 #1401) +#2301 := [def-axiom]: #2300 +#2693 := [unit-resolution #2301 #2682]: #1401 +#2694 := (not #1977) +#2695 := (or #2628 #1396 #2694) +#2696 := [th-lemma]: #2695 +#2697 := [unit-resolution #2696 #2693 #2692]: #2628 +#2565 := (<= #2564 0::int) +#2552 := (+ uf_6 #1381) +#2553 := (>= #2552 0::int) +#2699 := (not #2553) +#2266 := (or #2346 #854) +#2267 := [def-axiom]: #2266 +#2698 := [unit-resolution #2267 #2661]: #854 +#2700 := (or #2699 #1383 #2617 #850) +#2701 := [th-lemma]: #2700 +#2702 := [unit-resolution #2701 #2683 #2671 #2698]: #2699 +#2704 := (or #2553 #2565) +#2291 := (or #1640 #1192) +#1965 := [def-axiom]: #2291 +#2703 := [unit-resolution #1965 #2682]: #1192 +#2573 := (or #2317 #1625 #2553 #2565) +#2541 := (+ ?x3!1 #924) +#2542 := (>= #2541 0::int) +#2543 := (+ #1195 #1024) +#2544 := (<= #2543 0::int) +#2545 := (or #1625 #2544 #2542) +#2574 := (or #2317 #2545) +#2581 := (iff #2574 #2573) +#2570 := (or #1625 #2553 #2565) +#2576 := (or #2317 #2570) +#2579 := (iff #2576 #2573) +#2580 := [rewrite]: #2579 +#2577 := (iff #2574 #2576) +#2571 := (iff #2545 #2570) +#2568 := (iff #2542 #2565) +#2558 := (+ #924 ?x3!1) +#2561 := (>= #2558 0::int) +#2566 := (iff #2561 #2565) +#2567 := [rewrite]: #2566 +#2562 := (iff #2542 #2561) +#2559 := (= #2541 #2558) +#2560 := [rewrite]: #2559 +#2563 := [monotonicity #2560]: #2562 +#2569 := [trans #2563 #2567]: #2568 +#2556 := (iff #2544 #2553) +#2546 := (+ #1024 #1195) +#2549 := (<= #2546 0::int) +#2554 := (iff #2549 #2553) +#2555 := [rewrite]: #2554 +#2550 := (iff #2544 #2549) +#2547 := (= #2543 #2546) +#2548 := [rewrite]: #2547 +#2551 := [monotonicity #2548]: #2550 +#2557 := [trans #2551 #2555]: #2556 +#2572 := [monotonicity #2557 #2569]: #2571 +#2578 := [monotonicity #2572]: #2577 +#2582 := [trans #2578 #2580]: #2581 +#2575 := [quant-inst]: #2574 +#2583 := [mp #2575 #2582]: #2573 +#2705 := [unit-resolution #2583 #1912 #2703]: #2704 +#2706 := [unit-resolution #2705 #2702]: #2565 +#2708 := (not #2628) +#2707 := (not #2565) +#2709 := (or #2631 #2707 #2708) +#2710 := [th-lemma]: #2709 +#2711 := [unit-resolution #2710 #2706 #2697]: #2631 +#2658 := (not #2631) +#2659 := (or #2658 #2629 #376) +#2654 := (= #1195 #58) +#2652 := (= ?x3!1 uf_7) +#2648 := [hypothesis]: #361 +#2650 := (= ?x3!1 uf_4) +#2649 := [hypothesis]: #2631 +#2651 := [symm #2649]: #2650 +#2653 := [trans #2651 #2648]: #2652 +#2655 := [monotonicity #2653]: #2654 +#2656 := [symm #2655]: #2629 +#2647 := [hypothesis]: #2642 +#2657 := [unit-resolution #2647 #2656]: false +#2660 := [lemma #2657]: #2659 +#2712 := [unit-resolution #2660 #2711 #2689 #2662]: false +#2713 := [lemma #2712]: #2346 +#2766 := [unit-resolution #2256 #2713 #2416]: #2343 +#1928 := (or #2340 #2334) +#1929 := [def-axiom]: #1928 +#2767 := [unit-resolution #1929 #2766]: #2334 +#2515 := (= #36 #58) +#2772 := (= #58 #36) +#1937 := (or #2340 #191) +#2278 := [def-axiom]: #1937 +#2768 := [unit-resolution #2278 #2766]: #191 +#2769 := [symm #2768]: #42 +#2773 := [monotonicity #2769]: #2772 +#2774 := [symm #2773]: #2515 +#2775 := (= uf_8 #36) +#1941 := (or #2340 #194) +#1942 := [def-axiom]: #1941 +#2770 := [unit-resolution #1942 #2766]: #194 +#2771 := [symm #2770]: #44 +#2776 := [trans #2771 #2476]: #2775 +#2777 := [trans #2776 #2774]: #221 +#2778 := [unit-resolution #1960 #2777]: #2328 +#2779 := [unit-resolution #2294 #2778 #2767]: #1645 +#2780 := [unit-resolution #2301 #2779]: #1401 +#2781 := [unit-resolution #2696 #2780 #2692]: #2628 +#2510 := (+ uf_6 #799) +#2511 := (<= #2510 0::int) +#2782 := (or #301 #2511) +#2783 := [th-lemma]: #2782 +#2784 := [unit-resolution #2783 #2770]: #2511 +#2785 := [unit-resolution #2299 #2779]: #1966 +#2786 := (not #2511) +#2787 := (or #2699 #1383 #2786) +#2788 := [th-lemma]: #2787 +#2789 := [unit-resolution #2788 #2785 #2784]: #2699 +#2790 := [unit-resolution #1965 #2779]: #1192 +#2791 := [unit-resolution #2583 #1912 #2790 #2789]: #2565 +#2792 := [unit-resolution #2710 #2791 #2781]: #2631 +#2793 := [monotonicity #2792]: #2762 +#2794 := (not #2762) +#2795 := (or #2794 #2765) +#2796 := [th-lemma]: #2795 +#2797 := [unit-resolution #2796 #2793]: #2765 +#2286 := (or #2340 #850) +#2288 := [def-axiom]: #2286 +#2798 := [unit-resolution #2288 #2766]: #850 +[th-lemma #2798 #2785 #2784 #2797]: false +unsat diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/Boogie/Examples/cert/VCC_b_maximum --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Boogie/Examples/cert/VCC_b_maximum Wed Nov 04 11:40:59 2009 +0100 @@ -0,0 +1,687 @@ +(benchmark Isabelle +:extrasorts ( T2 T5 T8 T3 T15 T16 T4 T1 T6 T17 T11 T18 T7 T9 T13 T14 T12 T10 T19) +:extrafuns ( + (uf_9 T2) + (uf_48 T5 T3 T2) + (uf_26 T5) + (uf_72 T3 Int Int Int) + (uf_126 T5 T15 T5) + (uf_66 T5 Int T3 T5) + (uf_43 T3 Int T5) + (uf_116 T5 Int) + (uf_15 T5 T3) + (uf_81 Int Int Int) + (uf_80 Int Int Int) + (uf_70 T3 Int Int Int) + (uf_69 Int Int Int) + (uf_73 T3 Int Int) + (uf_101 T3 Int Int Int) + (uf_100 Int Int Int) + (uf_71 T3 Int Int Int) + (uf_46 T4 T4 T5 T3 T2) + (uf_121 T5) + (uf_53 T4 T5 T6) + (uf_163 T5 T6) + (uf_79 Int Int) + (uf_124 T3 Int T3) + (uf_259 T3 T3 T3) + (uf_25 T4 T5 T5) + (uf_27 T4 T5 T2) + (uf_255 T3) + (uf_254 T3) + (uf_94 T3) + (uf_90 T3) + (uf_87 T3) + (uf_83 T3) + (uf_7 T3) + (uf_91 T3) + (uf_4 T3) + (uf_84 T3) + (uf_24 T4 T5 T2) + (uf_10 T4 T5 T6) + (uf_128 T4 T5 T6) + (uf_253 Int) + (uf_20 T4 T9) + (uf_6 T3 T3) + (uf_224 T17 T17 T2) + (uf_153 T6 T6 T2) + (uf_13 T5 T6 T2) + (uf_138 T3 Int) + (uf_136 T14 T5) + (uf_5 T3) + (uf_291 T1) + (uf_122 T2 T2) + (uf_207 T4 T4 T5 T5 T2) + (uf_14 T3 T8) + (uf_61 T4 T5 T2) + (uf_114 T4 T5 Int) + (uf_113 T4 T5 Int) + (uf_112 T4 T5 Int) + (uf_111 T4 T5 Int) + (uf_110 T4 T5 Int) + (uf_109 T4 T5 Int) + (uf_108 T4 T5 Int) + (uf_107 T4 T5 Int) + (uf_38 T4 T5 Int) + (uf_169 T4 T4 T5 T5 T4) + (uf_145 T5 T6 T2) + (uf_147 T5 T6 T2) + (uf_59 T4 T13) + (uf_232 T4 T5 T18) + (uf_258 T3) + (uf_240 T3) + (uf_284 T16) + (uf_188 T4 T5 T5 T5 T5) + (uf_65 T4 T5 T3 Int T2) + (uf_95 Int) + (uf_92 Int) + (uf_88 Int) + (uf_85 Int) + (uf_78 Int) + (uf_77 Int) + (uf_76 Int) + (uf_75 Int) + (uf_96 Int) + (uf_93 Int) + (uf_89 Int) + (uf_86 Int) + (uf_42 T5) + (uf_230 T17) + (uf_173 T4 T5 T5 T11) + (uf_215 T11 T5) + (uf_266 T3 T3) + (uf_233 T18 T4) + (uf_37 T3) + (uf_279 T1) + (uf_281 T1) + (uf_287 T1) + (uf_99 Int Int Int Int Int Int) + (uf_55 T4 T2) + (uf_60 Int T3 T5) + (uf_246 Int T5) + (uf_220 T5 T15 Int) + (uf_196 T4 T5 T5 T2) + (uf_264 T3 T3) + (uf_142 T3 Int) + (uf_117 T5 Int) + (uf_119 T5 Int) + (uf_118 T5 Int) + (uf_120 T5 Int) + (uf_222 T17 T15 Int) + (uf_152 T6) + (uf_157 T6 T6 T6) + (uf_41 T4 T12) + (uf_174 T4 T5 T5 T4) + (uf_170 T4 T5 Int) + (uf_82 T3 Int Int) + (uf_106 T3 Int Int Int) + (uf_103 T3 Int Int Int) + (uf_102 T3 Int Int Int) + (uf_104 T3 Int Int Int) + (uf_105 T3 Int Int Int) + (uf_241 T15 Int T15) + (uf_50 T5 T5 T2) + (uf_245 Int T15) + (uf_51 T4 T2) + (uf_74 T3 Int T2) + (uf_195 T4 T5 T5 T2) + (uf_28 Int T5) + (uf_262 T8) + (uf_161 T5 Int T5) + (uf_265 T3 T3) + (uf_47 T4 T5 T2) + (uf_29 T5 Int) + (uf_201 T4 T5 T3 T5) + (uf_229 T17 T15 Int T17) + (uf_179 T4 T4 T5 T3 T2) + (uf_154 T6 T6 T2) + (uf_39 T11 Int) + (uf_172 T12 T5 T11 T12) + (uf_251 T13 T5 T14 T13) + (uf_175 T4 T5 T5 T11) + (uf_176 T4 T5 Int T4) + (uf_192 T7 T6) + (uf_257 T3) + (uf_132 T5 T3 Int T6) + (uf_139 T5 T5 T2) + (uf_276 T19 Int) + (uf_130 T5 T6) + (uf_44 T4 T2) + (uf_261 T8) + (uf_248 T3 T3 Int) + (uf_249 T3 T3 Int) + (uf_181 T4 T4 T2) + (uf_221 Int Int T2) + (uf_160 T5 Int T5) + (uf_40 T12 T5 T11) + (uf_58 T13 T5 T14) + (uf_178 T9 T5 Int T9) + (uf_235 T18) + (uf_49 T4 T5 T2) + (uf_234 T18 Int) + (uf_267 T3) + (uf_143 T3 Int) + (uf_243 T15 T15) + (uf_242 T15 Int) + (uf_54 T5 T5 T2) + (uf_144 T3 T3) + (uf_237 T15 Int) + (uf_148 T5 T2) + (uf_283 Int T5 T2) + (uf_125 T5 T5 Int) + (uf_141 T3 T2) + (uf_260 T3 T2) + (uf_57 T3 T2) + (uf_23 T3 T2) + (uf_159 T5 T5 T5) + (uf_12 T4 T5 T7) + (uf_19 T9 T5 Int) + (uf_131 T6 T6 T2) + (uf_149 T6) + (uf_217 T11 Int) + (uf_67 T4 T5 T2) + (uf_219 T3) + (uf_268 T3) + (uf_289 T1) + (uf_134 T5 T3 Int T6) + (uf_189 T5 T7) + (uf_183 T10 T5 Int) + (uf_62 Int Int) + (uf_63 Int Int) + (uf_200 T4 T5 T5 T16 T2) + (uf_140 T5 T3 T5) + (uf_34 Int T6) + (uf_225 Int T17) + (uf_56 T3 T2) + (uf_208 T3 T2) + (uf_35 T6 Int) + (uf_231 T17 T15 Int Int Int Int T17) + (uf_226 T17 Int) + (uf_151 T5 T6) + (uf_162 T4 T5 T6) + (uf_256 T3) + (uf_45 T4 T5 T2) + (uf_203 T4 T2) + (uf_202 T1 T4 T2) + (uf_198 T4 T5 T5 T16 T2) + (uf_32 Int T7) + (uf_185 T3 T15 T15 T2) + (uf_211 T4 T5 T2) + (uf_228 T3 T2) + (uf_263 T8) + (uf_16 T8) + (uf_214 T3 T15) + (uf_156 T6 T6 T6) + (uf_206 T4 T4 T5 T3 T2) + (uf_135 T14 T2) + (uf_33 T7 Int) + (uf_275 T1) + (uf_177 T4 T4 T2) + (uf_133 T5 T6 T6 Int) + (uf_186 T5 T5 T2) + (uf_247 T3 T3 Int Int T2) + (uf_227 T3 T15 T3 T2) + (uf_127 T4 T5 T6) + (uf_150 T6 Int) + (uf_286 T1) + (uf_288 T1) + (uf_295 T1) + (uf_290 T1) + (uf_305 T1) + (uf_18 T5 T2) + (uf_22 T3 T2) + (uf_184 T4 T5 T10) + (uf_155 T6 T6 T6) + (uf_303 T1) + (uf_306 T1) + (uf_97 Int Int Int Int Int) + (uf_236 T5 T15 T5) + (uf_171 T4 Int) + (uf_8 T4 T4 T5 T6 T2) + (uf_11 T7 T5 Int) + (uf_238 T15 T3) + (uf_210 T4 T5 T2) + (uf_180 T3 T15 T2) + (uf_252 T3) + (uf_64 Int Int T5) + (uf_30 Int T10) + (uf_31 T10 Int) + (uf_98 Int Int Int Int Int) + (uf_277 Int) + (uf_164 T4 T2) + (uf_21 T4 T4 T6 T2) + (uf_115 T5 T5 Int) + (uf_167 T5) + (uf_168 Int) + (uf_129 T5 T3 Int T6) + (uf_123 T4 T4 T5 T3 T2) + (uf_17 T4 T4 T6 T2) + (uf_239 T5 T15 Int) + (uf_166 T3) + (uf_223 T15 T15) + (uf_191 T4 T2) + (uf_137 T4 T5 T3 Int T2 T2) + (uf_158 T5 T6) + (uf_204 T4 T4 T5 T3 T2) + (uf_187 T15 Int T2) + (uf_190 T15 T2) + (uf_2 T1) + (uf_194 T15 Int T3 T2) + (uf_273 T4) + (uf_270 Int) + (uf_269 Int) + (uf_274 Int) + (uf_272 Int) + (uf_294 Int) + (uf_302 Int) + (uf_297 Int) + (uf_285 Int) + (uf_292 Int) + (uf_304 Int) + (uf_300 Int) + (uf_296 Int) + (uf_271 Int) + (uf_299 Int) + (uf_293 Int) + (uf_301 Int) + (uf_298 Int) + (uf_282 Int) + ) +:extrapreds ( + (up_199 T4 T5 T16) + (up_146 T5 T6) + (up_213 T14) + (up_209 T4 T5 T3) + (up_250 T3 T3) + (up_218 T11) + (up_1 Int T1) + (up_36 T3) + (up_3 Int T3) + (up_244 T15) + (up_212 T11) + (up_280 T4 T1 T1 Int T3) + (up_182 Int) + (up_216) + (up_68 T14) + (up_193 T2) + (up_52 T6) + (up_278 T4 T1 T1 T5 T3) + (up_197 T3) + (up_165 T4) + (up_205 T4 T4 T5 T3) + ) +:assumption (up_1 1 uf_2) +:assumption (up_3 1 uf_4) +:assumption (= uf_5 (uf_6 uf_7)) +:assumption (forall (?x1 T4) (?x2 T4) (?x3 T5) (?x4 T6) (iff (= (uf_8 ?x1 ?x2 ?x3 ?x4) uf_9) (and (= (uf_10 ?x1 ?x3) (uf_10 ?x2 ?x3)) (forall (?x5 T5) (implies (and (not (= (uf_13 ?x5 ?x4) uf_9)) (= (uf_14 (uf_15 ?x5)) uf_16)) (= (uf_11 (uf_12 ?x1 ?x3) ?x5) (uf_11 (uf_12 ?x2 ?x3) ?x5))) :pat { (uf_11 (uf_12 ?x2 ?x3) ?x5) }))) :pat { (uf_8 ?x1 ?x2 ?x3 ?x4) }) +:assumption (forall (?x6 T4) (?x7 T4) (?x8 T6) (implies (forall (?x9 T5) (implies (and (not (= (uf_14 (uf_15 ?x9)) uf_16)) (= (uf_13 ?x9 ?x8) uf_9)) (or (= (uf_8 ?x6 ?x7 ?x9 ?x8) uf_9) (= (uf_19 (uf_20 ?x6) ?x9) (uf_19 (uf_20 ?x7) ?x9)))) :pat { (uf_18 ?x9) }) (= (uf_17 ?x6 ?x7 ?x8) uf_9)) :pat { (uf_17 ?x6 ?x7 ?x8) }) +:assumption (forall (?x10 T4) (?x11 T4) (?x12 T6) (implies (forall (?x13 T5) (implies (or (= (uf_22 (uf_15 ?x13)) uf_9) (= (uf_23 (uf_15 ?x13)) uf_9)) (implies (and (not (or (and (= (uf_24 ?x10 ?x13) uf_9) (= (uf_14 (uf_15 ?x13)) uf_16)) (not (= (uf_25 ?x10 ?x13) uf_26)))) (= (uf_27 ?x10 ?x13) uf_9)) (or (= (uf_13 ?x13 ?x12) uf_9) (= (uf_19 (uf_20 ?x10) ?x13) (uf_19 (uf_20 ?x11) ?x13))))) :pat { (uf_18 ?x13) }) (= (uf_21 ?x10 ?x11 ?x12) uf_9)) :pat { (uf_21 ?x10 ?x11 ?x12) }) +:assumption (forall (?x14 T5) (= (uf_28 (uf_29 ?x14)) ?x14)) +:assumption (forall (?x15 T10) (= (uf_30 (uf_31 ?x15)) ?x15)) +:assumption (forall (?x16 T7) (= (uf_32 (uf_33 ?x16)) ?x16)) +:assumption (forall (?x17 T6) (= (uf_34 (uf_35 ?x17)) ?x17)) +:assumption (up_36 uf_37) +:assumption (forall (?x18 T4) (?x19 T5) (= (uf_38 ?x18 ?x19) (uf_39 (uf_40 (uf_41 ?x18) ?x19))) :pat { (uf_38 ?x18 ?x19) }) +:assumption (= uf_42 (uf_43 uf_37 0)) +:assumption (forall (?x20 T4) (?x21 T5) (implies (and (= (uf_45 ?x20 ?x21) uf_9) (= (uf_44 ?x20) uf_9)) (= (uf_46 ?x20 ?x20 ?x21 (uf_15 ?x21)) uf_9)) :pat { (uf_44 ?x20) (uf_45 ?x20 ?x21) }) +:assumption (forall (?x22 T4) (?x23 T5) (iff (= (uf_45 ?x22 ?x23) uf_9) (= (uf_24 ?x22 ?x23) uf_9)) :pat { (uf_45 ?x22 ?x23) }) +:assumption (forall (?x24 T4) (?x25 T5) (iff (= (uf_47 ?x24 ?x25) uf_9) (and (or (= (uf_38 ?x24 ?x25) 0) (not (up_36 (uf_15 ?x25)))) (and (= (uf_22 (uf_15 ?x25)) uf_9) (and (not (= (uf_14 (uf_15 ?x25)) uf_16)) (and (= (uf_27 ?x24 ?x25) uf_9) (and (= (uf_48 ?x25 (uf_15 ?x25)) uf_9) (and (= (uf_25 ?x24 ?x25) uf_26) (= (uf_24 ?x24 ?x25) uf_9)))))))) :pat { (uf_47 ?x24 ?x25) }) +:assumption (forall (?x26 T4) (?x27 T5) (?x28 Int) (implies (and (= (uf_50 ?x27 (uf_43 uf_37 ?x28)) uf_9) (= (uf_49 ?x26 ?x27) uf_9)) (= (uf_49 ?x26 (uf_43 uf_37 ?x28)) uf_9)) :pat { (uf_49 ?x26 ?x27) (uf_50 ?x27 (uf_43 uf_37 ?x28)) }) +:assumption (forall (?x29 T4) (?x30 T5) (?x31 T5) (implies (and (= (uf_50 ?x30 ?x31) uf_9) (= (uf_49 ?x29 ?x30) uf_9)) (= (uf_46 ?x29 ?x29 ?x31 (uf_15 ?x31)) uf_9)) :pat { (uf_49 ?x29 ?x30) (uf_50 ?x30 ?x31) }) +:assumption (forall (?x32 T4) (?x33 T5) (?x34 T5) (implies (= (uf_51 ?x32) uf_9) (implies (and (= (uf_24 ?x32 ?x33) uf_9) (= (uf_50 ?x33 ?x34) uf_9)) (and (< 0 (uf_38 ?x32 ?x34)) (and (= (uf_24 ?x32 ?x34) uf_9) (up_52 (uf_53 ?x32 ?x34)))))) :pat { (uf_24 ?x32 ?x33) (uf_50 ?x33 ?x34) }) +:assumption (forall (?x35 T4) (?x36 T5) (?x37 T5) (implies (and (= (uf_54 ?x36 ?x37) uf_9) (= (uf_49 ?x35 ?x36) uf_9)) (= (uf_49 ?x35 ?x37) uf_9)) :pat { (uf_49 ?x35 ?x36) (uf_54 ?x36 ?x37) }) +:assumption (forall (?x38 T5) (?x39 T5) (implies (and (forall (?x40 T4) (implies (= (uf_49 ?x40 ?x38) uf_9) (= (uf_24 ?x40 ?x39) uf_9))) (and (= (uf_48 ?x39 uf_37) uf_9) (= (uf_48 ?x38 uf_37) uf_9))) (= (uf_54 ?x38 ?x39) uf_9)) :pat { (uf_54 ?x38 ?x39) }) +:assumption (forall (?x41 T4) (?x42 T5) (implies (= (uf_49 ?x41 ?x42) uf_9) (and (= (uf_44 ?x41) uf_9) (= (uf_24 ?x41 ?x42) uf_9))) :pat { (uf_49 ?x41 ?x42) }) +:assumption (forall (?x43 T4) (?x44 T5) (implies (and (= (uf_24 ?x43 ?x44) uf_9) (= (uf_55 ?x43) uf_9)) (= (uf_49 ?x43 ?x44) uf_9)) :pat { (uf_55 ?x43) (uf_49 ?x43 ?x44) }) +:assumption (forall (?x45 T3) (implies (= (uf_56 ?x45) uf_9) (= (uf_23 ?x45) uf_9)) :pat { (uf_56 ?x45) }) +:assumption (forall (?x46 T3) (implies (= (uf_57 ?x46) uf_9) (= (uf_23 ?x46) uf_9)) :pat { (uf_57 ?x46) }) +:assumption (forall (?x47 T4) (?x48 Int) (?x49 T3) (implies (and (= (uf_51 ?x47) uf_9) (= (uf_56 ?x49) uf_9)) (= (uf_61 ?x47 (uf_60 ?x48 ?x49)) uf_9)) :pat { (uf_58 (uf_59 ?x47) (uf_60 ?x48 ?x49)) } :pat { (uf_40 (uf_41 ?x47) (uf_60 ?x48 ?x49)) }) +:assumption (forall (?x50 Int) (= (uf_62 (uf_63 ?x50)) ?x50)) +:assumption (forall (?x51 Int) (?x52 T3) (= (uf_60 ?x51 ?x52) (uf_43 ?x52 (uf_63 ?x51))) :pat { (uf_60 ?x51 ?x52) }) +:assumption (forall (?x53 Int) (?x54 Int) (?x55 T4) (implies (= (uf_51 ?x55) uf_9) (and (forall (?x56 Int) (implies (and (< ?x56 ?x54) (<= 0 ?x56)) (and (= (uf_67 ?x55 (uf_66 (uf_64 ?x53 ?x54) ?x56 uf_7)) uf_9) (and (= (uf_48 (uf_66 (uf_64 ?x53 ?x54) ?x56 uf_7) uf_7) uf_9) (up_68 (uf_58 (uf_59 ?x55) (uf_66 (uf_64 ?x53 ?x54) ?x56 uf_7)))))) :pat { (uf_40 (uf_41 ?x55) (uf_66 (uf_64 ?x53 ?x54) ?x56 uf_7)) } :pat { (uf_58 (uf_59 ?x55) (uf_66 (uf_64 ?x53 ?x54) ?x56 uf_7)) }) (= (uf_27 ?x55 (uf_64 ?x53 ?x54)) uf_9))) :pat { (uf_27 ?x55 (uf_64 ?x53 ?x54)) } :pat { (uf_65 ?x55 (uf_64 ?x53 ?x54) uf_7 ?x54) }) +:assumption (forall (?x57 Int) (?x58 Int) (= (uf_48 (uf_64 ?x57 ?x58) uf_7) uf_9) :pat { (uf_64 ?x57 ?x58) }) +:assumption (forall (?x59 Int) (?x60 Int) (= (uf_69 ?x59 ?x60) (+ ?x59 ?x60)) :pat { (uf_69 ?x59 ?x60) }) +:assumption (forall (?x61 T3) (?x62 Int) (?x63 Int) (= (uf_70 ?x61 ?x62 ?x63) (uf_70 ?x61 ?x63 ?x62)) :pat { (uf_70 ?x61 ?x62 ?x63) }) +:assumption (forall (?x64 T3) (?x65 Int) (?x66 Int) (= (uf_71 ?x64 ?x65 ?x66) (uf_71 ?x64 ?x66 ?x65)) :pat { (uf_71 ?x64 ?x65 ?x66) }) +:assumption (forall (?x67 T3) (?x68 Int) (?x69 Int) (= (uf_72 ?x67 ?x68 ?x69) (uf_72 ?x67 ?x69 ?x68)) :pat { (uf_72 ?x67 ?x68 ?x69) }) +:assumption (forall (?x70 T3) (?x71 Int) (implies (= (uf_74 ?x70 ?x71) uf_9) (= (uf_73 ?x70 (uf_73 ?x70 ?x71)) ?x71)) :pat { (uf_73 ?x70 (uf_73 ?x70 ?x71)) }) +:assumption (forall (?x72 T3) (?x73 Int) (= (uf_71 ?x72 ?x73 (uf_73 ?x72 0)) (uf_73 ?x72 ?x73)) :pat { (uf_71 ?x72 ?x73 (uf_73 ?x72 0)) }) +:assumption (forall (?x74 T3) (?x75 Int) (= (uf_71 ?x74 ?x75 ?x75) 0) :pat { (uf_71 ?x74 ?x75 ?x75) }) +:assumption (forall (?x76 T3) (?x77 Int) (implies (= (uf_74 ?x76 ?x77) uf_9) (= (uf_71 ?x76 ?x77 0) ?x77)) :pat { (uf_71 ?x76 ?x77 0) }) +:assumption (forall (?x78 T3) (?x79 Int) (?x80 Int) (= (uf_70 ?x78 (uf_72 ?x78 ?x79 ?x80) ?x79) ?x79) :pat { (uf_70 ?x78 (uf_72 ?x78 ?x79 ?x80) ?x79) }) +:assumption (forall (?x81 T3) (?x82 Int) (?x83 Int) (= (uf_70 ?x81 (uf_72 ?x81 ?x82 ?x83) ?x83) ?x83) :pat { (uf_70 ?x81 (uf_72 ?x81 ?x82 ?x83) ?x83) }) +:assumption (forall (?x84 T3) (?x85 Int) (implies (= (uf_74 ?x84 ?x85) uf_9) (= (uf_70 ?x84 ?x85 ?x85) ?x85)) :pat { (uf_70 ?x84 ?x85 ?x85) }) +:assumption (forall (?x86 T3) (?x87 Int) (implies (= (uf_74 ?x86 ?x87) uf_9) (= (uf_70 ?x86 ?x87 (uf_73 ?x86 0)) ?x87)) :pat { (uf_70 ?x86 ?x87 (uf_73 ?x86 0)) }) +:assumption (forall (?x88 T3) (?x89 Int) (= (uf_70 ?x88 ?x89 0) 0) :pat { (uf_70 ?x88 ?x89 0) }) +:assumption (forall (?x90 T3) (?x91 Int) (implies (= (uf_74 ?x90 ?x91) uf_9) (= (uf_72 ?x90 ?x91 ?x91) ?x91)) :pat { (uf_72 ?x90 ?x91 ?x91) }) +:assumption (forall (?x92 T3) (?x93 Int) (= (uf_72 ?x92 ?x93 (uf_73 ?x92 0)) (uf_73 ?x92 0)) :pat { (uf_72 ?x92 ?x93 (uf_73 ?x92 0)) }) +:assumption (forall (?x94 T3) (?x95 Int) (implies (= (uf_74 ?x94 ?x95) uf_9) (= (uf_72 ?x94 ?x95 0) ?x95)) :pat { (uf_72 ?x94 ?x95 0) }) +:assumption (forall (?x96 T3) (?x97 Int) (= (uf_70 ?x96 ?x97 (uf_73 ?x96 ?x97)) 0) :pat { (uf_70 ?x96 ?x97 (uf_73 ?x96 ?x97)) }) +:assumption (forall (?x98 T3) (?x99 Int) (= (uf_72 ?x98 ?x99 (uf_73 ?x98 ?x99)) (uf_73 ?x98 0)) :pat { (uf_72 ?x98 ?x99 (uf_73 ?x98 ?x99)) }) +:assumption (forall (?x100 T3) (?x101 Int) (= (uf_74 ?x100 (uf_73 ?x100 ?x101)) uf_9) :pat { (uf_73 ?x100 ?x101) }) +:assumption (forall (?x102 T3) (?x103 Int) (?x104 Int) (implies (and (<= ?x104 uf_75) (and (<= 0 ?x104) (and (<= ?x103 uf_75) (<= 0 ?x103)))) (and (<= (uf_71 ?x102 ?x103 ?x104) uf_75) (<= 0 (uf_71 ?x102 ?x103 ?x104)))) :pat { (uf_71 ?x102 ?x103 ?x104) }) +:assumption (forall (?x105 T3) (?x106 Int) (?x107 Int) (implies (and (<= ?x107 uf_76) (and (<= 0 ?x107) (and (<= ?x106 uf_76) (<= 0 ?x106)))) (and (<= (uf_71 ?x105 ?x106 ?x107) uf_76) (<= 0 (uf_71 ?x105 ?x106 ?x107)))) :pat { (uf_71 ?x105 ?x106 ?x107) }) +:assumption (forall (?x108 T3) (?x109 Int) (?x110 Int) (implies (and (<= ?x110 uf_77) (and (<= 0 ?x110) (and (<= ?x109 uf_77) (<= 0 ?x109)))) (and (<= (uf_71 ?x108 ?x109 ?x110) uf_77) (<= 0 (uf_71 ?x108 ?x109 ?x110)))) :pat { (uf_71 ?x108 ?x109 ?x110) }) +:assumption (forall (?x111 T3) (?x112 Int) (?x113 Int) (implies (and (<= ?x113 uf_78) (and (<= 0 ?x113) (and (<= ?x112 uf_78) (<= 0 ?x112)))) (and (<= (uf_71 ?x111 ?x112 ?x113) uf_78) (<= 0 (uf_71 ?x111 ?x112 ?x113)))) :pat { (uf_71 ?x111 ?x112 ?x113) }) +:assumption (forall (?x114 T3) (?x115 Int) (?x116 Int) (implies (and (<= ?x116 uf_75) (and (<= 0 ?x116) (and (<= ?x115 uf_75) (<= 0 ?x115)))) (and (<= (uf_70 ?x114 ?x115 ?x116) uf_75) (<= 0 (uf_70 ?x114 ?x115 ?x116)))) :pat { (uf_70 ?x114 ?x115 ?x116) }) +:assumption (forall (?x117 T3) (?x118 Int) (?x119 Int) (implies (and (<= ?x119 uf_76) (and (<= 0 ?x119) (and (<= ?x118 uf_76) (<= 0 ?x118)))) (and (<= (uf_70 ?x117 ?x118 ?x119) uf_76) (<= 0 (uf_70 ?x117 ?x118 ?x119)))) :pat { (uf_70 ?x117 ?x118 ?x119) }) +:assumption (forall (?x120 T3) (?x121 Int) (?x122 Int) (implies (and (<= ?x122 uf_77) (and (<= 0 ?x122) (and (<= ?x121 uf_77) (<= 0 ?x121)))) (and (<= (uf_70 ?x120 ?x121 ?x122) uf_77) (<= 0 (uf_70 ?x120 ?x121 ?x122)))) :pat { (uf_70 ?x120 ?x121 ?x122) }) +:assumption (forall (?x123 T3) (?x124 Int) (?x125 Int) (implies (and (<= ?x125 uf_78) (and (<= 0 ?x125) (and (<= ?x124 uf_78) (<= 0 ?x124)))) (and (<= (uf_70 ?x123 ?x124 ?x125) uf_78) (<= 0 (uf_70 ?x123 ?x124 ?x125)))) :pat { (uf_70 ?x123 ?x124 ?x125) }) +:assumption (forall (?x126 T3) (?x127 Int) (?x128 Int) (implies (and (<= ?x128 uf_75) (and (<= 0 ?x128) (and (<= ?x127 uf_75) (<= 0 ?x127)))) (and (<= (uf_72 ?x126 ?x127 ?x128) uf_75) (<= 0 (uf_72 ?x126 ?x127 ?x128)))) :pat { (uf_72 ?x126 ?x127 ?x128) }) +:assumption (forall (?x129 T3) (?x130 Int) (?x131 Int) (implies (and (<= ?x131 uf_76) (and (<= 0 ?x131) (and (<= ?x130 uf_76) (<= 0 ?x130)))) (and (<= (uf_72 ?x129 ?x130 ?x131) uf_76) (<= 0 (uf_72 ?x129 ?x130 ?x131)))) :pat { (uf_72 ?x129 ?x130 ?x131) }) +:assumption (forall (?x132 T3) (?x133 Int) (?x134 Int) (implies (and (<= ?x134 uf_77) (and (<= 0 ?x134) (and (<= ?x133 uf_77) (<= 0 ?x133)))) (and (<= (uf_72 ?x132 ?x133 ?x134) uf_77) (<= 0 (uf_72 ?x132 ?x133 ?x134)))) :pat { (uf_72 ?x132 ?x133 ?x134) }) +:assumption (forall (?x135 T3) (?x136 Int) (?x137 Int) (implies (and (<= ?x137 uf_78) (and (<= 0 ?x137) (and (<= ?x136 uf_78) (<= 0 ?x136)))) (and (<= (uf_72 ?x135 ?x136 ?x137) uf_78) (<= 0 (uf_72 ?x135 ?x136 ?x137)))) :pat { (uf_72 ?x135 ?x136 ?x137) }) +:assumption (forall (?x138 T3) (?x139 Int) (?x140 Int) (?x141 Int) (implies (and (= (uf_74 ?x138 ?x140) uf_9) (and (= (uf_74 ?x138 ?x139) uf_9) (and (< ?x140 (uf_79 ?x141)) (and (< ?x139 (uf_79 ?x141)) (and (< ?x141 64) (and (<= 0 ?x141) (and (<= 0 ?x140) (<= 0 ?x139)))))))) (< (uf_72 ?x138 ?x139 ?x140) (uf_79 ?x141))) :pat { (uf_72 ?x138 ?x139 ?x140) (uf_79 ?x141) }) +:assumption (forall (?x142 T3) (?x143 Int) (?x144 Int) (implies (and (= (uf_74 ?x142 ?x144) uf_9) (and (= (uf_74 ?x142 ?x143) uf_9) (and (<= 0 ?x144) (<= 0 ?x143)))) (and (<= ?x144 (uf_72 ?x142 ?x143 ?x144)) (<= ?x143 (uf_72 ?x142 ?x143 ?x144)))) :pat { (uf_72 ?x142 ?x143 ?x144) }) +:assumption (forall (?x145 T3) (?x146 Int) (?x147 Int) (implies (and (= (uf_74 ?x145 ?x147) uf_9) (and (= (uf_74 ?x145 ?x146) uf_9) (and (<= 0 ?x147) (<= 0 ?x146)))) (and (<= (uf_72 ?x145 ?x146 ?x147) (+ ?x146 ?x147)) (<= 0 (uf_72 ?x145 ?x146 ?x147)))) :pat { (uf_72 ?x145 ?x146 ?x147) }) +:assumption (forall (?x148 T3) (?x149 Int) (?x150 Int) (implies (and (= (uf_74 ?x148 ?x150) uf_9) (and (= (uf_74 ?x148 ?x149) uf_9) (and (<= 0 ?x150) (<= 0 ?x149)))) (and (<= (uf_70 ?x148 ?x149 ?x150) ?x150) (<= (uf_70 ?x148 ?x149 ?x150) ?x149))) :pat { (uf_70 ?x148 ?x149 ?x150) }) +:assumption (forall (?x151 T3) (?x152 Int) (?x153 Int) (implies (and (= (uf_74 ?x151 ?x152) uf_9) (<= 0 ?x152)) (and (<= (uf_70 ?x151 ?x152 ?x153) ?x152) (<= 0 (uf_70 ?x151 ?x152 ?x153)))) :pat { (uf_70 ?x151 ?x152 ?x153) }) +:assumption (forall (?x154 Int) (?x155 Int) (implies (and (< ?x155 0) (<= ?x154 0)) (and (<= (uf_80 ?x154 ?x155) 0) (< ?x155 (uf_80 ?x154 ?x155)))) :pat { (uf_80 ?x154 ?x155) }) +:assumption (forall (?x156 Int) (?x157 Int) (implies (and (< 0 ?x157) (<= ?x156 0)) (and (<= (uf_80 ?x156 ?x157) 0) (< (+ 0 ?x157) (uf_80 ?x156 ?x157)))) :pat { (uf_80 ?x156 ?x157) }) +:assumption (forall (?x158 Int) (?x159 Int) (implies (and (< ?x159 0) (<= 0 ?x158)) (and (< (uf_80 ?x158 ?x159) (+ 0 ?x159)) (<= 0 (uf_80 ?x158 ?x159)))) :pat { (uf_80 ?x158 ?x159) }) +:assumption (forall (?x160 Int) (?x161 Int) (implies (and (< 0 ?x161) (<= 0 ?x160)) (and (< (uf_80 ?x160 ?x161) ?x161) (<= 0 (uf_80 ?x160 ?x161)))) :pat { (uf_80 ?x160 ?x161) }) +:assumption (forall (?x162 Int) (?x163 Int) (= (uf_80 ?x162 ?x163) (+ ?x162 (+ (uf_81 ?x162 ?x163) ?x163))) :pat { (uf_80 ?x162 ?x163) } :pat { (uf_81 ?x162 ?x163) }) +:assumption (forall (?x164 Int) (implies (not (= ?x164 0)) (= (uf_81 ?x164 ?x164) 1)) :pat { (uf_81 ?x164 ?x164) }) +:assumption (forall (?x165 Int) (?x166 Int) (implies (and (< 0 ?x166) (< 0 ?x165)) (and (<= (+ (uf_81 ?x165 ?x166) ?x166) ?x165) (< (+ ?x165 ?x166) (+ (uf_81 ?x165 ?x166) ?x166)))) :pat { (uf_81 ?x165 ?x166) }) +:assumption (forall (?x167 Int) (?x168 Int) (implies (and (< 0 ?x168) (<= 0 ?x167)) (<= (uf_81 ?x167 ?x168) ?x167)) :pat { (uf_81 ?x167 ?x168) }) +:assumption (forall (?x169 T3) (?x170 Int) (?x171 Int) (?x172 Int) (implies (and (<= 0 ?x170) (and (= (uf_74 ?x169 (+ (uf_79 ?x171) 1)) uf_9) (= (uf_74 ?x169 ?x170) uf_9))) (= (uf_80 ?x170 (uf_79 ?x171)) (uf_70 ?x169 ?x170 (+ (uf_79 ?x171) 1)))) :pat { (uf_80 ?x170 (uf_79 ?x171)) (uf_70 ?x169 ?x170 ?x172) }) +:assumption (forall (?x173 Int) (implies (and (<= ?x173 uf_85) (<= uf_86 ?x173)) (= (uf_82 uf_83 (uf_82 uf_84 ?x173)) ?x173)) :pat { (uf_82 uf_83 (uf_82 uf_84 ?x173)) }) +:assumption (forall (?x174 Int) (implies (and (<= ?x174 uf_88) (<= uf_89 ?x174)) (= (uf_82 uf_87 (uf_82 uf_4 ?x174)) ?x174)) :pat { (uf_82 uf_87 (uf_82 uf_4 ?x174)) }) +:assumption (forall (?x175 Int) (implies (and (<= ?x175 uf_92) (<= uf_93 ?x175)) (= (uf_82 uf_90 (uf_82 uf_91 ?x175)) ?x175)) :pat { (uf_82 uf_90 (uf_82 uf_91 ?x175)) }) +:assumption (forall (?x176 Int) (implies (and (<= ?x176 uf_95) (<= uf_96 ?x176)) (= (uf_82 uf_94 (uf_82 uf_7 ?x176)) ?x176)) :pat { (uf_82 uf_94 (uf_82 uf_7 ?x176)) }) +:assumption (forall (?x177 Int) (implies (and (<= ?x177 uf_75) (<= 0 ?x177)) (= (uf_82 uf_84 (uf_82 uf_83 ?x177)) ?x177)) :pat { (uf_82 uf_84 (uf_82 uf_83 ?x177)) }) +:assumption (forall (?x178 Int) (implies (and (<= ?x178 uf_76) (<= 0 ?x178)) (= (uf_82 uf_4 (uf_82 uf_87 ?x178)) ?x178)) :pat { (uf_82 uf_4 (uf_82 uf_87 ?x178)) }) +:assumption (forall (?x179 Int) (implies (and (<= ?x179 uf_77) (<= 0 ?x179)) (= (uf_82 uf_91 (uf_82 uf_90 ?x179)) ?x179)) :pat { (uf_82 uf_91 (uf_82 uf_90 ?x179)) }) +:assumption (forall (?x180 Int) (implies (and (<= ?x180 uf_78) (<= 0 ?x180)) (= (uf_82 uf_7 (uf_82 uf_94 ?x180)) ?x180)) :pat { (uf_82 uf_7 (uf_82 uf_94 ?x180)) }) +:assumption (forall (?x181 T3) (?x182 Int) (= (uf_74 ?x181 (uf_82 ?x181 ?x182)) uf_9) :pat { (uf_82 ?x181 ?x182) }) +:assumption (forall (?x183 T3) (?x184 Int) (implies (= (uf_74 ?x183 ?x184) uf_9) (= (uf_82 ?x183 ?x184) ?x184)) :pat { (uf_82 ?x183 ?x184) }) +:assumption (forall (?x185 Int) (iff (= (uf_74 uf_84 ?x185) uf_9) (and (<= ?x185 uf_75) (<= 0 ?x185))) :pat { (uf_74 uf_84 ?x185) }) +:assumption (forall (?x186 Int) (iff (= (uf_74 uf_4 ?x186) uf_9) (and (<= ?x186 uf_76) (<= 0 ?x186))) :pat { (uf_74 uf_4 ?x186) }) +:assumption (forall (?x187 Int) (iff (= (uf_74 uf_91 ?x187) uf_9) (and (<= ?x187 uf_77) (<= 0 ?x187))) :pat { (uf_74 uf_91 ?x187) }) +:assumption (forall (?x188 Int) (iff (= (uf_74 uf_7 ?x188) uf_9) (and (<= ?x188 uf_78) (<= 0 ?x188))) :pat { (uf_74 uf_7 ?x188) }) +:assumption (forall (?x189 Int) (iff (= (uf_74 uf_83 ?x189) uf_9) (and (<= ?x189 uf_85) (<= uf_86 ?x189))) :pat { (uf_74 uf_83 ?x189) }) +:assumption (forall (?x190 Int) (iff (= (uf_74 uf_87 ?x190) uf_9) (and (<= ?x190 uf_88) (<= uf_89 ?x190))) :pat { (uf_74 uf_87 ?x190) }) +:assumption (forall (?x191 Int) (iff (= (uf_74 uf_90 ?x191) uf_9) (and (<= ?x191 uf_92) (<= uf_93 ?x191))) :pat { (uf_74 uf_90 ?x191) }) +:assumption (forall (?x192 Int) (iff (= (uf_74 uf_94 ?x192) uf_9) (and (<= ?x192 uf_95) (<= uf_96 ?x192))) :pat { (uf_74 uf_94 ?x192) }) +:assumption (forall (?x193 Int) (?x194 Int) (?x195 Int) (?x196 Int) (implies (and (<= (uf_79 (+ (+ ?x194 ?x193) 1)) (uf_80 (uf_81 ?x195 (uf_79 ?x193)) (uf_79 (+ ?x194 ?x193)))) (and (<= 0 ?x195) (and (<= ?x194 ?x196) (and (< ?x193 ?x194) (<= 0 ?x193))))) (= (uf_97 ?x195 ?x196 ?x193 ?x194) (+ (uf_79 (+ (+ ?x194 ?x193) 1)) (uf_80 (uf_81 ?x195 (uf_79 ?x193)) (uf_79 (+ ?x194 ?x193)))))) :pat { (uf_97 ?x195 ?x196 ?x193 ?x194) }) +:assumption (forall (?x197 Int) (?x198 Int) (?x199 Int) (?x200 Int) (implies (and (< (uf_80 (uf_81 ?x199 (uf_79 ?x197)) (uf_79 (+ ?x198 ?x197))) (uf_79 (+ (+ ?x198 ?x197) 1))) (and (<= 0 ?x199) (and (<= ?x198 ?x200) (and (< ?x197 ?x198) (<= 0 ?x197))))) (= (uf_97 ?x199 ?x200 ?x197 ?x198) (uf_80 (uf_81 ?x199 (uf_79 ?x197)) (uf_79 (+ ?x198 ?x197))))) :pat { (uf_97 ?x199 ?x200 ?x197 ?x198) }) +:assumption (forall (?x201 Int) (?x202 Int) (?x203 Int) (?x204 Int) (implies (and (<= 0 ?x203) (and (<= ?x202 ?x204) (and (< ?x201 ?x202) (<= 0 ?x201)))) (= (uf_98 ?x203 ?x204 ?x201 ?x202) (uf_80 (uf_81 ?x203 (uf_79 ?x201)) (uf_79 (+ ?x202 ?x201))))) :pat { (uf_98 ?x203 ?x204 ?x201 ?x202) }) +:assumption (forall (?x205 Int) (?x206 Int) (?x207 Int) (implies (and (<= ?x206 ?x207) (and (< ?x205 ?x206) (<= 0 ?x205))) (= (uf_98 0 ?x207 ?x205 ?x206) 0)) :pat { (uf_98 0 ?x207 ?x205 ?x206) }) +:assumption (forall (?x208 Int) (?x209 Int) (?x210 Int) (implies (and (<= ?x209 ?x210) (and (< ?x208 ?x209) (<= 0 ?x208))) (= (uf_97 0 ?x210 ?x208 ?x209) 0)) :pat { (uf_97 0 ?x210 ?x208 ?x209) }) +:assumption (forall (?x211 Int) (?x212 Int) (?x213 Int) (?x214 Int) (?x215 Int) (?x216 Int) (?x217 Int) (implies (and (<= ?x212 ?x215) (and (< ?x211 ?x212) (<= 0 ?x211))) (implies (and (<= ?x217 ?x215) (and (< ?x216 ?x217) (<= 0 ?x216))) (implies (or (<= ?x212 ?x216) (<= ?x217 ?x211)) (= (uf_98 (uf_99 ?x214 ?x215 ?x211 ?x212 ?x213) ?x215 ?x216 ?x217) (uf_98 ?x214 ?x215 ?x216 ?x217))))) :pat { (uf_98 (uf_99 ?x214 ?x215 ?x211 ?x212 ?x213) ?x215 ?x216 ?x217) }) +:assumption (forall (?x218 Int) (?x219 Int) (?x220 Int) (?x221 Int) (?x222 Int) (?x223 Int) (?x224 Int) (implies (and (<= ?x219 ?x222) (and (< ?x218 ?x219) (<= 0 ?x218))) (implies (and (<= ?x224 ?x222) (and (< ?x223 ?x224) (<= 0 ?x223))) (implies (or (<= ?x219 ?x223) (<= ?x224 ?x218)) (= (uf_97 (uf_99 ?x221 ?x222 ?x218 ?x219 ?x220) ?x222 ?x223 ?x224) (uf_97 ?x221 ?x222 ?x223 ?x224))))) :pat { (uf_97 (uf_99 ?x221 ?x222 ?x218 ?x219 ?x220) ?x222 ?x223 ?x224) }) +:assumption (forall (?x225 Int) (?x226 Int) (?x227 Int) (?x228 Int) (implies (and (<= ?x226 ?x228) (and (< ?x225 ?x226) (<= 0 ?x225))) (and (<= (uf_98 ?x227 ?x228 ?x225 ?x226) (+ (uf_79 (+ ?x226 ?x225)) 1)) (<= 0 (uf_98 ?x227 ?x228 ?x225 ?x226)))) :pat { (uf_98 ?x227 ?x228 ?x225 ?x226) }) +:assumption (forall (?x229 Int) (?x230 Int) (?x231 Int) (?x232 Int) (implies (and (<= ?x230 ?x232) (and (< ?x229 ?x230) (<= 0 ?x229))) (and (<= (uf_97 ?x231 ?x232 ?x229 ?x230) (+ (uf_79 (+ (+ ?x230 ?x229) 1)) 1)) (<= (+ 0 (uf_79 (+ (+ ?x230 ?x229) 1))) (uf_97 ?x231 ?x232 ?x229 ?x230)))) :pat { (uf_97 ?x231 ?x232 ?x229 ?x230) }) +:assumption (forall (?x233 Int) (?x234 Int) (?x235 Int) (?x236 Int) (?x237 Int) (implies (and (<= ?x234 ?x237) (and (< ?x233 ?x234) (<= 0 ?x233))) (implies (and (< ?x235 (uf_79 (+ ?x234 ?x233))) (<= 0 ?x235)) (= (uf_98 (uf_99 ?x236 ?x237 ?x233 ?x234 ?x235) ?x237 ?x233 ?x234) ?x235))) :pat { (uf_98 (uf_99 ?x236 ?x237 ?x233 ?x234 ?x235) ?x237 ?x233 ?x234) }) +:assumption (forall (?x238 Int) (?x239 Int) (?x240 Int) (?x241 Int) (?x242 Int) (implies (and (<= ?x239 ?x242) (and (< ?x238 ?x239) (<= 0 ?x238))) (implies (and (< ?x240 (uf_79 (+ (+ ?x239 ?x238) 1))) (<= (+ 0 (uf_79 (+ (+ ?x239 ?x238) 1))) ?x240)) (= (uf_97 (uf_99 ?x241 ?x242 ?x238 ?x239 ?x240) ?x242 ?x238 ?x239) ?x240))) :pat { (uf_97 (uf_99 ?x241 ?x242 ?x238 ?x239 ?x240) ?x242 ?x238 ?x239) }) +:assumption (forall (?x243 Int) (?x244 Int) (?x245 Int) (implies (and (<= ?x244 ?x245) (and (< ?x243 ?x244) (<= 0 ?x243))) (= (uf_99 0 ?x245 ?x243 ?x244 0) 0)) :pat { (uf_99 0 ?x245 ?x243 ?x244 0) }) +:assumption (forall (?x246 Int) (?x247 Int) (?x248 Int) (?x249 Int) (?x250 Int) (implies (and (<= ?x248 ?x249) (and (< ?x247 ?x248) (<= 0 ?x247))) (implies (and (< ?x250 (uf_79 (+ ?x248 ?x247))) (<= 0 ?x250)) (and (< (uf_99 ?x246 ?x249 ?x247 ?x248 ?x250) (uf_79 ?x249)) (<= 0 (uf_99 ?x246 ?x249 ?x247 ?x248 ?x250))))) :pat { (uf_99 ?x246 ?x249 ?x247 ?x248 ?x250) }) +:assumption (forall (?x251 Int) (?x252 Int) (= (uf_100 ?x251 ?x252) (uf_81 ?x251 (uf_79 ?x252))) :pat { (uf_100 ?x251 ?x252) }) +:assumption (forall (?x253 T3) (?x254 Int) (?x255 Int) (= (uf_101 ?x253 ?x254 ?x255) (uf_82 ?x253 (+ ?x254 (uf_79 ?x255)))) :pat { (uf_101 ?x253 ?x254 ?x255) }) +:assumption (forall (?x256 T3) (?x257 Int) (?x258 Int) (= (uf_102 ?x256 ?x257 ?x258) (uf_82 ?x256 (uf_80 ?x257 ?x258))) :pat { (uf_102 ?x256 ?x257 ?x258) }) +:assumption (forall (?x259 T3) (?x260 Int) (?x261 Int) (= (uf_103 ?x259 ?x260 ?x261) (uf_82 ?x259 (uf_81 ?x260 ?x261))) :pat { (uf_103 ?x259 ?x260 ?x261) }) +:assumption (forall (?x262 T3) (?x263 Int) (?x264 Int) (= (uf_104 ?x262 ?x263 ?x264) (uf_82 ?x262 (+ ?x263 ?x264))) :pat { (uf_104 ?x262 ?x263 ?x264) }) +:assumption (forall (?x265 T3) (?x266 Int) (?x267 Int) (= (uf_105 ?x265 ?x266 ?x267) (uf_82 ?x265 (+ ?x266 ?x267))) :pat { (uf_105 ?x265 ?x266 ?x267) }) +:assumption (forall (?x268 T3) (?x269 Int) (?x270 Int) (= (uf_106 ?x268 ?x269 ?x270) (uf_82 ?x268 (+ ?x269 ?x270))) :pat { (uf_106 ?x268 ?x269 ?x270) }) +:assumption (and (= (uf_79 63) 9223372036854775808) (and (= (uf_79 62) 4611686018427387904) (and (= (uf_79 61) 2305843009213693952) (and (= (uf_79 60) 1152921504606846976) (and (= (uf_79 59) 576460752303423488) (and (= (uf_79 58) 288230376151711744) (and (= (uf_79 57) 144115188075855872) (and (= (uf_79 56) 72057594037927936) (and (= (uf_79 55) 36028797018963968) (and (= (uf_79 54) 18014398509481984) (and (= (uf_79 53) 9007199254740992) (and (= (uf_79 52) 4503599627370496) (and (= (uf_79 51) 2251799813685248) (and (= (uf_79 50) 1125899906842624) (and (= (uf_79 49) 562949953421312) (and (= (uf_79 48) 281474976710656) (and (= (uf_79 47) 140737488355328) (and (= (uf_79 46) 70368744177664) (and (= (uf_79 45) 35184372088832) (and (= (uf_79 44) 17592186044416) (and (= (uf_79 43) 8796093022208) (and (= (uf_79 42) 4398046511104) (and (= (uf_79 41) 2199023255552) (and (= (uf_79 40) 1099511627776) (and (= (uf_79 39) 549755813888) (and (= (uf_79 38) 274877906944) (and (= (uf_79 37) 137438953472) (and (= (uf_79 36) 68719476736) (and (= (uf_79 35) 34359738368) (and (= (uf_79 34) 17179869184) (and (= (uf_79 33) 8589934592) (and (= (uf_79 32) 4294967296) (and (= (uf_79 31) 2147483648) (and (= (uf_79 30) 1073741824) (and (= (uf_79 29) 536870912) (and (= (uf_79 28) 268435456) (and (= (uf_79 27) 134217728) (and (= (uf_79 26) 67108864) (and (= (uf_79 25) 33554432) (and (= (uf_79 24) 16777216) (and (= (uf_79 23) 8388608) (and (= (uf_79 22) 4194304) (and (= (uf_79 21) 2097152) (and (= (uf_79 20) 1048576) (and (= (uf_79 19) 524288) (and (= (uf_79 18) 262144) (and (= (uf_79 17) 131072) (and (= (uf_79 16) 65536) (and (= (uf_79 15) 32768) (and (= (uf_79 14) 16384) (and (= (uf_79 13) 8192) (and (= (uf_79 12) 4096) (and (= (uf_79 11) 2048) (and (= (uf_79 10) 1024) (and (= (uf_79 9) 512) (and (= (uf_79 8) 256) (and (= (uf_79 7) 128) (and (= (uf_79 6) 64) (and (= (uf_79 5) 32) (and (= (uf_79 4) 16) (and (= (uf_79 3) 8) (and (= (uf_79 2) 4) (and (= (uf_79 1) 2) (= (uf_79 0) 1)))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) +:assumption (forall (?x271 T4) (?x272 T5) (implies (= (uf_51 ?x271) uf_9) (and (<= (uf_107 ?x271 ?x272) uf_75) (<= 0 (uf_107 ?x271 ?x272)))) :pat { (uf_107 ?x271 ?x272) }) +:assumption (forall (?x273 T4) (?x274 T5) (implies (= (uf_51 ?x273) uf_9) (and (<= (uf_108 ?x273 ?x274) uf_76) (<= 0 (uf_108 ?x273 ?x274)))) :pat { (uf_108 ?x273 ?x274) }) +:assumption (forall (?x275 T4) (?x276 T5) (implies (= (uf_51 ?x275) uf_9) (and (<= (uf_109 ?x275 ?x276) uf_77) (<= 0 (uf_109 ?x275 ?x276)))) :pat { (uf_109 ?x275 ?x276) }) +:assumption (forall (?x277 T4) (?x278 T5) (implies (= (uf_51 ?x277) uf_9) (and (<= (uf_110 ?x277 ?x278) uf_78) (<= 0 (uf_110 ?x277 ?x278)))) :pat { (uf_110 ?x277 ?x278) }) +:assumption (forall (?x279 T4) (?x280 T5) (implies (= (uf_51 ?x279) uf_9) (and (<= (uf_111 ?x279 ?x280) uf_85) (<= uf_86 (uf_111 ?x279 ?x280)))) :pat { (uf_111 ?x279 ?x280) }) +:assumption (forall (?x281 T4) (?x282 T5) (implies (= (uf_51 ?x281) uf_9) (and (<= (uf_112 ?x281 ?x282) uf_88) (<= uf_89 (uf_112 ?x281 ?x282)))) :pat { (uf_112 ?x281 ?x282) }) +:assumption (forall (?x283 T4) (?x284 T5) (implies (= (uf_51 ?x283) uf_9) (and (<= (uf_113 ?x283 ?x284) uf_92) (<= uf_93 (uf_113 ?x283 ?x284)))) :pat { (uf_113 ?x283 ?x284) }) +:assumption (forall (?x285 T4) (?x286 T5) (implies (= (uf_51 ?x285) uf_9) (and (<= (uf_114 ?x285 ?x286) uf_95) (<= uf_96 (uf_114 ?x285 ?x286)))) :pat { (uf_114 ?x285 ?x286) }) +:assumption (forall (?x287 T5) (?x288 T5) (= (uf_115 ?x287 ?x288) (+ (uf_116 ?x287) (uf_116 ?x288))) :pat { (uf_115 ?x287 ?x288) }) +:assumption (forall (?x289 T5) (implies (and (<= (uf_116 ?x289) uf_88) (<= uf_89 (uf_116 ?x289))) (= (uf_117 ?x289) (uf_116 ?x289))) :pat { (uf_117 ?x289) }) +:assumption (forall (?x290 T5) (implies (and (<= (uf_116 ?x290) uf_76) (<= 0 (uf_116 ?x290))) (= (uf_118 ?x290) (uf_116 ?x290))) :pat { (uf_118 ?x290) }) +:assumption (forall (?x291 T5) (implies (and (<= (uf_116 ?x291) uf_85) (<= uf_86 (uf_116 ?x291))) (= (uf_119 ?x291) (uf_116 ?x291))) :pat { (uf_119 ?x291) }) +:assumption (forall (?x292 T5) (implies (and (<= (uf_116 ?x292) uf_75) (<= 0 (uf_116 ?x292))) (= (uf_120 ?x292) (uf_116 ?x292))) :pat { (uf_120 ?x292) }) +:assumption (= (uf_117 uf_121) 0) +:assumption (= (uf_118 uf_121) 0) +:assumption (= (uf_119 uf_121) 0) +:assumption (= (uf_120 uf_121) 0) +:assumption (forall (?x293 T4) (?x294 T5) (= (uf_107 ?x293 ?x294) (uf_19 (uf_20 ?x293) ?x294)) :pat { (uf_107 ?x293 ?x294) }) +:assumption (forall (?x295 T4) (?x296 T5) (= (uf_108 ?x295 ?x296) (uf_19 (uf_20 ?x295) ?x296)) :pat { (uf_108 ?x295 ?x296) }) +:assumption (forall (?x297 T4) (?x298 T5) (= (uf_109 ?x297 ?x298) (uf_19 (uf_20 ?x297) ?x298)) :pat { (uf_109 ?x297 ?x298) }) +:assumption (forall (?x299 T4) (?x300 T5) (= (uf_110 ?x299 ?x300) (uf_19 (uf_20 ?x299) ?x300)) :pat { (uf_110 ?x299 ?x300) }) +:assumption (forall (?x301 T4) (?x302 T5) (= (uf_111 ?x301 ?x302) (uf_19 (uf_20 ?x301) ?x302)) :pat { (uf_111 ?x301 ?x302) }) +:assumption (forall (?x303 T4) (?x304 T5) (= (uf_112 ?x303 ?x304) (uf_19 (uf_20 ?x303) ?x304)) :pat { (uf_112 ?x303 ?x304) }) +:assumption (forall (?x305 T4) (?x306 T5) (= (uf_113 ?x305 ?x306) (uf_19 (uf_20 ?x305) ?x306)) :pat { (uf_113 ?x305 ?x306) }) +:assumption (forall (?x307 T4) (?x308 T5) (= (uf_114 ?x307 ?x308) (uf_19 (uf_20 ?x307) ?x308)) :pat { (uf_114 ?x307 ?x308) }) +:assumption (= uf_75 (+ (+ (+ (+ 65536 65536) 65536) 65536) 1)) +:assumption (= uf_76 (+ (+ 65536 65536) 1)) +:assumption (= uf_77 65535) +:assumption (= uf_78 255) +:assumption (= uf_85 (+ (+ (+ (+ 65536 65536) 65536) 32768) 1)) +:assumption (= uf_86 (+ 0 (+ (+ (+ 65536 65536) 65536) 32768))) +:assumption (= uf_88 (+ (+ 65536 32768) 1)) +:assumption (= uf_89 (+ 0 (+ 65536 32768))) +:assumption (= uf_92 32767) +:assumption (= uf_93 (+ 0 32768)) +:assumption (= uf_95 127) +:assumption (= uf_96 (+ 0 128)) +:assumption (forall (?x309 T2) (iff (= (uf_122 ?x309) uf_9) (= ?x309 uf_9)) :pat { (uf_122 ?x309) }) +:assumption (forall (?x310 T4) (?x311 T4) (?x312 T5) (?x313 T3) (?x314 Int) (implies (= (uf_23 ?x313) uf_9) (implies (= (uf_123 ?x310 ?x311 ?x312 (uf_124 ?x313 ?x314)) uf_9) (forall (?x315 Int) (implies (and (< ?x315 ?x314) (<= 0 ?x315)) (= (uf_19 (uf_20 ?x310) (uf_66 (uf_43 ?x313 (uf_116 ?x312)) ?x315 ?x313)) (uf_19 (uf_20 ?x311) (uf_66 (uf_43 ?x313 (uf_116 ?x312)) ?x315 ?x313)))) :pat { (uf_19 (uf_20 ?x311) (uf_66 (uf_43 ?x313 (uf_116 ?x312)) ?x315 ?x313)) }))) :pat { (uf_123 ?x310 ?x311 ?x312 (uf_124 ?x313 ?x314)) (uf_23 ?x313) }) +:assumption (forall (?x316 T5) (?x317 Int) (?x318 T15) (= (uf_125 (uf_126 (uf_66 ?x316 ?x317 (uf_15 ?x316)) ?x318) ?x316) ?x317) :pat { (uf_125 (uf_126 (uf_66 ?x316 ?x317 (uf_15 ?x316)) ?x318) ?x316) }) +:assumption (forall (?x319 T5) (?x320 Int) (= (uf_125 (uf_66 ?x319 ?x320 (uf_15 ?x319)) ?x319) ?x320) :pat { (uf_66 ?x319 ?x320 (uf_15 ?x319)) }) +:assumption (forall (?x321 T5) (?x322 T4) (?x323 T5) (iff (= (uf_13 ?x321 (uf_127 ?x322 ?x323)) uf_9) (and (= (uf_13 ?x321 (uf_128 ?x322 ?x323)) uf_9) (not (= (uf_116 ?x323) (uf_116 uf_121))))) :pat { (uf_13 ?x321 (uf_127 ?x322 ?x323)) }) +:assumption (forall (?x324 T5) (?x325 Int) (?x326 T3) (?x327 Int) (iff (= (uf_13 ?x324 (uf_129 (uf_43 ?x326 ?x325) ?x326 ?x327)) uf_9) (and (= (uf_13 ?x324 (uf_130 (uf_66 (uf_43 ?x326 ?x325) (uf_125 ?x324 (uf_43 ?x326 ?x325)) ?x326))) uf_9) (and (<= (uf_125 ?x324 (uf_43 ?x326 ?x325)) (+ ?x327 1)) (and (<= 0 (uf_125 ?x324 (uf_43 ?x326 ?x325))) (not (= ?x325 0)))))) :pat { (uf_13 ?x324 (uf_129 (uf_43 ?x326 ?x325) ?x326 ?x327)) }) +:assumption (forall (?x328 T5) (?x329 T3) (?x330 Int) (?x331 Int) (?x332 T6) (implies (and (< ?x331 ?x330) (<= 0 ?x331)) (= (uf_133 (uf_66 ?x328 ?x331 ?x329) ?x332 (uf_132 ?x328 ?x329 ?x330)) 2)) :pat { (uf_66 ?x328 ?x331 ?x329) (uf_131 ?x332 (uf_132 ?x328 ?x329 ?x330)) }) +:assumption (forall (?x333 T5) (?x334 T3) (?x335 Int) (?x336 Int) (?x337 T6) (implies (and (< ?x336 ?x335) (<= 0 ?x336)) (= (uf_133 (uf_66 ?x333 ?x336 ?x334) (uf_132 ?x333 ?x334 ?x335) ?x337) 1)) :pat { (uf_66 ?x333 ?x336 ?x334) (uf_131 (uf_132 ?x333 ?x334 ?x335) ?x337) }) +:assumption (forall (?x338 T5) (?x339 Int) (?x340 T3) (?x341 Int) (iff (= (uf_13 ?x338 (uf_132 (uf_43 ?x340 ?x339) ?x340 ?x341)) uf_9) (and (= (uf_13 ?x338 (uf_130 (uf_66 (uf_43 ?x340 ?x339) (uf_125 ?x338 (uf_43 ?x340 ?x339)) ?x340))) uf_9) (and (<= (uf_125 ?x338 (uf_43 ?x340 ?x339)) (+ ?x341 1)) (<= 0 (uf_125 ?x338 (uf_43 ?x340 ?x339)))))) :pat { (uf_13 ?x338 (uf_132 (uf_43 ?x340 ?x339) ?x340 ?x341)) }) +:assumption (forall (?x342 T5) (?x343 T3) (?x344 Int) (?x345 T5) (iff (= (uf_13 ?x345 (uf_134 ?x342 ?x343 ?x344)) uf_9) (and (= ?x345 (uf_66 ?x342 (uf_125 ?x345 ?x342) ?x343)) (and (<= (uf_125 ?x345 ?x342) (+ ?x344 1)) (<= 0 (uf_125 ?x345 ?x342))))) :pat { (uf_13 ?x345 (uf_134 ?x342 ?x343 ?x344)) }) +:assumption (forall (?x346 T4) (?x347 Int) (?x348 T3) (?x349 Int) (?x350 Int) (implies (= (uf_27 ?x346 (uf_43 (uf_124 ?x348 ?x349) ?x347)) uf_9) (implies (and (< ?x350 ?x349) (<= 0 ?x350)) (and (= (uf_27 ?x346 (uf_66 (uf_43 (uf_124 ?x348 ?x349) ?x347) ?x350 ?x348)) uf_9) (and (up_68 (uf_58 (uf_59 ?x346) (uf_66 (uf_43 (uf_124 ?x348 ?x349) ?x347) ?x350 ?x348))) (and (not (= (uf_135 (uf_58 (uf_59 ?x346) (uf_66 (uf_43 (uf_124 ?x348 ?x349) ?x347) ?x350 ?x348))) uf_9)) (= (uf_136 (uf_58 (uf_59 ?x346) (uf_66 (uf_43 (uf_124 ?x348 ?x349) ?x347) ?x350 ?x348))) (uf_43 (uf_124 ?x348 ?x349) ?x347))))))) :pat { (uf_40 (uf_41 ?x346) (uf_66 (uf_43 ?x348 ?x347) ?x350 ?x348)) (uf_43 (uf_124 ?x348 ?x349) ?x347) } :pat { (uf_58 (uf_59 ?x346) (uf_66 (uf_43 ?x348 ?x347) ?x350 ?x348)) (uf_43 (uf_124 ?x348 ?x349) ?x347) }) +:assumption (forall (?x351 T4) (?x352 T5) (?x353 Int) (?x354 T3) (?x355 Int) (iff (= (uf_13 ?x352 (uf_128 ?x351 (uf_43 (uf_124 ?x354 ?x355) ?x353))) uf_9) (or (and (= (uf_13 ?x352 (uf_128 ?x351 (uf_66 (uf_43 ?x354 ?x353) (uf_125 ?x352 (uf_43 ?x354 ?x353)) ?x354))) uf_9) (and (<= (uf_125 ?x352 (uf_43 ?x354 ?x353)) (+ ?x355 1)) (<= 0 (uf_125 ?x352 (uf_43 ?x354 ?x353))))) (= ?x352 (uf_43 (uf_124 ?x354 ?x355) ?x353)))) :pat { (uf_13 ?x352 (uf_128 ?x351 (uf_43 (uf_124 ?x354 ?x355) ?x353))) }) +:assumption (forall (?x356 T5) (?x357 Int) (?x358 T3) (?x359 Int) (iff (= (uf_13 ?x356 (uf_130 (uf_43 (uf_124 ?x358 ?x359) ?x357))) uf_9) (or (and (= (uf_13 ?x356 (uf_130 (uf_66 (uf_43 ?x358 ?x357) (uf_125 ?x356 (uf_43 ?x358 ?x357)) ?x358))) uf_9) (and (<= (uf_125 ?x356 (uf_43 ?x358 ?x357)) (+ ?x359 1)) (<= 0 (uf_125 ?x356 (uf_43 ?x358 ?x357))))) (= ?x356 (uf_43 (uf_124 ?x358 ?x359) ?x357)))) :pat { (uf_13 ?x356 (uf_130 (uf_43 (uf_124 ?x358 ?x359) ?x357))) }) +:assumption (forall (?x360 T4) (?x361 T5) (?x362 T3) (?x363 Int) (iff (= (uf_65 ?x360 ?x361 ?x362 ?x363) uf_9) (and (forall (?x364 Int) (implies (and (< ?x364 ?x363) (<= 0 ?x364)) (and (= (uf_27 ?x360 (uf_66 ?x361 ?x364 ?x362)) uf_9) (up_68 (uf_58 (uf_59 ?x360) (uf_66 ?x361 ?x364 ?x362))))) :pat { (uf_40 (uf_41 ?x360) (uf_66 ?x361 ?x364 ?x362)) } :pat { (uf_58 (uf_59 ?x360) (uf_66 ?x361 ?x364 ?x362)) } :pat { (uf_19 (uf_20 ?x360) (uf_66 ?x361 ?x364 ?x362)) }) (= (uf_48 ?x361 ?x362) uf_9))) :pat { (uf_65 ?x360 ?x361 ?x362 ?x363) }) +:assumption (forall (?x365 T4) (?x366 T5) (?x367 T3) (?x368 Int) (?x369 T2) (iff (= (uf_137 ?x365 ?x366 ?x367 ?x368 ?x369) uf_9) (and (forall (?x370 Int) (implies (and (< ?x370 ?x368) (<= 0 ?x370)) (and (= (uf_27 ?x365 (uf_66 ?x366 ?x370 ?x367)) uf_9) (and (up_68 (uf_58 (uf_59 ?x365) (uf_66 ?x366 ?x370 ?x367))) (iff (= (uf_135 (uf_58 (uf_59 ?x365) (uf_66 ?x366 ?x370 ?x367))) uf_9) (= ?x369 uf_9))))) :pat { (uf_40 (uf_41 ?x365) (uf_66 ?x366 ?x370 ?x367)) } :pat { (uf_58 (uf_59 ?x365) (uf_66 ?x366 ?x370 ?x367)) } :pat { (uf_19 (uf_20 ?x365) (uf_66 ?x366 ?x370 ?x367)) }) (= (uf_48 ?x366 ?x367) uf_9))) :pat { (uf_137 ?x365 ?x366 ?x367 ?x368 ?x369) }) +:assumption (forall (?x371 T5) (?x372 Int) (?x373 Int) (?x374 T3) (implies (and (not (= ?x373 0)) (not (= ?x372 0))) (= (uf_66 (uf_66 ?x371 ?x372 ?x374) ?x373 ?x374) (uf_66 ?x371 (+ ?x372 ?x373) ?x374))) :pat { (uf_66 (uf_66 ?x371 ?x372 ?x374) ?x373 ?x374) }) +:assumption (forall (?x375 T5) (?x376 Int) (?x377 T3) (and (= (uf_66 ?x375 ?x376 ?x377) (uf_43 ?x377 (+ (uf_116 ?x375) (+ ?x376 (uf_138 ?x377))))) (= (uf_139 (uf_66 ?x375 ?x376 ?x377) ?x375) uf_9)) :pat { (uf_66 ?x375 ?x376 ?x377) }) +:assumption (forall (?x378 T5) (?x379 T3) (= (uf_140 ?x378 ?x379) ?x378) :pat { (uf_140 ?x378 ?x379) }) +:assumption (forall (?x380 T3) (?x381 Int) (not (up_36 (uf_124 ?x380 ?x381))) :pat { (uf_124 ?x380 ?x381) }) +:assumption (forall (?x382 T3) (?x383 Int) (= (uf_141 (uf_124 ?x382 ?x383)) uf_9) :pat { (uf_124 ?x382 ?x383) }) +:assumption (forall (?x384 T3) (?x385 Int) (= (uf_142 (uf_124 ?x384 ?x385)) 0) :pat { (uf_124 ?x384 ?x385) }) +:assumption (forall (?x386 T3) (?x387 Int) (= (uf_143 (uf_124 ?x386 ?x387)) ?x387) :pat { (uf_124 ?x386 ?x387) }) +:assumption (forall (?x388 T3) (?x389 Int) (= (uf_144 (uf_124 ?x388 ?x389)) ?x388) :pat { (uf_124 ?x388 ?x389) }) +:assumption (forall (?x390 T5) (?x391 T6) (iff (= (uf_13 ?x390 ?x391) uf_9) (= (uf_145 ?x390 ?x391) uf_9)) :pat { (uf_145 ?x390 ?x391) }) +:assumption (forall (?x392 T5) (?x393 T6) (iff (= (uf_13 ?x392 ?x393) uf_9) (up_146 ?x392 ?x393)) :pat { (uf_13 ?x392 ?x393) }) +:assumption (forall (?x394 T5) (?x395 T6) (iff (= (uf_13 ?x394 ?x395) uf_9) (= (uf_147 ?x394 ?x395) uf_9)) :pat { (uf_13 ?x394 ?x395) }) +:assumption (forall (?x396 T5) (?x397 T4) (?x398 T5) (iff (= (uf_13 ?x396 (uf_53 ?x397 ?x398)) uf_9) (= (uf_147 ?x396 (uf_53 ?x397 ?x398)) uf_9)) :pat { (uf_147 ?x396 (uf_53 ?x397 ?x398)) (uf_148 ?x396) }) +:assumption (forall (?x399 T5) (?x400 T4) (?x401 T5) (implies (= (uf_13 ?x399 (uf_53 ?x400 ?x401)) uf_9) (= (uf_148 ?x399) uf_9)) :pat { (uf_13 ?x399 (uf_53 ?x400 ?x401)) }) +:assumption (forall (?x402 T6) (?x403 T6) (implies (forall (?x404 T5) (and (implies (= (uf_13 ?x404 ?x403) uf_9) (not (= (uf_13 ?x404 ?x402) uf_9))) (implies (= (uf_13 ?x404 ?x402) uf_9) (not (= (uf_13 ?x404 ?x403) uf_9)))) :pat { (uf_18 ?x404) }) (= (uf_131 ?x402 ?x403) uf_9)) :pat { (uf_131 ?x402 ?x403) }) +:assumption (forall (?x405 T5) (?x406 T6) (?x407 T6) (implies (and (= (uf_13 ?x405 ?x407) uf_9) (= (uf_131 ?x406 ?x407) uf_9)) (= (uf_133 ?x405 ?x406 ?x407) 2)) :pat { (uf_131 ?x406 ?x407) (uf_13 ?x405 ?x407) }) +:assumption (forall (?x408 T5) (?x409 T6) (?x410 T6) (implies (and (= (uf_13 ?x408 ?x409) uf_9) (= (uf_131 ?x409 ?x410) uf_9)) (= (uf_133 ?x408 ?x409 ?x410) 1)) :pat { (uf_131 ?x409 ?x410) (uf_13 ?x408 ?x409) }) +:assumption (forall (?x411 T5) (= (uf_13 ?x411 uf_149) uf_9) :pat { (uf_13 ?x411 uf_149) }) +:assumption (forall (?x412 T5) (= (uf_150 (uf_151 ?x412)) 1)) +:assumption (= (uf_150 uf_152) 0) +:assumption (forall (?x413 T6) (?x414 T6) (implies (= (uf_153 ?x413 ?x414) uf_9) (= ?x413 ?x414)) :pat { (uf_153 ?x413 ?x414) }) +:assumption (forall (?x415 T6) (?x416 T6) (implies (forall (?x417 T5) (iff (= (uf_13 ?x417 ?x415) uf_9) (= (uf_13 ?x417 ?x416) uf_9)) :pat { (uf_18 ?x417) }) (= (uf_153 ?x415 ?x416) uf_9)) :pat { (uf_153 ?x415 ?x416) }) +:assumption (forall (?x418 T6) (?x419 T6) (iff (= (uf_154 ?x418 ?x419) uf_9) (forall (?x420 T5) (implies (= (uf_13 ?x420 ?x418) uf_9) (= (uf_13 ?x420 ?x419) uf_9)) :pat { (uf_13 ?x420 ?x418) } :pat { (uf_13 ?x420 ?x419) })) :pat { (uf_154 ?x418 ?x419) }) +:assumption (forall (?x421 T6) (?x422 T6) (?x423 T5) (iff (= (uf_13 ?x423 (uf_155 ?x421 ?x422)) uf_9) (and (= (uf_13 ?x423 ?x422) uf_9) (= (uf_13 ?x423 ?x421) uf_9))) :pat { (uf_13 ?x423 (uf_155 ?x421 ?x422)) }) +:assumption (forall (?x424 T6) (?x425 T6) (?x426 T5) (iff (= (uf_13 ?x426 (uf_156 ?x424 ?x425)) uf_9) (and (not (= (uf_13 ?x426 ?x425) uf_9)) (= (uf_13 ?x426 ?x424) uf_9))) :pat { (uf_13 ?x426 (uf_156 ?x424 ?x425)) }) +:assumption (forall (?x427 T6) (?x428 T6) (?x429 T5) (iff (= (uf_13 ?x429 (uf_157 ?x427 ?x428)) uf_9) (or (= (uf_13 ?x429 ?x428) uf_9) (= (uf_13 ?x429 ?x427) uf_9))) :pat { (uf_13 ?x429 (uf_157 ?x427 ?x428)) }) +:assumption (forall (?x430 T5) (?x431 T5) (iff (= (uf_13 ?x431 (uf_158 ?x430)) uf_9) (and (not (= (uf_116 ?x430) (uf_116 uf_121))) (= ?x430 ?x431))) :pat { (uf_13 ?x431 (uf_158 ?x430)) }) +:assumption (forall (?x432 T5) (?x433 T5) (iff (= (uf_13 ?x433 (uf_151 ?x432)) uf_9) (= ?x432 ?x433)) :pat { (uf_13 ?x433 (uf_151 ?x432)) }) +:assumption (forall (?x434 T5) (not (= (uf_13 ?x434 uf_152) uf_9)) :pat { (uf_13 ?x434 uf_152) }) +:assumption (forall (?x435 T5) (?x436 T5) (= (uf_159 ?x435 ?x436) (uf_43 (uf_124 (uf_144 (uf_15 ?x435)) (+ (uf_143 (uf_15 ?x435)) (uf_143 (uf_15 ?x436)))) (uf_116 ?x435))) :pat { (uf_159 ?x435 ?x436) }) +:assumption (forall (?x437 T5) (?x438 Int) (= (uf_160 ?x437 ?x438) (uf_43 (uf_124 (uf_144 (uf_15 ?x437)) (+ (uf_143 (uf_15 ?x437)) ?x438)) (uf_116 (uf_66 (uf_43 (uf_144 (uf_15 ?x437)) (uf_116 ?x437)) ?x438 (uf_144 (uf_15 ?x437)))))) :pat { (uf_160 ?x437 ?x438) }) +:assumption (forall (?x439 T5) (?x440 Int) (= (uf_161 ?x439 ?x440) (uf_43 (uf_124 (uf_144 (uf_15 ?x439)) ?x440) (uf_116 ?x439))) :pat { (uf_161 ?x439 ?x440) }) +:assumption (forall (?x441 T4) (?x442 T5) (?x443 T5) (iff (= (uf_13 ?x442 (uf_162 ?x441 ?x443)) uf_9) (or (and (= (uf_13 ?x442 (uf_163 ?x443)) uf_9) (= (uf_135 (uf_58 (uf_59 ?x441) ?x442)) uf_9)) (= ?x442 ?x443))) :pat { (uf_13 ?x442 (uf_162 ?x441 ?x443)) }) +:assumption (forall (?x444 T4) (implies (= (uf_164 ?x444) uf_9) (up_165 ?x444)) :pat { (uf_164 ?x444) }) +:assumption (= (uf_142 uf_166) 0) +:assumption (= uf_167 (uf_43 uf_166 uf_168)) +:assumption (forall (?x445 T4) (?x446 T4) (?x447 T5) (?x448 T5) (and true (and (= (uf_170 (uf_169 ?x445 ?x446 ?x447 ?x448) ?x448) (uf_171 ?x445)) (and (= (uf_38 (uf_169 ?x445 ?x446 ?x447 ?x448) ?x448) (uf_38 ?x446 ?x448)) (and (= (uf_25 (uf_169 ?x445 ?x446 ?x447 ?x448) ?x448) uf_26) (and (= (uf_24 (uf_169 ?x445 ?x446 ?x447 ?x448) ?x448) uf_9) (= (uf_41 (uf_169 ?x445 ?x446 ?x447 ?x448)) (uf_172 (uf_41 ?x446) ?x448 (uf_173 ?x446 ?x447 ?x448)))))))) :pat { (uf_169 ?x445 ?x446 ?x447 ?x448) }) +:assumption (forall (?x449 T4) (?x450 T5) (?x451 T5) (implies (not (= (uf_14 (uf_15 ?x450)) uf_16)) (and true (and (= (uf_38 (uf_174 ?x449 ?x450 ?x451) ?x451) (uf_38 ?x449 ?x451)) (and (= (uf_25 (uf_174 ?x449 ?x450 ?x451) ?x451) ?x450) (and (= (uf_24 (uf_174 ?x449 ?x450 ?x451) ?x451) uf_9) (= (uf_41 (uf_174 ?x449 ?x450 ?x451)) (uf_172 (uf_41 ?x449) ?x451 (uf_175 ?x449 ?x450 ?x451)))))))) :pat { (uf_174 ?x449 ?x450 ?x451) }) +:assumption (forall (?x452 T4) (?x453 T5) (?x454 Int) (and (= (uf_177 ?x452 (uf_176 ?x452 ?x453 ?x454)) uf_9) (and (forall (?x455 T5) (<= (uf_170 ?x452 ?x455) (uf_170 (uf_176 ?x452 ?x455 ?x454) ?x455)) :pat { (uf_170 (uf_176 ?x452 ?x455 ?x454) ?x455) }) (and (< (uf_171 ?x452) (uf_171 (uf_176 ?x452 ?x453 ?x454))) (and (= (uf_20 (uf_176 ?x452 ?x453 ?x454)) (uf_178 (uf_20 ?x452) ?x453 ?x454)) (and (= (uf_41 (uf_176 ?x452 ?x453 ?x454)) (uf_41 ?x452)) (= (uf_59 (uf_176 ?x452 ?x453 ?x454)) (uf_59 ?x452))))))) :pat { (uf_176 ?x452 ?x453 ?x454) }) +:assumption (forall (?x456 T4) (implies (= (uf_51 ?x456) uf_9) (forall (?x457 T5) (?x458 T5) (implies (and (= (uf_24 ?x456 ?x458) uf_9) (and (= (uf_13 ?x457 (uf_53 ?x456 ?x458)) uf_9) (= (uf_51 ?x456) uf_9))) (and (not (= (uf_116 ?x457) 0)) (= (uf_24 ?x456 ?x457) uf_9))) :pat { (uf_13 ?x457 (uf_53 ?x456 ?x458)) })) :pat { (uf_51 ?x456) }) +:assumption (forall (?x459 T4) (?x460 T5) (?x461 T3) (implies (and (= (uf_24 ?x459 ?x460) uf_9) (= (uf_44 ?x459) uf_9)) (= (uf_46 ?x459 ?x459 ?x460 ?x461) uf_9)) :pat { (uf_46 ?x459 ?x459 ?x460 ?x461) }) +:assumption (forall (?x462 T4) (?x463 Int) (?x464 T3) (implies (= (uf_51 ?x462) uf_9) (implies (= (uf_141 ?x464) uf_9) (= (uf_53 ?x462 (uf_43 ?x464 ?x463)) uf_152))) :pat { (uf_53 ?x462 (uf_43 ?x464 ?x463)) (uf_141 ?x464) }) +:assumption (forall (?x465 T4) (?x466 T4) (?x467 T5) (?x468 T3) (implies (and (= (uf_15 ?x467) ?x468) (= (uf_141 ?x468) uf_9)) (and (= (uf_179 ?x465 ?x466 ?x467 ?x468) uf_9) (iff (= (uf_46 ?x465 ?x466 ?x467 ?x468) uf_9) (= (uf_27 ?x466 ?x467) uf_9)))) :pat { (uf_141 ?x468) (uf_46 ?x465 ?x466 ?x467 ?x468) }) +:assumption (forall (?x469 T4) (?x470 T5) (?x471 T5) (implies (and (= (uf_22 (uf_15 ?x470)) uf_9) (and (= (uf_24 ?x469 ?x471) uf_9) (= (uf_51 ?x469) uf_9))) (iff (= (uf_13 ?x470 (uf_53 ?x469 ?x471)) uf_9) (= (uf_25 ?x469 ?x470) ?x471))) :pat { (uf_13 ?x470 (uf_53 ?x469 ?x471)) (uf_22 (uf_15 ?x470)) }) +:assumption (forall (?x472 T4) (?x473 T4) (?x474 Int) (?x475 T3) (?x476 T15) (up_182 (uf_19 (uf_20 ?x473) (uf_126 (uf_43 ?x475 ?x474) ?x476))) :pat { (uf_180 ?x475 ?x476) (uf_181 ?x472 ?x473) (uf_19 (uf_20 ?x472) (uf_126 (uf_43 ?x475 ?x474) ?x476)) }) +:assumption (forall (?x477 T4) (?x478 Int) (?x479 T3) (?x480 T15) (implies (and (= (uf_25 ?x477 (uf_43 ?x479 ?x478)) uf_26) (and (= (uf_180 ?x479 ?x480) uf_9) (and (= (uf_24 ?x477 (uf_43 ?x479 ?x478)) uf_9) (= (uf_55 ?x477) uf_9)))) (= (uf_19 (uf_20 ?x477) (uf_126 (uf_43 ?x479 ?x478) ?x480)) (uf_183 (uf_184 ?x477 (uf_43 ?x479 ?x478)) (uf_126 (uf_43 ?x479 ?x478) ?x480)))) :pat { (uf_180 ?x479 ?x480) (uf_19 (uf_20 ?x477) (uf_126 (uf_43 ?x479 ?x478) ?x480)) }) +:assumption (forall (?x481 T4) (?x482 Int) (?x483 T3) (?x484 T15) (?x485 T15) (implies (and (or (= (uf_28 (uf_183 (uf_184 ?x481 (uf_43 ?x483 ?x482)) (uf_126 (uf_43 ?x483 ?x482) ?x484))) uf_26) (= (uf_28 (uf_19 (uf_20 ?x481) (uf_126 (uf_43 ?x483 ?x482) ?x484))) uf_26)) (and (= (uf_24 ?x481 (uf_43 ?x483 ?x482)) uf_9) (and (= (uf_185 ?x483 ?x484 ?x485) uf_9) (= (uf_55 ?x481) uf_9)))) (= (uf_19 (uf_20 ?x481) (uf_126 (uf_43 ?x483 ?x482) ?x485)) (uf_183 (uf_184 ?x481 (uf_43 ?x483 ?x482)) (uf_126 (uf_43 ?x483 ?x482) ?x485)))) :pat { (uf_185 ?x483 ?x484 ?x485) (uf_19 (uf_20 ?x481) (uf_126 (uf_43 ?x483 ?x482) ?x485)) }) +:assumption (forall (?x486 T4) (?x487 T5) (= (uf_184 ?x486 ?x487) (uf_30 (uf_19 (uf_20 ?x486) ?x487))) :pat { (uf_184 ?x486 ?x487) }) +:assumption (forall (?x488 T4) (?x489 T5) (?x490 T5) (?x491 T15) (?x492 Int) (?x493 Int) (?x494 T3) (implies (and (< ?x492 ?x493) (and (<= 0 ?x492) (and (= (uf_187 ?x491 ?x493) uf_9) (and (= (uf_186 ?x489 ?x490) uf_9) (and (= (uf_24 ?x488 ?x490) uf_9) (= (uf_51 ?x488) uf_9)))))) (= (uf_19 (uf_20 ?x488) (uf_66 (uf_126 ?x489 ?x491) ?x492 ?x494)) (uf_11 (uf_189 ?x490) (uf_66 (uf_126 ?x489 ?x491) ?x492 ?x494)))) :pat { (uf_49 ?x488 ?x490) (uf_186 ?x489 ?x490) (uf_19 (uf_20 ?x488) (uf_66 (uf_126 ?x489 ?x491) ?x492 ?x494)) (uf_187 ?x491 ?x493) } :pat { (uf_188 ?x488 ?x490 ?x489 (uf_66 (uf_126 ?x489 ?x491) ?x492 ?x494)) (uf_187 ?x491 ?x493) }) +:assumption (forall (?x495 T4) (?x496 T5) (?x497 T5) (?x498 T15) (implies (and (= (uf_190 ?x498) uf_9) (and (= (uf_186 ?x496 ?x497) uf_9) (and (= (uf_24 ?x495 ?x497) uf_9) (= (uf_51 ?x495) uf_9)))) (and (= (uf_19 (uf_20 ?x495) (uf_126 ?x496 ?x498)) (uf_11 (uf_189 ?x497) (uf_126 ?x496 ?x498))) (= (uf_186 ?x496 ?x497) uf_9))) :pat { (uf_186 ?x496 ?x497) (uf_19 (uf_20 ?x495) (uf_126 ?x496 ?x498)) } :pat { (uf_188 ?x495 ?x497 ?x496 (uf_126 ?x496 ?x498)) }) +:assumption (forall (?x499 T4) (?x500 T5) (?x501 T5) (?x502 T5) (= (uf_188 ?x499 ?x500 ?x501 ?x502) ?x502) :pat { (uf_188 ?x499 ?x500 ?x501 ?x502) }) +:assumption (forall (?x503 T5) (?x504 T5) (implies (forall (?x505 T4) (implies (= (uf_49 ?x505 ?x504) uf_9) (= (uf_24 ?x505 ?x503) uf_9)) :pat { (uf_191 ?x505) }) (= (uf_186 ?x503 ?x504) uf_9)) :pat { (uf_186 ?x503 ?x504) }) +:assumption (forall (?x506 T5) (?x507 T4) (?x508 T4) (?x509 T5) (up_193 (uf_13 ?x509 (uf_192 (uf_12 ?x508 ?x506)))) :pat { (uf_13 ?x509 (uf_192 (uf_12 ?x507 ?x506))) (uf_177 ?x507 ?x508) }) +:assumption (forall (?x510 T5) (?x511 T4) (?x512 T4) (?x513 T5) (up_193 (uf_13 ?x513 (uf_10 ?x512 ?x510))) :pat { (uf_13 ?x513 (uf_10 ?x511 ?x510)) (uf_177 ?x511 ?x512) }) +:assumption (forall (?x514 T4) (?x515 T5) (?x516 T15) (?x517 Int) (?x518 Int) (?x519 T3) (implies (and (< ?x518 ?x517) (and (<= 0 ?x518) (and (= (uf_194 ?x516 ?x517 ?x519) uf_9) (= (uf_51 ?x514) uf_9)))) (= (uf_135 (uf_58 (uf_59 ?x514) (uf_66 (uf_126 ?x515 ?x516) ?x518 ?x519))) uf_9)) :pat { (uf_194 ?x516 ?x517 ?x519) (uf_135 (uf_58 (uf_59 ?x514) (uf_66 (uf_126 ?x515 ?x516) ?x518 ?x519))) }) +:assumption (forall (?x520 T4) (?x521 Int) (?x522 T5) (?x523 Int) (?x524 Int) (?x525 T3) (implies (and (< ?x524 ?x523) (and (<= 0 ?x524) (and (= (uf_13 (uf_43 (uf_124 ?x525 ?x523) ?x521) (uf_10 ?x520 ?x522)) uf_9) (and (= (uf_23 ?x525) uf_9) (= (uf_55 ?x520) uf_9))))) (= (uf_19 (uf_20 ?x520) (uf_66 (uf_43 ?x525 ?x521) ?x524 ?x525)) (uf_11 (uf_12 ?x520 ?x522) (uf_66 (uf_43 ?x525 ?x521) ?x524 ?x525)))) :pat { (uf_13 (uf_43 (uf_124 ?x525 ?x523) ?x521) (uf_10 ?x520 ?x522)) (uf_19 (uf_20 ?x520) (uf_66 (uf_43 ?x525 ?x521) ?x524 ?x525)) (uf_23 ?x525) }) +:assumption (forall (?x526 T4) (?x527 Int) (?x528 T5) (?x529 Int) (?x530 Int) (?x531 T3) (implies (and (< ?x530 ?x529) (and (<= 0 ?x530) (and (= (uf_13 (uf_43 (uf_124 ?x531 ?x529) ?x527) (uf_10 ?x526 ?x528)) uf_9) (and (= (uf_23 ?x531) uf_9) (= (uf_55 ?x526) uf_9))))) (and (not (= (uf_135 (uf_58 (uf_59 ?x526) (uf_66 (uf_43 ?x531 ?x527) ?x530 ?x531))) uf_9)) (= (uf_27 ?x526 (uf_66 (uf_43 ?x531 ?x527) ?x530 ?x531)) uf_9))) :pat { (uf_13 (uf_43 (uf_124 ?x531 ?x529) ?x527) (uf_10 ?x526 ?x528)) (uf_58 (uf_59 ?x526) (uf_66 (uf_43 ?x531 ?x527) ?x530 ?x531)) (uf_23 ?x531) } :pat { (uf_13 (uf_43 (uf_124 ?x531 ?x529) ?x527) (uf_10 ?x526 ?x528)) (uf_25 ?x526 (uf_66 (uf_43 ?x531 ?x527) ?x530 ?x531)) (uf_23 ?x531) }) +:assumption (forall (?x532 T4) (?x533 T5) (?x534 T5) (?x535 T15) (?x536 Int) (?x537 Int) (?x538 T3) (implies (and (< ?x537 ?x536) (and (<= 0 ?x537) (and (= (uf_187 ?x535 ?x536) uf_9) (and (= (uf_13 ?x533 (uf_10 ?x532 ?x534)) uf_9) (= (uf_55 ?x532) uf_9))))) (and (not (= (uf_135 (uf_58 (uf_59 ?x532) (uf_66 (uf_126 ?x533 ?x535) ?x537 ?x538))) uf_9)) (= (uf_27 ?x532 (uf_66 (uf_126 ?x533 ?x535) ?x537 ?x538)) uf_9))) :pat { (uf_13 ?x533 (uf_10 ?x532 ?x534)) (uf_187 ?x535 ?x536) (uf_58 (uf_59 ?x532) (uf_66 (uf_126 ?x533 ?x535) ?x537 ?x538)) } :pat { (uf_13 ?x533 (uf_10 ?x532 ?x534)) (uf_187 ?x535 ?x536) (uf_25 ?x532 (uf_66 (uf_126 ?x533 ?x535) ?x537 ?x538)) }) +:assumption (forall (?x539 T4) (?x540 T5) (?x541 T5) (?x542 T15) (?x543 Int) (?x544 Int) (?x545 T3) (implies (and (< ?x544 ?x543) (and (<= 0 ?x544) (and (= (uf_187 ?x542 ?x543) uf_9) (and (= (uf_13 ?x540 (uf_10 ?x539 ?x541)) uf_9) (= (uf_55 ?x539) uf_9))))) (= (uf_19 (uf_20 ?x539) (uf_66 (uf_126 ?x540 ?x542) ?x544 ?x545)) (uf_11 (uf_12 ?x539 ?x541) (uf_66 (uf_126 ?x540 ?x542) ?x544 ?x545)))) :pat { (uf_13 ?x540 (uf_10 ?x539 ?x541)) (uf_187 ?x542 ?x543) (uf_19 (uf_20 ?x539) (uf_66 (uf_126 ?x540 ?x542) ?x544 ?x545)) }) +:assumption (forall (?x546 T4) (?x547 T5) (?x548 T5) (?x549 T15) (implies (and (= (uf_190 ?x549) uf_9) (and (= (uf_13 ?x547 (uf_10 ?x546 ?x548)) uf_9) (= (uf_55 ?x546) uf_9))) (and (not (= (uf_135 (uf_58 (uf_59 ?x546) (uf_126 ?x547 ?x549))) uf_9)) (= (uf_27 ?x546 (uf_126 ?x547 ?x549)) uf_9))) :pat { (uf_13 ?x547 (uf_10 ?x546 ?x548)) (uf_190 ?x549) (uf_25 ?x546 (uf_126 ?x547 ?x549)) } :pat { (uf_13 ?x547 (uf_10 ?x546 ?x548)) (uf_190 ?x549) (uf_58 (uf_59 ?x546) (uf_126 ?x547 ?x549)) }) +:assumption (forall (?x550 T4) (?x551 T5) (?x552 T5) (implies (and (= (uf_13 ?x551 (uf_10 ?x550 ?x552)) uf_9) (= (uf_55 ?x550) uf_9)) (and (not (= (uf_135 (uf_58 (uf_59 ?x550) ?x551)) uf_9)) (= (uf_27 ?x550 ?x551) uf_9))) :pat { (uf_55 ?x550) (uf_13 ?x551 (uf_10 ?x550 ?x552)) (uf_40 (uf_41 ?x550) ?x551) } :pat { (uf_55 ?x550) (uf_13 ?x551 (uf_10 ?x550 ?x552)) (uf_58 (uf_59 ?x550) ?x551) }) +:assumption (forall (?x553 T4) (?x554 T5) (?x555 T5) (?x556 T15) (implies (and (= (uf_190 ?x556) uf_9) (= (uf_13 ?x554 (uf_10 ?x553 ?x555)) uf_9)) (= (uf_19 (uf_20 ?x553) (uf_126 ?x554 ?x556)) (uf_11 (uf_12 ?x553 ?x555) (uf_126 ?x554 ?x556)))) :pat { (uf_13 ?x554 (uf_10 ?x553 ?x555)) (uf_190 ?x556) (uf_19 (uf_20 ?x553) (uf_126 ?x554 ?x556)) }) +:assumption (forall (?x557 T4) (?x558 T5) (?x559 T5) (implies (= (uf_195 ?x557 ?x558 ?x559) uf_9) (= (uf_196 ?x557 ?x558 ?x559) uf_9)) :pat { (uf_195 ?x557 ?x558 ?x559) }) +:assumption (forall (?x560 T4) (?x561 T5) (?x562 T5) (?x563 T5) (implies (and (forall (?x564 T4) (implies (and (= (uf_10 ?x564 ?x561) (uf_10 ?x560 ?x561)) (and (= (uf_12 ?x564 ?x561) (uf_12 ?x560 ?x561)) (= (uf_46 ?x564 ?x564 ?x562 (uf_15 ?x562)) uf_9))) (= (uf_145 ?x563 (uf_53 ?x564 ?x562)) uf_9))) (and (= (uf_13 ?x562 (uf_10 ?x560 ?x561)) uf_9) (up_197 (uf_15 ?x562)))) (and (= (uf_145 ?x563 (uf_53 ?x560 ?x562)) uf_9) (= (uf_195 ?x560 ?x563 ?x561) uf_9))) :pat { (uf_13 ?x562 (uf_10 ?x560 ?x561)) (uf_195 ?x560 ?x563 ?x561) }) +:assumption (forall (?x565 T4) (?x566 T5) (?x567 T5) (?x568 T5) (implies (and (= (uf_145 ?x568 (uf_53 ?x565 ?x567)) uf_9) (and (= (uf_13 ?x567 (uf_10 ?x565 ?x566)) uf_9) (not (up_197 (uf_15 ?x567))))) (and (= (uf_145 ?x568 (uf_53 ?x565 ?x567)) uf_9) (= (uf_196 ?x565 ?x568 ?x566) uf_9))) :pat { (uf_13 ?x567 (uf_10 ?x565 ?x566)) (uf_196 ?x565 ?x568 ?x566) }) +:assumption (forall (?x569 T4) (?x570 T5) (?x571 T5) (implies (and (= (uf_13 ?x571 (uf_10 ?x569 ?x570)) uf_9) (= (uf_55 ?x569) uf_9)) (= (uf_196 ?x569 ?x571 ?x570) uf_9)) :pat { (uf_196 ?x569 ?x571 ?x570) }) +:assumption (forall (?x572 T4) (?x573 T5) (implies (and (= (uf_22 (uf_15 ?x573)) uf_9) (and (not (= (uf_14 (uf_15 ?x573)) uf_16)) (and (= (uf_27 ?x572 ?x573) uf_9) (and (= (uf_48 ?x573 (uf_15 ?x573)) uf_9) (and (= (uf_25 ?x572 ?x573) uf_26) (and (= (uf_24 ?x572 ?x573) uf_9) (= (uf_55 ?x572) uf_9))))))) (= (uf_196 ?x572 ?x573 ?x573) uf_9)) :pat { (uf_196 ?x572 ?x573 ?x573) }) +:assumption (forall (?x574 T4) (?x575 T5) (?x576 T5) (implies (= (uf_196 ?x574 ?x575 ?x576) uf_9) (and (forall (?x577 T5) (implies (and (= (uf_13 ?x577 (uf_53 ?x574 ?x575)) uf_9) (not (up_197 (uf_15 ?x575)))) (= (uf_147 ?x577 (uf_192 (uf_12 ?x574 ?x576))) uf_9)) :pat { (uf_13 ?x577 (uf_53 ?x574 ?x575)) }) (and (= (uf_24 ?x574 ?x575) uf_9) (= (uf_13 ?x575 (uf_10 ?x574 ?x576)) uf_9)))) :pat { (uf_196 ?x574 ?x575 ?x576) }) +:assumption (forall (?x578 T4) (?x579 T5) (?x580 T5) (?x581 T16) (iff (= (uf_198 ?x578 ?x579 ?x580 ?x581) uf_9) (= (uf_195 ?x578 ?x579 ?x580) uf_9)) :pat { (uf_198 ?x578 ?x579 ?x580 ?x581) }) +:assumption (forall (?x582 T4) (?x583 T5) (?x584 T5) (?x585 T16) (implies (= (uf_198 ?x582 ?x583 ?x584 ?x585) uf_9) (up_199 ?x582 ?x583 ?x585)) :pat { (uf_198 ?x582 ?x583 ?x584 ?x585) }) +:assumption (forall (?x586 T4) (?x587 T5) (?x588 T5) (?x589 T16) (iff (= (uf_200 ?x586 ?x587 ?x588 ?x589) uf_9) (= (uf_196 ?x586 ?x587 ?x588) uf_9)) :pat { (uf_200 ?x586 ?x587 ?x588 ?x589) }) +:assumption (forall (?x590 T4) (?x591 T5) (?x592 T5) (?x593 T16) (implies (= (uf_200 ?x590 ?x591 ?x592 ?x593) uf_9) (up_199 ?x590 ?x591 ?x593)) :pat { (uf_200 ?x590 ?x591 ?x592 ?x593) }) +:assumption (forall (?x594 T4) (?x595 T5) (= (uf_10 ?x594 ?x595) (uf_192 (uf_12 ?x594 ?x595))) :pat { (uf_10 ?x594 ?x595) }) +:assumption (forall (?x596 T4) (?x597 T5) (= (uf_12 ?x596 ?x597) (uf_32 (uf_19 (uf_20 ?x596) ?x597))) :pat { (uf_12 ?x596 ?x597) }) +:assumption (forall (?x598 T4) (?x599 Int) (?x600 T3) (= (uf_43 ?x600 (uf_19 (uf_20 ?x598) (uf_43 (uf_6 ?x600) ?x599))) (uf_201 ?x598 (uf_43 (uf_6 ?x600) ?x599) ?x600)) :pat { (uf_43 ?x600 (uf_19 (uf_20 ?x598) (uf_43 (uf_6 ?x600) ?x599))) }) +:assumption (forall (?x601 T1) (?x602 T4) (implies (= (uf_202 ?x601 ?x602) uf_9) (= (uf_51 ?x602) uf_9)) :pat { (uf_202 ?x601 ?x602) }) +:assumption (forall (?x603 T4) (implies (= (uf_44 ?x603) uf_9) (= (uf_51 ?x603) uf_9)) :pat { (uf_44 ?x603) }) +:assumption (forall (?x604 T4) (implies (= (uf_55 ?x604) uf_9) (and (= (uf_44 ?x604) uf_9) (= (uf_51 ?x604) uf_9))) :pat { (uf_55 ?x604) }) +:assumption (forall (?x605 T4) (implies (= (uf_203 ?x605) uf_9) (and (<= 0 (uf_171 ?x605)) (= (uf_55 ?x605) uf_9))) :pat { (uf_203 ?x605) }) +:assumption (forall (?x606 T3) (implies (= (uf_23 ?x606) uf_9) (forall (?x607 T4) (?x608 Int) (?x609 T5) (iff (= (uf_13 ?x609 (uf_128 ?x607 (uf_43 ?x606 ?x608))) uf_9) (= ?x609 (uf_43 ?x606 ?x608))) :pat { (uf_13 ?x609 (uf_128 ?x607 (uf_43 ?x606 ?x608))) })) :pat { (uf_23 ?x606) }) +:assumption (forall (?x610 T3) (implies (= (uf_23 ?x610) uf_9) (forall (?x611 Int) (?x612 T5) (iff (= (uf_13 ?x612 (uf_130 (uf_43 ?x610 ?x611))) uf_9) (= ?x612 (uf_43 ?x610 ?x611))) :pat { (uf_13 ?x612 (uf_130 (uf_43 ?x610 ?x611))) })) :pat { (uf_23 ?x610) }) +:assumption (forall (?x613 T4) (?x614 T4) (?x615 T5) (?x616 T3) (iff (= (uf_204 ?x613 ?x614 ?x615 ?x616) uf_9) (and (up_205 ?x613 ?x614 ?x615 ?x616) (and (= (uf_58 (uf_59 ?x613) ?x615) (uf_58 (uf_59 ?x614) ?x615)) (= (uf_12 ?x613 ?x615) (uf_12 ?x614 ?x615))))) :pat { (uf_204 ?x613 ?x614 ?x615 ?x616) }) +:assumption (forall (?x617 T4) (?x618 T4) (?x619 T5) (?x620 T3) (iff (= (uf_206 ?x617 ?x618 ?x619 ?x620) uf_9) (and (= (uf_123 ?x617 ?x618 ?x619 ?x620) uf_9) (and (= (uf_58 (uf_59 ?x617) ?x619) (uf_58 (uf_59 ?x618) ?x619)) (and (= (uf_53 ?x617 ?x619) (uf_53 ?x618 ?x619)) (= (uf_12 ?x617 ?x619) (uf_12 ?x618 ?x619)))))) :pat { (uf_206 ?x617 ?x618 ?x619 ?x620) }) +:assumption (forall (?x621 T4) (?x622 T4) (?x623 T5) (?x624 T5) (iff (= (uf_207 ?x621 ?x622 ?x623 ?x624) uf_9) (or (= (uf_208 (uf_15 ?x623)) uf_9) (or (and (= (uf_204 ?x621 ?x622 ?x623 (uf_15 ?x623)) uf_9) (= (uf_46 ?x621 ?x622 ?x623 (uf_15 ?x623)) uf_9)) (or (and (not (= (uf_24 ?x622 ?x623) uf_9)) (not (= (uf_24 ?x621 ?x623) uf_9))) (= (uf_206 ?x621 ?x622 ?x624 (uf_15 ?x624)) uf_9))))) :pat { (uf_207 ?x621 ?x622 ?x623 ?x624) }) +:assumption (forall (?x625 T4) (?x626 T4) (?x627 T5) (?x628 T3) (iff (= (uf_179 ?x625 ?x626 ?x627 ?x628) uf_9) (implies (and (= (uf_24 ?x626 ?x627) uf_9) (= (uf_24 ?x625 ?x627) uf_9)) (= (uf_206 ?x625 ?x626 ?x627 ?x628) uf_9))) :pat { (uf_179 ?x625 ?x626 ?x627 ?x628) }) +:assumption (forall (?x629 T4) (?x630 T5) (?x631 T3) (implies (up_209 ?x629 ?x630 ?x631) (= (uf_46 ?x629 ?x629 ?x630 ?x631) uf_9)) :pat { (uf_46 ?x629 ?x629 ?x630 ?x631) }) +:assumption (forall (?x632 T4) (?x633 T5) (iff (= (uf_67 ?x632 ?x633) uf_9) (and (or (and (or (= (uf_210 ?x632 ?x633) uf_9) (= (uf_25 ?x632 ?x633) uf_26)) (not (= (uf_14 (uf_15 ?x633)) uf_16))) (and (or (= (uf_210 ?x632 (uf_136 (uf_58 (uf_59 ?x632) ?x633))) uf_9) (= (uf_25 ?x632 (uf_136 (uf_58 (uf_59 ?x632) ?x633))) uf_26)) (and (not (= (uf_14 (uf_15 (uf_136 (uf_58 (uf_59 ?x632) ?x633)))) uf_16)) (and (or (not (= (uf_24 ?x632 (uf_136 (uf_58 (uf_59 ?x632) ?x633))) uf_9)) (not (= (uf_135 (uf_58 (uf_59 ?x632) ?x633)) uf_9))) (= (uf_14 (uf_15 ?x633)) uf_16))))) (= (uf_27 ?x632 ?x633) uf_9))) :pat { (uf_67 ?x632 ?x633) }) +:assumption (forall (?x634 T4) (?x635 T5) (iff (= (uf_210 ?x634 ?x635) uf_9) (exists (?x636 T5) (and (= (uf_211 ?x634 ?x636) uf_9) (and (= (uf_22 (uf_15 ?x636)) uf_9) (and (not (= (uf_14 (uf_15 ?x636)) uf_16)) (and (= (uf_27 ?x634 ?x636) uf_9) (and (= (uf_48 ?x636 (uf_15 ?x636)) uf_9) (and (= (uf_25 ?x634 ?x636) uf_26) (and (= (uf_24 ?x634 ?x636) uf_9) (= (uf_13 ?x635 (uf_192 (uf_12 ?x634 ?x636))) uf_9)))))))) :pat { (uf_147 ?x635 (uf_192 (uf_12 ?x634 ?x636))) })) :pat { (uf_210 ?x634 ?x635) }) +:assumption (forall (?x637 T4) (?x638 T5) (iff (= (uf_211 ?x637 ?x638) uf_9) true) :pat { (uf_211 ?x637 ?x638) }) +:assumption (forall (?x639 T4) (?x640 T4) (?x641 T5) (implies (= (uf_177 ?x639 ?x640) uf_9) (up_212 (uf_40 (uf_41 ?x639) ?x641))) :pat { (uf_40 (uf_41 ?x640) ?x641) (uf_177 ?x639 ?x640) }) +:assumption (forall (?x642 T4) (?x643 T5) (implies (and (= (uf_27 ?x642 ?x643) uf_9) (= (uf_51 ?x642) uf_9)) (< 0 (uf_116 ?x643))) :pat { (uf_27 ?x642 ?x643) }) +:assumption (forall (?x644 T4) (?x645 T5) (implies (= (uf_51 ?x644) uf_9) (iff (= (uf_27 ?x644 ?x645) uf_9) (up_213 (uf_58 (uf_59 ?x644) ?x645)))) :pat { (uf_27 ?x644 ?x645) }) +:assumption (forall (?x646 T4) (?x647 T5) (iff (= (uf_61 ?x646 ?x647) uf_9) (and (not (= (uf_24 ?x646 ?x647) uf_9)) (and (= (uf_25 ?x646 ?x647) uf_26) (= (uf_27 ?x646 ?x647) uf_9)))) :pat { (uf_61 ?x646 ?x647) }) +:assumption (forall (?x648 T4) (?x649 T5) (= (uf_53 ?x648 ?x649) (uf_34 (uf_19 (uf_20 ?x648) (uf_126 ?x649 (uf_214 (uf_15 ?x649)))))) :pat { (uf_53 ?x648 ?x649) }) +:assumption (forall (?x650 T11) (and (= (uf_22 (uf_15 (uf_215 ?x650))) uf_9) (not (= (uf_14 (uf_15 (uf_215 ?x650))) uf_16))) :pat { (uf_215 ?x650) }) +:assumption up_216 +:assumption (forall (?x651 T4) (?x652 T5) (implies (= (uf_22 (uf_15 ?x652)) uf_9) (= (uf_170 ?x651 ?x652) (uf_217 (uf_40 (uf_41 ?x651) ?x652)))) :pat { (uf_22 (uf_15 ?x652)) (uf_170 ?x651 ?x652) }) +:assumption (forall (?x653 T4) (?x654 T5) (implies (= (uf_23 (uf_15 ?x654)) uf_9) (= (uf_170 ?x653 ?x654) (uf_217 (uf_40 (uf_41 ?x653) (uf_136 (uf_58 (uf_59 ?x653) ?x654)))))) :pat { (uf_23 (uf_15 ?x654)) (uf_170 ?x653 ?x654) }) +:assumption (forall (?x655 T4) (?x656 T5) (implies (= (uf_22 (uf_15 ?x656)) uf_9) (iff (= (uf_24 ?x655 ?x656) uf_9) (up_218 (uf_40 (uf_41 ?x655) ?x656)))) :pat { (uf_22 (uf_15 ?x656)) (uf_24 ?x655 ?x656) }) +:assumption (forall (?x657 T4) (?x658 T5) (implies (= (uf_23 (uf_15 ?x658)) uf_9) (iff (= (uf_24 ?x657 ?x658) uf_9) (up_218 (uf_40 (uf_41 ?x657) (uf_136 (uf_58 (uf_59 ?x657) ?x658)))))) :pat { (uf_23 (uf_15 ?x658)) (uf_24 ?x657 ?x658) }) +:assumption (forall (?x659 T4) (?x660 T5) (implies (= (uf_22 (uf_15 ?x660)) uf_9) (= (uf_25 ?x659 ?x660) (uf_215 (uf_40 (uf_41 ?x659) ?x660)))) :pat { (uf_22 (uf_15 ?x660)) (uf_25 ?x659 ?x660) }) +:assumption (forall (?x661 T4) (?x662 T5) (implies (= (uf_23 (uf_15 ?x662)) uf_9) (= (uf_25 ?x661 ?x662) (uf_25 ?x661 (uf_136 (uf_58 (uf_59 ?x661) ?x662))))) :pat { (uf_23 (uf_15 ?x662)) (uf_25 ?x661 ?x662) }) +:assumption (forall (?x663 T5) (?x664 T3) (= (uf_126 ?x663 (uf_214 ?x664)) (uf_43 uf_219 (uf_220 ?x663 (uf_214 ?x664)))) :pat { (uf_126 ?x663 (uf_214 ?x664)) }) +:assumption (up_197 uf_37) +:assumption (forall (?x665 T17) (?x666 T17) (?x667 T15) (implies (= (uf_224 (uf_225 (uf_222 ?x665 ?x667)) (uf_225 (uf_222 ?x666 ?x667))) uf_9) (= (uf_221 (uf_222 ?x665 ?x667) (uf_222 ?x666 ?x667)) uf_9)) :pat { (uf_221 (uf_222 ?x665 ?x667) (uf_222 ?x666 (uf_223 ?x667))) }) +:assumption (forall (?x668 T17) (?x669 T17) (implies (forall (?x670 T15) (= (uf_221 (uf_222 ?x668 ?x670) (uf_222 ?x669 ?x670)) uf_9)) (= (uf_224 ?x668 ?x669) uf_9)) :pat { (uf_224 ?x668 ?x669) }) +:assumption (forall (?x671 T17) (= (uf_225 (uf_226 ?x671)) ?x671)) +:assumption (forall (?x672 Int) (?x673 Int) (iff (= (uf_221 ?x672 ?x673) uf_9) (= ?x672 ?x673)) :pat { (uf_221 ?x672 ?x673) }) +:assumption (forall (?x674 T17) (?x675 T17) (iff (= (uf_224 ?x674 ?x675) uf_9) (= ?x674 ?x675)) :pat { (uf_224 ?x674 ?x675) }) +:assumption (forall (?x676 T3) (?x677 T15) (?x678 T3) (implies (and (= (uf_228 ?x678) uf_9) (= (uf_227 ?x676 ?x677 ?x678) uf_9)) (= (uf_223 ?x677) ?x677)) :pat { (uf_227 ?x676 ?x677 ?x678) (uf_228 ?x678) }) +:assumption (forall (?x679 T3) (implies (= (uf_228 ?x679) uf_9) (= (uf_23 ?x679) uf_9)) :pat { (uf_228 ?x679) }) +:assumption (forall (?x680 T17) (?x681 T15) (?x682 T15) (?x683 Int) (or (= ?x681 ?x682) (= (uf_222 (uf_229 ?x680 ?x681 ?x683) ?x682) (uf_222 ?x680 ?x682))) :pat { (uf_222 (uf_229 ?x680 ?x681 ?x683) ?x682) }) +:assumption (forall (?x684 T17) (?x685 T15) (?x686 Int) (= (uf_222 (uf_229 ?x684 ?x685 ?x686) ?x685) ?x686) :pat { (uf_222 (uf_229 ?x684 ?x685 ?x686) ?x685) }) +:assumption (forall (?x687 T15) (= (uf_222 uf_230 ?x687) 0)) +:assumption (forall (?x688 T17) (?x689 T15) (?x690 Int) (?x691 Int) (?x692 Int) (?x693 Int) (= (uf_231 ?x688 ?x689 ?x690 ?x691 ?x692 ?x693) (uf_229 ?x688 ?x689 (uf_99 (uf_222 ?x688 ?x689) ?x690 ?x691 ?x692 ?x693))) :pat { (uf_231 ?x688 ?x689 ?x690 ?x691 ?x692 ?x693) }) +:assumption (forall (?x694 T4) (?x695 T5) (implies (= (uf_51 ?x694) uf_9) (and (= (uf_233 (uf_232 ?x694 ?x695)) ?x694) (= (uf_234 (uf_232 ?x694 ?x695)) (uf_116 ?x695)))) :pat { (uf_232 ?x694 ?x695) }) +:assumption (forall (?x696 T18) (= (uf_51 (uf_233 ?x696)) uf_9)) +:assumption (= (uf_51 (uf_233 uf_235)) uf_9) +:assumption (forall (?x697 T4) (?x698 T5) (or (not (up_213 (uf_58 (uf_59 ?x697) ?x698))) (<= (uf_170 ?x697 ?x698) (uf_171 ?x697))) :pat { (uf_40 (uf_41 ?x697) ?x698) }) +:assumption (forall (?x699 T4) (?x700 T5) (implies (and (= (uf_135 (uf_58 (uf_59 ?x699) ?x700)) uf_9) (= (uf_51 ?x699) uf_9)) (= (uf_14 (uf_15 ?x700)) uf_16)) :pat { (uf_135 (uf_58 (uf_59 ?x699) ?x700)) }) +:assumption (forall (?x701 T4) (?x702 T5) (implies (= (uf_27 ?x701 ?x702) uf_9) (= (uf_27 ?x701 (uf_136 (uf_58 (uf_59 ?x701) ?x702))) uf_9)) :pat { (uf_27 ?x701 ?x702) (uf_58 (uf_59 ?x701) (uf_136 (uf_58 (uf_59 ?x701) ?x702))) }) +:assumption (forall (?x703 T14) (and (= (uf_22 (uf_15 (uf_136 ?x703))) uf_9) (not (= (uf_14 (uf_15 (uf_136 ?x703))) uf_16))) :pat { (uf_136 ?x703) }) +:assumption (forall (?x704 T5) (?x705 T15) (implies (<= 0 (uf_237 ?x705)) (= (uf_116 (uf_126 (uf_236 ?x704 ?x705) ?x705)) (uf_116 ?x704))) :pat { (uf_126 (uf_236 ?x704 ?x705) ?x705) }) +:assumption (forall (?x706 T5) (?x707 T15) (= (uf_236 ?x706 ?x707) (uf_43 (uf_238 ?x707) (uf_239 ?x706 ?x707))) :pat { (uf_236 ?x706 ?x707) }) +:assumption (forall (?x708 Int) (?x709 T15) (= (uf_236 (uf_126 (uf_43 (uf_238 ?x709) ?x708) ?x709) ?x709) (uf_43 (uf_238 ?x709) ?x708)) :pat { (uf_236 (uf_126 (uf_43 (uf_238 ?x709) ?x708) ?x709) ?x709) }) +:assumption (forall (?x710 T5) (?x711 T3) (implies (= (uf_48 ?x710 ?x711) uf_9) (= ?x710 (uf_43 ?x711 (uf_116 ?x710)))) :pat { (uf_48 ?x710 ?x711) }) +:assumption (forall (?x712 T5) (?x713 T3) (iff (= (uf_48 ?x712 ?x713) uf_9) (= (uf_15 ?x712) ?x713))) +:assumption (= uf_121 (uf_43 uf_240 0)) +:assumption (forall (?x714 T15) (?x715 Int) (and (= (uf_242 (uf_241 ?x714 ?x715)) ?x715) (and (= (uf_243 (uf_241 ?x714 ?x715)) ?x714) (not (up_244 (uf_241 ?x714 ?x715))))) :pat { (uf_241 ?x714 ?x715) }) +:assumption (forall (?x716 T5) (?x717 T15) (and (= (uf_245 (uf_220 ?x716 ?x717)) ?x717) (= (uf_246 (uf_220 ?x716 ?x717)) ?x716)) :pat { (uf_220 ?x716 ?x717) }) +:assumption (forall (?x718 T3) (?x719 Int) (= (uf_116 (uf_43 ?x718 ?x719)) ?x719)) +:assumption (forall (?x720 T3) (?x721 Int) (= (uf_15 (uf_43 ?x720 ?x721)) ?x720)) +:assumption (forall (?x722 T3) (?x723 T3) (?x724 Int) (?x725 Int) (iff (= (uf_247 ?x722 ?x723 ?x724 ?x725) uf_9) (and (= (uf_248 ?x722 ?x723) ?x725) (and (= (uf_249 ?x722 ?x723) ?x724) (up_250 ?x722 ?x723)))) :pat { (uf_247 ?x722 ?x723 ?x724 ?x725) }) +:assumption (forall (?x726 T5) (= (uf_139 ?x726 ?x726) uf_9) :pat { (uf_15 ?x726) }) +:assumption (forall (?x727 T5) (?x728 T5) (?x729 T5) (implies (and (= (uf_139 ?x728 ?x729) uf_9) (= (uf_139 ?x727 ?x728) uf_9)) (= (uf_139 ?x727 ?x729) uf_9)) :pat { (uf_139 ?x727 ?x728) (uf_139 ?x728 ?x729) }) +:assumption (forall (?x730 T12) (?x731 T5) (?x732 T5) (?x733 T11) (or (= (uf_40 (uf_172 ?x730 ?x731 ?x733) ?x732) (uf_40 ?x730 ?x732)) (= ?x731 ?x732))) +:assumption (forall (?x734 T12) (?x735 T5) (?x736 T11) (= (uf_40 (uf_172 ?x734 ?x735 ?x736) ?x735) ?x736)) +:assumption (forall (?x737 T13) (?x738 T5) (?x739 T5) (?x740 T14) (or (= (uf_58 (uf_251 ?x737 ?x738 ?x740) ?x739) (uf_58 ?x737 ?x739)) (= ?x738 ?x739))) +:assumption (forall (?x741 T13) (?x742 T5) (?x743 T14) (= (uf_58 (uf_251 ?x741 ?x742 ?x743) ?x742) ?x743)) +:assumption (forall (?x744 T9) (?x745 T5) (?x746 T5) (?x747 Int) (or (= (uf_19 (uf_178 ?x744 ?x745 ?x747) ?x746) (uf_19 ?x744 ?x746)) (= ?x745 ?x746))) +:assumption (forall (?x748 T9) (?x749 T5) (?x750 Int) (= (uf_19 (uf_178 ?x748 ?x749 ?x750) ?x749) ?x750)) +:assumption (= uf_26 (uf_43 uf_252 uf_253)) +:assumption (= (uf_23 uf_254) uf_9) +:assumption (= (uf_23 uf_255) uf_9) +:assumption (= (uf_23 uf_84) uf_9) +:assumption (= (uf_23 uf_4) uf_9) +:assumption (= (uf_23 uf_91) uf_9) +:assumption (= (uf_23 uf_7) uf_9) +:assumption (= (uf_23 uf_83) uf_9) +:assumption (= (uf_23 uf_87) uf_9) +:assumption (= (uf_23 uf_90) uf_9) +:assumption (= (uf_23 uf_94) uf_9) +:assumption (= (uf_208 uf_252) uf_9) +:assumption (= (uf_23 uf_256) uf_9) +:assumption (= (uf_23 uf_219) uf_9) +:assumption (= (uf_23 uf_257) uf_9) +:assumption (= (uf_23 uf_258) uf_9) +:assumption (= (uf_23 uf_240) uf_9) +:assumption (forall (?x751 T3) (implies (= (uf_23 ?x751) uf_9) (not (up_36 ?x751))) :pat { (uf_23 ?x751) }) +:assumption (forall (?x752 T3) (= (uf_23 (uf_6 ?x752)) uf_9) :pat { (uf_6 ?x752) }) +:assumption (forall (?x753 T3) (?x754 T3) (= (uf_23 (uf_259 ?x753 ?x754)) uf_9) :pat { (uf_259 ?x753 ?x754) }) +:assumption (forall (?x755 T3) (implies (= (uf_208 ?x755) uf_9) (= (uf_22 ?x755) uf_9)) :pat { (uf_208 ?x755) }) +:assumption (forall (?x756 T3) (implies (= (uf_141 ?x756) uf_9) (= (uf_22 ?x756) uf_9)) :pat { (uf_141 ?x756) }) +:assumption (forall (?x757 T3) (implies (= (uf_260 ?x757) uf_9) (= (uf_22 ?x757) uf_9)) :pat { (uf_260 ?x757) }) +:assumption (forall (?x758 T3) (iff (= (uf_208 ?x758) uf_9) (= (uf_14 ?x758) uf_261)) :pat { (uf_208 ?x758) }) +:assumption (forall (?x759 T3) (iff (= (uf_141 ?x759) uf_9) (= (uf_14 ?x759) uf_262)) :pat { (uf_141 ?x759) }) +:assumption (forall (?x760 T3) (iff (= (uf_260 ?x760) uf_9) (= (uf_14 ?x760) uf_263)) :pat { (uf_260 ?x760) }) +:assumption (forall (?x761 T3) (iff (= (uf_23 ?x761) uf_9) (= (uf_14 ?x761) uf_16)) :pat { (uf_23 ?x761) }) +:assumption (forall (?x762 T3) (?x763 T3) (= (uf_142 (uf_259 ?x762 ?x763)) (+ (uf_142 ?x762) 23)) :pat { (uf_259 ?x762 ?x763) }) +:assumption (forall (?x764 T3) (= (uf_142 (uf_6 ?x764)) (+ (uf_142 ?x764) 17)) :pat { (uf_6 ?x764) }) +:assumption (forall (?x765 T3) (?x766 T3) (= (uf_264 (uf_259 ?x765 ?x766)) ?x765) :pat { (uf_259 ?x765 ?x766) }) +:assumption (forall (?x767 T3) (?x768 T3) (= (uf_265 (uf_259 ?x767 ?x768)) ?x768) :pat { (uf_259 ?x767 ?x768) }) +:assumption (forall (?x769 T3) (= (uf_138 (uf_6 ?x769)) 8) :pat { (uf_6 ?x769) }) +:assumption (forall (?x770 T3) (= (uf_266 (uf_6 ?x770)) ?x770) :pat { (uf_6 ?x770) }) +:assumption (= (uf_260 uf_267) uf_9) +:assumption (= (uf_260 uf_37) uf_9) +:assumption (= (uf_142 uf_268) 0) +:assumption (= (uf_142 uf_256) 0) +:assumption (= (uf_142 uf_252) 0) +:assumption (= (uf_142 uf_219) 0) +:assumption (= (uf_142 uf_267) 0) +:assumption (= (uf_142 uf_37) 0) +:assumption (= (uf_142 uf_240) 0) +:assumption (= (uf_142 uf_258) 0) +:assumption (= (uf_142 uf_257) 0) +:assumption (= (uf_142 uf_254) 0) +:assumption (= (uf_142 uf_255) 0) +:assumption (= (uf_142 uf_84) 0) +:assumption (= (uf_142 uf_4) 0) +:assumption (= (uf_142 uf_91) 0) +:assumption (= (uf_142 uf_7) 0) +:assumption (= (uf_142 uf_83) 0) +:assumption (= (uf_142 uf_87) 0) +:assumption (= (uf_142 uf_90) 0) +:assumption (= (uf_142 uf_94) 0) +:assumption (= (uf_138 uf_219) 1) +:assumption (= (uf_138 uf_252) 1) +:assumption (= (uf_138 uf_254) 8) +:assumption (= (uf_138 uf_255) 4) +:assumption (= (uf_138 uf_84) 8) +:assumption (= (uf_138 uf_4) 4) +:assumption (= (uf_138 uf_91) 2) +:assumption (= (uf_138 uf_7) 1) +:assumption (= (uf_138 uf_83) 8) +:assumption (= (uf_138 uf_87) 4) +:assumption (= (uf_138 uf_90) 2) +:assumption (= (uf_138 uf_94) 1) +:assumption (not (implies true (implies (and (<= uf_269 uf_78) (<= 0 uf_269)) (implies (and (<= uf_270 uf_76) (<= 0 uf_270)) (implies (and (<= uf_271 uf_76) (<= 0 uf_271)) (implies (< uf_272 1099511627776) (implies (< 0 uf_272) (implies (and (= (uf_22 (uf_124 uf_7 uf_272)) uf_9) (and (not (= (uf_14 (uf_124 uf_7 uf_272)) uf_16)) (and (= (uf_27 uf_273 (uf_43 (uf_124 uf_7 uf_272) (uf_116 (uf_43 uf_7 uf_274)))) uf_9) (and (= (uf_48 (uf_43 (uf_124 uf_7 uf_272) (uf_116 (uf_43 uf_7 uf_274))) (uf_124 uf_7 uf_272)) uf_9) (and (= (uf_25 uf_273 (uf_43 (uf_124 uf_7 uf_272) (uf_116 (uf_43 uf_7 uf_274)))) uf_26) (= (uf_24 uf_273 (uf_43 (uf_124 uf_7 uf_272) (uf_116 (uf_43 uf_7 uf_274)))) uf_9)))))) (implies true (implies (= (uf_203 uf_273) uf_9) (implies (and (= (uf_55 uf_273) uf_9) (= (uf_202 uf_275 uf_273) uf_9)) (implies (forall (?x771 T19) (< (uf_276 ?x771) uf_277) :pat { (uf_276 ?x771) }) (implies (and (up_278 uf_273 uf_275 uf_279 (uf_43 uf_7 uf_274) (uf_6 uf_7)) (up_280 uf_273 uf_275 uf_279 (uf_29 (uf_43 uf_7 uf_274)) (uf_6 uf_7))) (implies (up_280 uf_273 uf_275 uf_281 uf_272 uf_4) (implies (= uf_282 (uf_171 uf_273)) (implies (forall (?x772 T5) (iff (= (uf_283 uf_282 ?x772) uf_9) false) :pat { (uf_283 uf_282 ?x772) }) (implies (and (<= uf_272 uf_76) (<= 0 uf_272)) (and (implies (= (uf_200 uf_273 (uf_43 (uf_124 uf_7 uf_272) (uf_116 (uf_43 uf_7 uf_274))) (uf_43 (uf_124 uf_7 uf_272) (uf_116 (uf_43 uf_7 uf_274))) uf_284) uf_9) (and (implies (and (= (uf_27 uf_273 (uf_66 (uf_43 uf_7 uf_274) 0 uf_7)) uf_9) (= (uf_48 (uf_66 (uf_43 uf_7 uf_274) 0 uf_7) uf_7) uf_9)) (and (implies (and (= (uf_67 uf_273 (uf_66 (uf_43 uf_7 uf_274) 0 uf_7)) uf_9) (= (uf_48 (uf_66 (uf_43 uf_7 uf_274) 0 uf_7) uf_7) uf_9)) (implies (= uf_285 (uf_110 uf_273 (uf_66 (uf_43 uf_7 uf_274) 0 uf_7))) (implies (up_280 uf_273 uf_286 uf_287 uf_285 uf_7) (implies (up_280 uf_273 uf_288 uf_289 0 uf_4) (implies (up_280 uf_273 uf_290 uf_291 1 uf_4) (implies (and (<= 0 0) (and (<= 0 0) (and (<= 1 1) (<= 1 1)))) (and (implies (<= 1 uf_272) (and (implies (forall (?x773 Int) (implies (and (<= ?x773 uf_76) (<= 0 ?x773)) (implies (< ?x773 1) (<= (uf_110 uf_273 (uf_66 (uf_43 uf_7 uf_274) ?x773 uf_7)) uf_285)))) (and (implies (and (= (uf_110 uf_273 (uf_66 (uf_43 uf_7 uf_274) 0 uf_7)) uf_285) (< 0 uf_272)) (implies true (implies (and (<= uf_292 uf_78) (<= 0 uf_292)) (implies (and (<= uf_293 uf_76) (<= 0 uf_293)) (implies (and (<= uf_294 uf_76) (<= 0 uf_294)) (implies (and (<= 0 uf_293) (<= 1 uf_294)) (implies (<= uf_294 uf_272) (implies (forall (?x774 Int) (implies (and (<= ?x774 uf_76) (<= 0 ?x774)) (implies (< ?x774 uf_294) (<= (uf_110 uf_273 (uf_66 (uf_43 uf_7 uf_274) ?x774 uf_7)) uf_292)))) (implies (and (= (uf_110 uf_273 (uf_66 (uf_43 uf_7 uf_274) uf_293 uf_7)) uf_292) (< uf_293 uf_272)) (implies (and (<= 0 uf_293) (<= 1 uf_294)) (and (implies true (implies (and (<= 0 uf_293) (<= 1 uf_294)) (implies true (implies (and (<= 0 uf_293) (<= 1 uf_294)) (implies true (implies (and (<= 0 uf_293) (<= 1 uf_294)) (implies (and (= (uf_177 uf_273 uf_273) uf_9) (and (forall (?x775 T5) (<= (uf_170 uf_273 ?x775) (uf_170 uf_273 ?x775)) :pat { (uf_170 uf_273 ?x775) }) (and (<= (uf_171 uf_273) (uf_171 uf_273)) (and (forall (?x776 T5) (implies (= (uf_67 uf_273 ?x776) uf_9) (and (= (uf_67 uf_273 ?x776) uf_9) (= (uf_58 (uf_59 uf_273) ?x776) (uf_58 (uf_59 uf_273) ?x776)))) :pat { (uf_58 (uf_59 uf_273) ?x776) }) (and (forall (?x777 T5) (implies (= (uf_67 uf_273 ?x777) uf_9) (and (= (uf_67 uf_273 ?x777) uf_9) (= (uf_40 (uf_41 uf_273) ?x777) (uf_40 (uf_41 uf_273) ?x777)))) :pat { (uf_40 (uf_41 uf_273) ?x777) }) (and (forall (?x778 T5) (implies (= (uf_67 uf_273 ?x778) uf_9) (and (= (uf_67 uf_273 ?x778) uf_9) (= (uf_19 (uf_20 uf_273) ?x778) (uf_19 (uf_20 uf_273) ?x778)))) :pat { (uf_19 (uf_20 uf_273) ?x778) }) (forall (?x779 T5) (implies (not (= (uf_14 (uf_15 (uf_25 uf_273 ?x779))) uf_261)) (not (= (uf_14 (uf_15 (uf_25 uf_273 ?x779))) uf_261))) :pat { (uf_40 (uf_41 uf_273) ?x779) }))))))) (implies (and (= (uf_177 uf_273 uf_273) uf_9) (and (forall (?x780 T5) (<= (uf_170 uf_273 ?x780) (uf_170 uf_273 ?x780)) :pat { (uf_170 uf_273 ?x780) }) (<= (uf_171 uf_273) (uf_171 uf_273)))) (implies (and (= (uf_55 uf_273) uf_9) (= (uf_202 uf_295 uf_273) uf_9)) (implies (up_280 uf_273 uf_295 uf_291 uf_294 uf_4) (implies (up_280 uf_273 uf_295 uf_289 uf_293 uf_4) (implies (up_280 uf_273 uf_295 uf_287 uf_292 uf_7) (implies (up_280 uf_273 uf_295 uf_281 uf_272 uf_4) (implies (and (up_278 uf_273 uf_295 uf_279 (uf_43 uf_7 uf_274) (uf_6 uf_7)) (up_280 uf_273 uf_295 uf_279 (uf_29 (uf_43 uf_7 uf_274)) (uf_6 uf_7))) (implies (and (= (uf_41 uf_273) (uf_41 uf_273)) (= (uf_59 uf_273) (uf_59 uf_273))) (implies (and (<= 0 uf_293) (<= 1 uf_294)) (and (implies true (implies (and (<= 0 uf_293) (<= 1 uf_294)) (implies (<= uf_272 uf_294) (implies (and (<= 0 uf_293) (<= 1 uf_294)) (implies true (implies (and (<= 0 uf_293) (<= 1 uf_294)) (implies (and (<= 0 uf_293) (<= 1 uf_294)) (implies true (implies (and (<= 0 uf_293) (<= 1 uf_294)) (and (implies up_216 (implies (and (<= 0 uf_293) (<= 1 uf_294)) (implies true (implies (and (<= 0 uf_293) (<= 1 uf_294)) (implies (and (<= 0 uf_293) (<= 1 uf_294)) (implies true (implies (= uf_296 uf_292) (implies (= uf_297 uf_294) (implies (= uf_298 uf_293) (implies (= uf_299 uf_292) (implies true (and (implies (forall (?x781 Int) (implies (and (<= ?x781 uf_76) (<= 0 ?x781)) (implies (< ?x781 uf_272) (<= (uf_110 uf_273 (uf_66 (uf_43 uf_7 uf_274) ?x781 uf_7)) uf_299)))) (and (implies (exists (?x782 Int) (and (= (uf_110 uf_273 (uf_66 (uf_43 uf_7 uf_274) ?x782 uf_7)) uf_299) (and (< ?x782 uf_272) (and (<= ?x782 uf_76) (<= 0 ?x782))))) true) (exists (?x783 Int) (and (= (uf_110 uf_273 (uf_66 (uf_43 uf_7 uf_274) ?x783 uf_7)) uf_299) (and (< ?x783 uf_272) (and (<= ?x783 uf_76) (<= 0 ?x783))))))) (forall (?x784 Int) (implies (and (<= ?x784 uf_76) (<= 0 ?x784)) (implies (< ?x784 uf_272) (<= (uf_110 uf_273 (uf_66 (uf_43 uf_7 uf_274) ?x784 uf_7)) uf_299)))))))))))))))) up_216)))))))))) (implies true (implies (and (<= 0 uf_293) (<= 1 uf_294)) (implies (< uf_294 uf_272) (implies (and (<= 0 uf_293) (<= 1 uf_294)) (implies true (implies (and (<= 0 uf_293) (<= 1 uf_294)) (and (implies (and (= (uf_27 uf_273 (uf_66 (uf_43 uf_7 uf_274) uf_294 uf_7)) uf_9) (= (uf_48 (uf_66 (uf_43 uf_7 uf_274) uf_294 uf_7) uf_7) uf_9)) (and (implies (and (= (uf_67 uf_273 (uf_66 (uf_43 uf_7 uf_274) uf_294 uf_7)) uf_9) (= (uf_48 (uf_66 (uf_43 uf_7 uf_274) uf_294 uf_7) uf_7) uf_9)) (implies (and (<= 0 uf_293) (<= 1 uf_294)) (and (implies true (implies (and (<= 0 uf_293) (<= 1 uf_294)) (implies (<= (uf_110 uf_273 (uf_66 (uf_43 uf_7 uf_274) uf_294 uf_7)) uf_292) (implies (and (<= 0 uf_293) (<= 1 uf_294)) (implies true (implies (and (<= 0 uf_293) (<= 1 uf_294)) (implies (and (<= 0 uf_293) (<= 1 uf_294)) (implies true (implies (= uf_300 uf_292) (implies (= uf_301 uf_293) (implies true (implies (and (<= 0 uf_301) (<= 1 uf_294)) (and (implies (and (<= (+ uf_294 1) uf_76) (<= 0 (+ uf_294 1))) (implies (= uf_302 (+ uf_294 1)) (implies (up_280 uf_273 uf_303 uf_291 uf_302 uf_4) (implies (and (<= 0 uf_301) (<= 2 uf_302)) (implies true (and (implies (<= uf_302 uf_272) (and (implies (forall (?x785 Int) (implies (and (<= ?x785 uf_76) (<= 0 ?x785)) (implies (< ?x785 uf_302) (<= (uf_110 uf_273 (uf_66 (uf_43 uf_7 uf_274) ?x785 uf_7)) uf_300)))) (and (implies (and (= (uf_110 uf_273 (uf_66 (uf_43 uf_7 uf_274) uf_301 uf_7)) uf_300) (< uf_301 uf_272)) (implies false true)) (and (= (uf_110 uf_273 (uf_66 (uf_43 uf_7 uf_274) uf_301 uf_7)) uf_300) (< uf_301 uf_272)))) (forall (?x786 Int) (implies (and (<= ?x786 uf_76) (<= 0 ?x786)) (implies (< ?x786 uf_302) (<= (uf_110 uf_273 (uf_66 (uf_43 uf_7 uf_274) ?x786 uf_7)) uf_300)))))) (<= uf_302 uf_272))))))) (and (<= (+ uf_294 1) uf_76) (<= 0 (+ uf_294 1)))))))))))))))) (implies true (implies (and (<= 0 uf_293) (<= 1 uf_294)) (implies (< uf_292 (uf_110 uf_273 (uf_66 (uf_43 uf_7 uf_274) uf_294 uf_7))) (implies (and (<= 0 uf_293) (<= 1 uf_294)) (implies true (implies (and (<= 0 uf_293) (<= 1 uf_294)) (and (implies (and (= (uf_27 uf_273 (uf_66 (uf_43 uf_7 uf_274) uf_294 uf_7)) uf_9) (= (uf_48 (uf_66 (uf_43 uf_7 uf_274) uf_294 uf_7) uf_7) uf_9)) (and (implies (and (= (uf_67 uf_273 (uf_66 (uf_43 uf_7 uf_274) uf_294 uf_7)) uf_9) (= (uf_48 (uf_66 (uf_43 uf_7 uf_274) uf_294 uf_7) uf_7) uf_9)) (implies (= uf_304 (uf_110 uf_273 (uf_66 (uf_43 uf_7 uf_274) uf_294 uf_7))) (implies (up_280 uf_273 uf_305 uf_287 uf_304 uf_7) (implies (up_280 uf_273 uf_306 uf_289 uf_294 uf_4) (implies (and (<= 1 uf_294) (<= 1 uf_294)) (implies true (implies (= uf_300 uf_304) (implies (= uf_301 uf_294) (implies true (implies (and (<= 0 uf_301) (<= 1 uf_294)) (and (implies (and (<= (+ uf_294 1) uf_76) (<= 0 (+ uf_294 1))) (implies (= uf_302 (+ uf_294 1)) (implies (up_280 uf_273 uf_303 uf_291 uf_302 uf_4) (implies (and (<= 0 uf_301) (<= 2 uf_302)) (implies true (and (implies (<= uf_302 uf_272) (and (implies (forall (?x787 Int) (implies (and (<= ?x787 uf_76) (<= 0 ?x787)) (implies (< ?x787 uf_302) (<= (uf_110 uf_273 (uf_66 (uf_43 uf_7 uf_274) ?x787 uf_7)) uf_300)))) (and (implies (and (= (uf_110 uf_273 (uf_66 (uf_43 uf_7 uf_274) uf_301 uf_7)) uf_300) (< uf_301 uf_272)) (implies false true)) (and (= (uf_110 uf_273 (uf_66 (uf_43 uf_7 uf_274) uf_301 uf_7)) uf_300) (< uf_301 uf_272)))) (forall (?x788 Int) (implies (and (<= ?x788 uf_76) (<= 0 ?x788)) (implies (< ?x788 uf_302) (<= (uf_110 uf_273 (uf_66 (uf_43 uf_7 uf_274) ?x788 uf_7)) uf_300)))))) (<= uf_302 uf_272))))))) (and (<= (+ uf_294 1) uf_76) (<= 0 (+ uf_294 1)))))))))))))) (and (= (uf_67 uf_273 (uf_66 (uf_43 uf_7 uf_274) uf_294 uf_7)) uf_9) (= (uf_48 (uf_66 (uf_43 uf_7 uf_274) uf_294 uf_7) uf_7) uf_9)))) (and (= (uf_27 uf_273 (uf_66 (uf_43 uf_7 uf_274) uf_294 uf_7)) uf_9) (= (uf_48 (uf_66 (uf_43 uf_7 uf_274) uf_294 uf_7) uf_7) uf_9)))))))))))) (and (= (uf_67 uf_273 (uf_66 (uf_43 uf_7 uf_274) uf_294 uf_7)) uf_9) (= (uf_48 (uf_66 (uf_43 uf_7 uf_274) uf_294 uf_7) uf_7) uf_9)))) (and (= (uf_27 uf_273 (uf_66 (uf_43 uf_7 uf_274) uf_294 uf_7)) uf_9) (= (uf_48 (uf_66 (uf_43 uf_7 uf_274) uf_294 uf_7) uf_7) uf_9)))))))))))))))))))))))))) (implies true (implies (and (<= 0 uf_293) (<= 1 uf_294)) (implies (not true) (implies (and (<= 0 uf_293) (<= 1 uf_294)) (implies true (implies (and (<= 0 uf_293) (<= 1 uf_294)) (implies (and (= (uf_55 uf_273) uf_9) (= (uf_202 uf_295 uf_273) uf_9)) (implies (and (<= 0 uf_293) (<= 1 uf_294)) (implies true (implies (and (<= 0 uf_293) (<= 1 uf_294)) (and (implies up_216 (implies (and (<= 0 uf_293) (<= 1 uf_294)) (implies true (implies (and (<= 0 uf_293) (<= 1 uf_294)) (implies (and (<= 0 uf_293) (<= 1 uf_294)) (implies true (implies (= uf_296 uf_292) (implies (= uf_297 uf_294) (implies (= uf_298 uf_293) (implies (= uf_299 uf_292) (implies true (and (implies (forall (?x789 Int) (implies (and (<= ?x789 uf_76) (<= 0 ?x789)) (implies (< ?x789 uf_272) (<= (uf_110 uf_273 (uf_66 (uf_43 uf_7 uf_274) ?x789 uf_7)) uf_299)))) (and (implies (exists (?x790 Int) (and (= (uf_110 uf_273 (uf_66 (uf_43 uf_7 uf_274) ?x790 uf_7)) uf_299) (and (< ?x790 uf_272) (and (<= ?x790 uf_76) (<= 0 ?x790))))) true) (exists (?x791 Int) (and (= (uf_110 uf_273 (uf_66 (uf_43 uf_7 uf_274) ?x791 uf_7)) uf_299) (and (< ?x791 uf_272) (and (<= ?x791 uf_76) (<= 0 ?x791))))))) (forall (?x792 Int) (implies (and (<= ?x792 uf_76) (<= 0 ?x792)) (implies (< ?x792 uf_272) (<= (uf_110 uf_273 (uf_66 (uf_43 uf_7 uf_274) ?x792 uf_7)) uf_299)))))))))))))))) up_216)))))))))))))))))))))) (and (= (uf_110 uf_273 (uf_66 (uf_43 uf_7 uf_274) 0 uf_7)) uf_285) (< 0 uf_272)))) (forall (?x793 Int) (implies (and (<= ?x793 uf_76) (<= 0 ?x793)) (implies (< ?x793 1) (<= (uf_110 uf_273 (uf_66 (uf_43 uf_7 uf_274) ?x793 uf_7)) uf_285)))))) (<= 1 uf_272)))))))) (and (= (uf_67 uf_273 (uf_66 (uf_43 uf_7 uf_274) 0 uf_7)) uf_9) (= (uf_48 (uf_66 (uf_43 uf_7 uf_274) 0 uf_7) uf_7) uf_9)))) (and (= (uf_27 uf_273 (uf_66 (uf_43 uf_7 uf_274) 0 uf_7)) uf_9) (= (uf_48 (uf_66 (uf_43 uf_7 uf_274) 0 uf_7) uf_7) uf_9)))) (= (uf_200 uf_273 (uf_43 (uf_124 uf_7 uf_272) (uf_116 (uf_43 uf_7 uf_274))) (uf_43 (uf_124 uf_7 uf_272) (uf_116 (uf_43 uf_7 uf_274))) uf_284) uf_9))))))))))))))))))) +:formula true +) diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/Boogie/Examples/cert/VCC_b_maximum.proof --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Boogie/Examples/cert/VCC_b_maximum.proof Wed Nov 04 11:40:59 2009 +0100 @@ -0,0 +1,8070 @@ +#2 := false +#121 := 0::int +decl uf_110 :: (-> T4 T5 int) +decl uf_66 :: (-> T5 int T3 T5) +decl uf_7 :: T3 +#10 := uf_7 +decl ?x785!14 :: int +#19054 := ?x785!14 +decl uf_43 :: (-> T3 int T5) +decl uf_274 :: int +#2959 := uf_274 +#2960 := (uf_43 uf_7 uf_274) +#19059 := (uf_66 #2960 ?x785!14 uf_7) +decl uf_273 :: T4 +#2958 := uf_273 +#19060 := (uf_110 uf_273 #19059) +#4076 := -1::int +#19385 := (* -1::int #19060) +decl uf_300 :: int +#3186 := uf_300 +#19386 := (+ uf_300 #19385) +#19387 := (>= #19386 0::int) +#23584 := (not #19387) +#19372 := (* -1::int ?x785!14) +decl uf_302 :: int +#3196 := uf_302 +#19373 := (+ uf_302 #19372) +#19374 := (<= #19373 0::int) +#19056 := (>= ?x785!14 0::int) +#22816 := (not #19056) +#7878 := 131073::int +#19055 := (<= ?x785!14 131073::int) +#22815 := (not #19055) +#22831 := (or #22815 #22816 #19374 #19387) +#22836 := (not #22831) +#161 := (:var 0 int) +#3039 := (uf_66 #2960 #161 uf_7) +#23745 := (pattern #3039) +#15606 := (<= #161 131073::int) +#20064 := (not #15606) +#14120 := (* -1::int uf_300) +#3040 := (uf_110 uf_273 #3039) +#14121 := (+ #3040 #14120) +#14122 := (<= #14121 0::int) +#14101 := (* -1::int uf_302) +#14110 := (+ #161 #14101) +#14109 := (>= #14110 0::int) +#4084 := (>= #161 0::int) +#5113 := (not #4084) +#22797 := (or #5113 #14109 #14122 #20064) +#23762 := (forall (vars (?x785 int)) (:pat #23745) #22797) +#23767 := (not #23762) +decl uf_301 :: int +#3188 := uf_301 +#14142 := (* -1::int uf_301) +decl uf_272 :: int +#2949 := uf_272 +#14143 := (+ uf_272 #14142) +#14144 := (<= #14143 0::int) +#3208 := (uf_66 #2960 uf_301 uf_7) +#3209 := (uf_110 uf_273 #3208) +#12862 := (= uf_300 #3209) +#22782 := (not #12862) +#22783 := (or #22782 #14144) +#22784 := (not #22783) +#23770 := (or #22784 #23767) +#14145 := (not #14144) +decl uf_294 :: int +#3055 := uf_294 +#14044 := (* -1::int uf_294) +#14045 := (+ uf_272 #14044) +#14046 := (<= #14045 0::int) +#14049 := (not #14046) +decl uf_125 :: (-> T5 T5 int) +decl uf_28 :: (-> int T5) +decl uf_29 :: (-> T5 int) +#2992 := (uf_29 #2960) +#23223 := (uf_28 #2992) +decl uf_15 :: (-> T5 T3) +#26404 := (uf_15 #23223) +decl uf_293 :: int +#3051 := uf_293 +#26963 := (uf_66 #23223 uf_293 #26404) +#26964 := (uf_125 #26963 #23223) +#27037 := (>= #26964 0::int) +#13947 := (>= uf_293 0::int) +decl ?x781!15 :: int +#19190 := ?x781!15 +#19195 := (uf_66 #2960 ?x781!15 uf_7) +#19196 := (uf_110 uf_273 #19195) +#19541 := (* -1::int #19196) +decl uf_299 :: int +#3138 := uf_299 +#19542 := (+ uf_299 #19541) +#19543 := (>= #19542 0::int) +#19528 := (* -1::int ?x781!15) +#19529 := (+ uf_272 #19528) +#19530 := (<= #19529 0::int) +#19192 := (>= ?x781!15 0::int) +#22993 := (not #19192) +#19191 := (<= ?x781!15 131073::int) +#22992 := (not #19191) +#23008 := (or #22992 #22993 #19530 #19543) +#23013 := (not #23008) +#13873 := (* -1::int uf_272) +#13960 := (+ #161 #13873) +#13959 := (>= #13960 0::int) +#3145 := (= #3040 uf_299) +#22966 := (not #3145) +#22967 := (or #22966 #5113 #13959 #20064) +#23886 := (forall (vars (?x782 int)) (:pat #23745) #22967) +#23891 := (not #23886) +#13970 := (* -1::int uf_299) +#13971 := (+ #3040 #13970) +#13972 := (<= #13971 0::int) +#22958 := (or #5113 #13959 #13972 #20064) +#23878 := (forall (vars (?x781 int)) (:pat #23745) #22958) +#23883 := (not #23878) +#23894 := (or #23883 #23891) +#23897 := (not #23894) +#23900 := (or #23897 #23013) +#23903 := (not #23900) +#4 := 1::int +#13950 := (>= uf_294 1::int) +#14243 := (not #13950) +#22873 := (not #13947) +decl uf_292 :: int +#3047 := uf_292 +#12576 := (= uf_292 uf_299) +#12644 := (not #12576) +decl uf_298 :: int +#3136 := uf_298 +#12573 := (= uf_293 uf_298) +#12653 := (not #12573) +decl uf_297 :: int +#3134 := uf_297 +#12570 := (= uf_294 uf_297) +#12662 := (not #12570) +decl uf_296 :: int +#3132 := uf_296 +#12567 := (= uf_292 uf_296) +#12671 := (not #12567) +#23906 := (or #12671 #12662 #12653 #12644 #22873 #14243 #14049 #23903) +#23909 := (not #23906) +#23773 := (not #23770) +#23776 := (or #23773 #22836) +#23779 := (not #23776) +#14102 := (+ uf_272 #14101) +#14100 := (>= #14102 0::int) +#14105 := (not #14100) +#23782 := (or #14105 #23779) +#23785 := (not #23782) +#23788 := (or #14105 #23785) +#23791 := (not #23788) +#1066 := 131072::int +#16368 := (<= uf_294 131072::int) +#19037 := (not #16368) +#14169 := (+ uf_294 #14101) +#14168 := (= #14169 -1::int) +#14172 := (not #14168) +#1120 := 2::int +#14092 := (>= uf_302 2::int) +#22859 := (not #14092) +#14088 := (>= uf_294 -1::int) +#19034 := (not #14088) +#14076 := (>= uf_301 0::int) +#22858 := (not #14076) +decl up_280 :: (-> T4 T1 T1 int T3 bool) +decl uf_4 :: T3 +#7 := uf_4 +decl uf_291 :: T1 +#3030 := uf_291 +decl uf_303 :: T1 +#3198 := uf_303 +#3199 := (up_280 uf_273 uf_303 uf_291 uf_302 uf_4) +#12942 := (not #3199) +#23794 := (or #12942 #22858 #19034 #22859 #14172 #19037 #23791) +#23797 := (not #23794) +#23800 := (or #19034 #19037 #23797) +#23803 := (not #23800) +#13075 := (= uf_294 uf_301) +#13081 := (not #13075) +decl uf_304 :: int +#3239 := uf_304 +#3175 := (uf_66 #2960 uf_294 uf_7) +#3184 := (uf_110 uf_273 #3175) +#13070 := (= #3184 uf_304) +#13133 := (not #13070) +decl uf_67 :: (-> T4 T5 T2) +#3181 := (uf_67 uf_273 #3175) +decl uf_9 :: T2 +#19 := uf_9 +#12812 := (= uf_9 #3181) +#19017 := (not #12812) +decl uf_48 :: (-> T5 T3 T2) +#3178 := (uf_48 #3175 uf_7) +#12806 := (= uf_9 #3178) +#19011 := (not #12806) +#3246 := (= uf_300 uf_304) +#13090 := (not #3246) +decl uf_289 :: T1 +#3027 := uf_289 +decl uf_306 :: T1 +#3243 := uf_306 +#3244 := (up_280 uf_273 uf_306 uf_289 uf_294 uf_4) +#13115 := (not #3244) +decl uf_287 :: T1 +#3024 := uf_287 +decl uf_305 :: T1 +#3241 := uf_305 +#3242 := (up_280 uf_273 uf_305 uf_287 uf_304 uf_7) +#13124 := (not #3242) +#23812 := (or #13124 #13115 #13090 #19011 #19017 #13133 #13081 #14243 #22858 #23803) +#23815 := (not #23812) +#23818 := (or #19011 #19017 #23815) +#23821 := (not #23818) +decl uf_27 :: (-> T4 T5 T2) +#3176 := (uf_27 uf_273 #3175) +#12803 := (= uf_9 #3176) +#19008 := (not #12803) +#23824 := (or #19008 #19011 #23821) +#23827 := (not #23824) +#23830 := (or #19008 #19011 #23827) +#23833 := (not #23830) +#14208 := (* -1::int #3184) +#14209 := (+ uf_292 #14208) +#14207 := (>= #14209 0::int) +#23836 := (or #22873 #14243 #14207 #23833) +#23839 := (not #23836) +#14211 := (not #14207) +#12826 := (= uf_293 uf_301) +#12993 := (not #12826) +#12823 := (= uf_292 uf_300) +#13002 := (not #12823) +#23806 := (or #13002 #12993 #22873 #14243 #22858 #14211 #23803) +#23809 := (not #23806) +#23842 := (or #23809 #23839) +#23845 := (not #23842) +#23848 := (or #19011 #19017 #22873 #14243 #23845) +#23851 := (not #23848) +#23854 := (or #19011 #19017 #23851) +#23857 := (not #23854) +#23860 := (or #19008 #19011 #23857) +#23863 := (not #23860) +#23866 := (or #19008 #19011 #23863) +#23869 := (not #23866) +#23872 := (or #22873 #14243 #14046 #23869) +#23875 := (not #23872) +#23912 := (or #23875 #23909) +#23915 := (not #23912) +#14431 := (* -1::int uf_292) +#14432 := (+ #3040 #14431) +#14433 := (<= #14432 0::int) +#14421 := (+ #161 #14044) +#14420 := (>= #14421 0::int) +#22774 := (or #5113 #14420 #14433 #20064) +#23754 := (forall (vars (?x774 int)) (:pat #23745) #22774) +#23759 := (not #23754) +#1322 := 255::int +#16349 := (<= uf_292 255::int) +#23043 := (not #16349) +#16332 := (<= uf_293 131073::int) +#23042 := (not #16332) +#16310 := (<= uf_294 131073::int) +#23041 := (not #16310) +#14490 := (>= uf_292 0::int) +#23039 := (not #14490) +#14462 := (>= uf_294 0::int) +#23038 := (not #14462) +#14453 := (>= #14045 0::int) +#14456 := (not #14453) +#14402 := (* -1::int uf_293) +#14403 := (+ uf_272 #14402) +#14404 := (<= #14403 0::int) +#13942 := (<= uf_272 0::int) +decl uf_202 :: (-> T1 T4 T2) +decl uf_295 :: T1 +#3117 := uf_295 +#3118 := (uf_202 uf_295 uf_273) +#12553 := (= uf_9 #3118) +#15709 := (not #12553) +decl uf_177 :: (-> T4 T4 T2) +#3072 := (uf_177 uf_273 uf_273) +#12437 := (= uf_9 #3072) +#14399 := (not #12437) +#3067 := (uf_66 #2960 uf_293 uf_7) +#3068 := (uf_110 uf_273 #3067) +#12426 := (= uf_292 #3068) +#23037 := (not #12426) +decl uf_6 :: (-> T3 T3) +#11 := (uf_6 uf_7) +decl uf_279 :: T1 +#2990 := uf_279 +#3126 := (up_280 uf_273 uf_295 uf_279 #2992 #11) +#23036 := (not #3126) +decl up_278 :: (-> T4 T1 T1 T5 T3 bool) +#3125 := (up_278 uf_273 uf_295 uf_279 #2960 #11) +#23035 := (not #3125) +decl uf_281 :: T1 +#2995 := uf_281 +#3124 := (up_280 uf_273 uf_295 uf_281 uf_272 uf_4) +#13340 := (not #3124) +#3123 := (up_280 uf_273 uf_295 uf_287 uf_292 uf_7) +#13349 := (not #3123) +#3122 := (up_280 uf_273 uf_295 uf_289 uf_293 uf_4) +#13358 := (not #3122) +#3121 := (up_280 uf_273 uf_295 uf_291 uf_294 uf_4) +#13367 := (not #3121) +#3011 := (uf_66 #2960 0::int uf_7) +#3021 := (uf_110 uf_273 #3011) +decl uf_285 :: int +#3020 := uf_285 +#3022 := (= uf_285 #3021) +#13672 := (not #3022) +#23918 := (or #13672 #13367 #13358 #13349 #13340 #23035 #23036 #23037 #14399 #15709 #13942 #22873 #14243 #14404 #14456 #23038 #23039 #23041 #23042 #23043 #23759 #23915) +#23921 := (not #23918) +#23924 := (or #13672 #13942 #23921) +#23927 := (not #23924) +#13922 := (* -1::int #3040) +#13923 := (+ uf_285 #13922) +#13921 := (>= #13923 0::int) +#13910 := (>= #161 1::int) +#22763 := (or #5113 #13910 #13921 #20064) +#23746 := (forall (vars (?x773 int)) (:pat #23745) #22763) +#23751 := (not #23746) +#23930 := (or #23751 #23927) +#23933 := (not #23930) +decl ?x773!13 :: int +#18929 := ?x773!13 +#18939 := (>= ?x773!13 1::int) +#18934 := (uf_66 #2960 ?x773!13 uf_7) +#18935 := (uf_110 uf_273 #18934) +#18936 := (* -1::int #18935) +#18937 := (+ uf_285 #18936) +#18938 := (>= #18937 0::int) +#18931 := (>= ?x773!13 0::int) +#22737 := (not #18931) +#18930 := (<= ?x773!13 131073::int) +#22736 := (not #18930) +#22752 := (or #22736 #22737 #18938 #18939) +#22757 := (not #22752) +#23936 := (or #22757 #23933) +#23939 := (not #23936) +#13903 := (>= uf_272 1::int) +#13906 := (not #13903) +#23942 := (or #13906 #23939) +#23945 := (not #23942) +#23948 := (or #13906 #23945) +#23951 := (not #23948) +#3017 := (uf_67 uf_273 #3011) +#12367 := (= uf_9 #3017) +#18906 := (not #12367) +#3014 := (uf_48 #3011 uf_7) +#12361 := (= uf_9 #3014) +#18900 := (not #12361) +decl uf_290 :: T1 +#3029 := uf_290 +#3031 := (up_280 uf_273 uf_290 uf_291 1::int uf_4) +#13645 := (not #3031) +decl uf_288 :: T1 +#3026 := uf_288 +#3028 := (up_280 uf_273 uf_288 uf_289 0::int uf_4) +#13654 := (not #3028) +decl uf_286 :: T1 +#3023 := uf_286 +#3025 := (up_280 uf_273 uf_286 uf_287 uf_285 uf_7) +#13663 := (not #3025) +#23954 := (or #13672 #13663 #13654 #13645 #18900 #18906 #23951) +#23957 := (not #23954) +#23960 := (or #18900 #18906 #23957) +#23963 := (not #23960) +#3012 := (uf_27 uf_273 #3011) +#12358 := (= uf_9 #3012) +#18897 := (not #12358) +#23966 := (or #18897 #18900 #23963) +#23969 := (not #23966) +#23972 := (or #18897 #18900 #23969) +#23975 := (not #23972) +decl uf_200 :: (-> T4 T5 T5 T16 T2) +decl uf_284 :: T16 +#3008 := uf_284 +decl uf_116 :: (-> T5 int) +#2961 := (uf_116 #2960) +decl uf_124 :: (-> T3 int T3) +#2952 := (uf_124 uf_7 uf_272) +#2962 := (uf_43 #2952 #2961) +#3009 := (uf_200 uf_273 #2962 #2962 uf_284) +#12355 := (= uf_9 #3009) +#13715 := (not #12355) +#23978 := (or #13715 #23975) +#23981 := (not #23978) +decl uf_14 :: (-> T3 T8) +#24016 := (uf_116 #2962) +#25404 := (uf_43 #2952 #24016) +#25815 := (uf_15 #25404) +#26092 := (uf_14 #25815) +decl uf_16 :: T8 +#35 := uf_16 +#26095 := (= uf_16 #26092) +#26297 := (not #26095) +#2955 := (uf_14 #2952) +#12296 := (= uf_16 #2955) +#12299 := (not #12296) +#26298 := (iff #12299 #26297) +#26293 := (iff #12296 #26095) +#26342 := (iff #26095 #12296) +#26340 := (= #26092 #2955) +#26338 := (= #25815 #2952) +#24234 := (uf_15 #2962) +#28358 := (= #24234 #2952) +#24237 := (= #2952 #24234) +#326 := (:var 1 T3) +#2692 := (uf_43 #326 #161) +#23682 := (pattern #2692) +#2696 := (uf_15 #2692) +#11677 := (= #326 #2696) +#23689 := (forall (vars (?x720 T3) (?x721 int)) (:pat #23682) #11677) +#11681 := (forall (vars (?x720 T3) (?x721 int)) #11677) +#23692 := (iff #11681 #23689) +#23690 := (iff #11677 #11677) +#23691 := [refl]: #23690 +#23693 := [quant-intro #23691]: #23692 +#18759 := (~ #11681 #11681) +#18757 := (~ #11677 #11677) +#18758 := [refl]: #18757 +#18760 := [nnf-pos #18758]: #18759 +#2697 := (= #2696 #326) +#2698 := (forall (vars (?x720 T3) (?x721 int)) #2697) +#11682 := (iff #2698 #11681) +#11679 := (iff #2697 #11677) +#11680 := [rewrite]: #11679 +#11683 := [quant-intro #11680]: #11682 +#11676 := [asserted]: #2698 +#11686 := [mp #11676 #11683]: #11681 +#18761 := [mp~ #11686 #18760]: #11681 +#23694 := [mp #18761 #23693]: #23689 +#24181 := (not #23689) +#24242 := (or #24181 #24237) +#24243 := [quant-inst]: #24242 +#28006 := [unit-resolution #24243 #23694]: #24237 +#28359 := [symm #28006]: #28358 +#26336 := (= #25815 #24234) +#27940 := (= #25404 #2962) +#25411 := (= #2962 #25404) +#2965 := (uf_48 #2962 #2952) +#12305 := (= uf_9 #2965) +decl uf_24 :: (-> T4 T5 T2) +#2969 := (uf_24 uf_273 #2962) +#12311 := (= uf_9 #2969) +decl uf_25 :: (-> T4 T5 T5) +#2967 := (uf_25 uf_273 #2962) +decl uf_26 :: T5 +#78 := uf_26 +#12308 := (= uf_26 #2967) +#2963 := (uf_27 uf_273 #2962) +#12302 := (= uf_9 #2963) +decl uf_22 :: (-> T3 T2) +#2953 := (uf_22 #2952) +#12293 := (= uf_9 #2953) +#14658 := (and #12293 #12299 #12302 #12305 #12308 #12311) +decl uf_269 :: int +#2937 := uf_269 +#14715 := (>= uf_269 0::int) +#14711 := (* -1::int uf_269) +decl uf_78 :: int +#429 := uf_78 +#14712 := (+ uf_78 #14711) +#14710 := (>= #14712 0::int) +#14718 := (and #14710 #14715) +#14721 := (not #14718) +decl uf_270 :: int +#2941 := uf_270 +#14701 := (>= uf_270 0::int) +#14697 := (* -1::int uf_270) +decl uf_76 :: int +#409 := uf_76 +#14698 := (+ uf_76 #14697) +#14696 := (>= #14698 0::int) +#14704 := (and #14696 #14701) +#14707 := (not #14704) +decl uf_271 :: int +#2945 := uf_271 +#14687 := (>= uf_271 0::int) +#14683 := (* -1::int uf_271) +#14684 := (+ uf_76 #14683) +#14682 := (>= #14684 0::int) +#14690 := (and #14682 #14687) +#14693 := (not #14690) +#974 := 1099511627776::int +#14671 := (>= uf_272 1099511627776::int) +#14661 := (not #14658) +decl uf_276 :: (-> T19 int) +#2984 := (:var 0 T19) +#2985 := (uf_276 #2984) +#2986 := (pattern #2985) +decl uf_277 :: int +#2987 := uf_277 +#14648 := (* -1::int uf_277) +#14649 := (+ #2985 #14648) +#14647 := (>= #14649 0::int) +#14646 := (not #14647) +#14652 := (forall (vars (?x771 T19)) (:pat #2986) #14646) +#14655 := (not #14652) +#13943 := (not #13942) +#14502 := (and #3022 #13943) +#14507 := (not #14502) +#14487 := (+ uf_78 #14431) +#14486 := (>= #14487 0::int) +#14493 := (and #14486 #14490) +#14496 := (not #14493) +#14472 := (+ uf_76 #14402) +#14471 := (>= #14472 0::int) +#14478 := (and #13947 #14471) +#14483 := (not #14478) +#14085 := (+ uf_76 #14044) +#14459 := (>= #14085 0::int) +#14465 := (and #14459 #14462) +#14468 := (not #14465) +#4413 := (* -1::int uf_76) +#4418 := (+ #161 #4413) +#4419 := (<= #4418 0::int) +#5736 := (and #4084 #4419) +#5739 := (not #5736) +#14442 := (or #5739 #14420 #14433) +#14447 := (forall (vars (?x774 int)) #14442) +#14450 := (not #14447) +#14405 := (not #14404) +#14411 := (and #12426 #14405) +#14416 := (not #14411) +#14084 := (>= #14085 1::int) +#14175 := (and #14084 #14088) +#14178 := (not #14175) +#14151 := (and #12862 #14145) +#14131 := (or #5739 #14109 #14122) +#14136 := (forall (vars (?x785 int)) #14131) +#14139 := (not #14136) +#14156 := (or #14139 #14151) +#14159 := (and #14136 #14156) +#14162 := (or #14105 #14159) +#14165 := (and #14100 #14162) +#14094 := (and #14076 #14092) +#14097 := (not #14094) +#14193 := (or #12942 #14097 #14165 #14172 #14178) +#14201 := (and #14084 #14088 #14193) +#14078 := (and #13950 #14076) +#14081 := (not #14078) +#12818 := (and #12806 #12812) +#13142 := (not #12818) +#14267 := (or #13124 #13115 #13090 #13142 #13133 #13081 #14243 #14081 #14201) +#14275 := (and #12806 #12812 #14267) +#12809 := (and #12803 #12806) +#13159 := (not #12809) +#14280 := (or #13159 #14275) +#14286 := (and #12803 #12806 #14280) +#13952 := (and #13947 #13950) +#13955 := (not #13952) +#14312 := (or #13955 #14207 #14286) +#14238 := (or #13002 #12993 #13955 #14081 #14201 #14211) +#14317 := (and #14238 #14312) +#14326 := (or #13142 #13955 #14317) +#14334 := (and #12806 #12812 #14326) +#14339 := (or #13159 #14334) +#14345 := (and #12803 #12806 #14339) +#14371 := (or #13955 #14046 #14345) +#13958 := (not #13959) +#13998 := (and #3145 #4084 #4419 #13958) +#14003 := (exists (vars (?x782 int)) #13998) +#13981 := (or #5739 #13959 #13972) +#13986 := (forall (vars (?x781 int)) #13981) +#13989 := (not #13986) +#14006 := (or #13989 #14003) +#14009 := (and #13986 #14006) +decl up_216 :: bool +#2477 := up_216 +#12719 := (not up_216) +#14036 := (or #12719 #12671 #12662 #12653 #12644 #13955 #14009) +#14041 := (and up_216 #14036) +#14070 := (or #13955 #14041 #14049) +#14376 := (and #14070 #14371) +decl uf_55 :: (-> T4 T2) +#2978 := (uf_55 uf_273) +#12332 := (= uf_9 #2978) +#12556 := (and #12332 #12553) +#13376 := (not #12556) +#3127 := (and #3125 #3126) +#13331 := (not #3127) +#14573 := (or #13367 #13358 #13349 #13340 #13331 #14399 #13376 #13955 #14376 #14416 #14450 #14456 #14468 #14483 #14496 #14507) +#14581 := (and #3022 #13943 #14573) +#13931 := (or #5739 #13910 #13921) +#13936 := (forall (vars (?x773 int)) #13931) +#13939 := (not #13936) +#14586 := (or #13939 #14581) +#14589 := (and #13936 #14586) +#14592 := (or #13906 #14589) +#14595 := (and #13903 #14592) +#12373 := (and #12361 #12367) +#13681 := (not #12373) +#14616 := (or #13672 #13663 #13654 #13645 #13681 #14595) +#14624 := (and #12361 #12367 #14616) +#12364 := (and #12358 #12361) +#13698 := (not #12364) +#14629 := (or #13698 #14624) +#14635 := (and #12358 #12361 #14629) +#14640 := (or #13715 #14635) +#14643 := (and #12355 #14640) +#13877 := (>= uf_272 0::int) +#13874 := (+ uf_76 #13873) +#13872 := (>= #13874 0::int) +#13880 := (and #13872 #13877) +#13883 := (not #13880) +decl uf_283 :: (-> int T5 T2) +#26 := (:var 0 T5) +decl uf_282 :: int +#2997 := uf_282 +#3000 := (uf_283 uf_282 #26) +#3001 := (pattern #3000) +#12341 := (= uf_9 #3000) +#12347 := (not #12341) +#12352 := (forall (vars (?x772 T5)) (:pat #3001) #12347) +#13741 := (not #12352) +decl uf_275 :: T1 +#2980 := uf_275 +#2981 := (uf_202 uf_275 uf_273) +#12335 := (= uf_9 #2981) +#12338 := (and #12332 #12335) +#13786 := (not #12338) +decl uf_203 :: (-> T4 T2) +#2976 := (uf_203 uf_273) +#12329 := (= uf_9 #2976) +#13795 := (not #12329) +decl uf_171 :: (-> T4 int) +#2998 := (uf_171 uf_273) +#2999 := (= uf_282 #2998) +#13750 := (not #2999) +#2996 := (up_280 uf_273 uf_275 uf_281 uf_272 uf_4) +#13759 := (not #2996) +#2993 := (up_280 uf_273 uf_275 uf_279 #2992 #11) +#2991 := (up_278 uf_273 uf_275 uf_279 #2960 #11) +#2994 := (and #2991 #2993) +#13768 := (not #2994) +#14766 := (or #13768 #13759 #13750 #13795 #13786 #13741 #13883 #13942 #14643 #14655 #14661 #14671 #14693 #14707 #14721) +#14771 := (not #14766) +#3010 := (= #3009 uf_9) +#3015 := (= #3014 uf_9) +#3013 := (= #3012 uf_9) +#3016 := (and #3013 #3015) +#3018 := (= #3017 uf_9) +#3019 := (and #3018 #3015) +#3037 := (<= 1::int uf_272) +#3041 := (<= #3040 uf_285) +#3038 := (< #161 1::int) +#3042 := (implies #3038 #3041) +#285 := (<= 0::int #161) +#410 := (<= #161 uf_76) +#645 := (and #410 #285) +#3043 := (implies #645 #3042) +#3044 := (forall (vars (?x773 int)) #3043) +#2951 := (< 0::int uf_272) +#3045 := (= #3021 uf_285) +#3046 := (and #3045 #2951) +#3141 := (<= #3040 uf_299) +#3140 := (< #161 uf_272) +#3142 := (implies #3140 #3141) +#3143 := (implies #645 #3142) +#3144 := (forall (vars (?x781 int)) #3143) +#3146 := (and #3140 #645) +#3147 := (and #3145 #3146) +#3148 := (exists (vars (?x782 int)) #3147) +#1 := true +#3149 := (implies #3148 true) +#3150 := (and #3149 #3148) +#3151 := (implies #3144 #3150) +#3152 := (and #3151 #3144) +#3153 := (implies true #3152) +#3139 := (= uf_299 uf_292) +#3154 := (implies #3139 #3153) +#3137 := (= uf_298 uf_293) +#3155 := (implies #3137 #3154) +#3135 := (= uf_297 uf_294) +#3156 := (implies #3135 #3155) +#3133 := (= uf_296 uf_292) +#3157 := (implies #3133 #3156) +#3158 := (implies true #3157) +#3059 := (<= 1::int uf_294) +#3053 := (<= 0::int uf_293) +#3060 := (and #3053 #3059) +#3159 := (implies #3060 #3158) +#3160 := (implies #3060 #3159) +#3161 := (implies true #3160) +#3162 := (implies #3060 #3161) +#3163 := (implies up_216 #3162) +#3164 := (and #3163 up_216) +#3165 := (implies #3060 #3164) +#3166 := (implies true #3165) +#3167 := (implies #3060 #3166) +#3119 := (= #3118 uf_9) +#2979 := (= #2978 uf_9) +#3120 := (and #2979 #3119) +#3295 := (implies #3120 #3167) +#3296 := (implies #3060 #3295) +#3297 := (implies true #3296) +#3298 := (implies #3060 #3297) +#3294 := (not true) +#3299 := (implies #3294 #3298) +#3300 := (implies #3060 #3299) +#3301 := (implies true #3300) +#3179 := (= #3178 uf_9) +#3177 := (= #3176 uf_9) +#3180 := (and #3177 #3179) +#3182 := (= #3181 uf_9) +#3183 := (and #3182 #3179) +#3192 := (+ uf_294 1::int) +#3194 := (<= 0::int #3192) +#3193 := (<= #3192 uf_76) +#3195 := (and #3193 #3194) +#3202 := (<= uf_302 uf_272) +#3204 := (<= #3040 uf_300) +#3203 := (< #161 uf_302) +#3205 := (implies #3203 #3204) +#3206 := (implies #645 #3205) +#3207 := (forall (vars (?x785 int)) #3206) +#3211 := (< uf_301 uf_272) +#3210 := (= #3209 uf_300) +#3212 := (and #3210 #3211) +#3213 := (implies false true) +#3214 := (implies #3212 #3213) +#3215 := (and #3214 #3212) +#3216 := (implies #3207 #3215) +#3217 := (and #3216 #3207) +#3218 := (implies #3202 #3217) +#3219 := (and #3218 #3202) +#3220 := (implies true #3219) +#3200 := (<= 2::int uf_302) +#3190 := (<= 0::int uf_301) +#3201 := (and #3190 #3200) +#3221 := (implies #3201 #3220) +#3222 := (implies #3199 #3221) +#3197 := (= uf_302 #3192) +#3223 := (implies #3197 #3222) +#3224 := (implies #3195 #3223) +#3225 := (and #3224 #3195) +#3191 := (and #3190 #3059) +#3226 := (implies #3191 #3225) +#3227 := (implies true #3226) +#3247 := (= uf_301 uf_294) +#3248 := (implies #3247 #3227) +#3249 := (implies #3246 #3248) +#3250 := (implies true #3249) +#3245 := (and #3059 #3059) +#3251 := (implies #3245 #3250) +#3252 := (implies #3244 #3251) +#3253 := (implies #3242 #3252) +#3240 := (= uf_304 #3184) +#3254 := (implies #3240 #3253) +#3255 := (implies #3183 #3254) +#3256 := (and #3255 #3183) +#3257 := (implies #3180 #3256) +#3258 := (and #3257 #3180) +#3259 := (implies #3060 #3258) +#3260 := (implies true #3259) +#3261 := (implies #3060 #3260) +#3238 := (< uf_292 #3184) +#3262 := (implies #3238 #3261) +#3263 := (implies #3060 #3262) +#3264 := (implies true #3263) +#3189 := (= uf_301 uf_293) +#3228 := (implies #3189 #3227) +#3187 := (= uf_300 uf_292) +#3229 := (implies #3187 #3228) +#3230 := (implies true #3229) +#3231 := (implies #3060 #3230) +#3232 := (implies #3060 #3231) +#3233 := (implies true #3232) +#3234 := (implies #3060 #3233) +#3185 := (<= #3184 uf_292) +#3235 := (implies #3185 #3234) +#3236 := (implies #3060 #3235) +#3237 := (implies true #3236) +#3265 := (and #3237 #3264) +#3266 := (implies #3060 #3265) +#3267 := (implies #3183 #3266) +#3268 := (and #3267 #3183) +#3269 := (implies #3180 #3268) +#3270 := (and #3269 #3180) +#3271 := (implies #3060 #3270) +#3272 := (implies true #3271) +#3273 := (implies #3060 #3272) +#3174 := (< uf_294 uf_272) +#3274 := (implies #3174 #3273) +#3275 := (implies #3060 #3274) +#3276 := (implies true #3275) +#3168 := (implies #3060 #3167) +#3169 := (implies true #3168) +#3170 := (implies #3060 #3169) +#3131 := (<= uf_272 uf_294) +#3171 := (implies #3131 #3170) +#3172 := (implies #3060 #3171) +#3173 := (implies true #3172) +#3277 := (and #3173 #3276) +#3278 := (implies #3060 #3277) +decl uf_59 :: (-> T4 T13) +#3079 := (uf_59 uf_273) +#3129 := (= #3079 #3079) +decl uf_41 :: (-> T4 T12) +#3088 := (uf_41 uf_273) +#3128 := (= #3088 #3088) +#3130 := (and #3128 #3129) +#3279 := (implies #3130 #3278) +#3280 := (implies #3127 #3279) +#3281 := (implies #3124 #3280) +#3282 := (implies #3123 #3281) +#3283 := (implies #3122 #3282) +#3284 := (implies #3121 #3283) +#3285 := (implies #3120 #3284) +#3078 := (<= #2998 #2998) +decl uf_170 :: (-> T4 T5 int) +#3074 := (uf_170 uf_273 #26) +#3075 := (pattern #3074) +#3076 := (<= #3074 #3074) +#3077 := (forall (vars (?x775 T5)) (:pat #3075) #3076) +#3115 := (and #3077 #3078) +#3073 := (= #3072 uf_9) +#3116 := (and #3073 #3115) +#3286 := (implies #3116 #3285) +decl uf_40 :: (-> T12 T5 T11) +#3089 := (uf_40 #3088 #26) +#3090 := (pattern #3089) +decl uf_261 :: T8 +#2832 := uf_261 +#3102 := (uf_25 uf_273 #26) +#3103 := (uf_15 #3102) +#3104 := (uf_14 #3103) +#3105 := (= #3104 uf_261) +#3106 := (not #3105) +#3107 := (implies #3106 #3106) +#3108 := (forall (vars (?x779 T5)) (:pat #3090) #3107) +decl uf_19 :: (-> T9 T5 int) +decl uf_20 :: (-> T4 T9) +#3095 := (uf_20 uf_273) +#3096 := (uf_19 #3095 #26) +#3097 := (pattern #3096) +#3098 := (= #3096 #3096) +#3082 := (uf_67 uf_273 #26) +#3083 := (= #3082 uf_9) +#3099 := (and #3083 #3098) +#3100 := (implies #3083 #3099) +#3101 := (forall (vars (?x778 T5)) (:pat #3097) #3100) +#3109 := (and #3101 #3108) +#3091 := (= #3089 #3089) +#3092 := (and #3083 #3091) +#3093 := (implies #3083 #3092) +#3094 := (forall (vars (?x777 T5)) (:pat #3090) #3093) +#3110 := (and #3094 #3109) +decl uf_58 :: (-> T13 T5 T14) +#3080 := (uf_58 #3079 #26) +#3081 := (pattern #3080) +#3084 := (= #3080 #3080) +#3085 := (and #3083 #3084) +#3086 := (implies #3083 #3085) +#3087 := (forall (vars (?x776 T5)) (:pat #3081) #3086) +#3111 := (and #3087 #3110) +#3112 := (and #3078 #3111) +#3113 := (and #3077 #3112) +#3114 := (and #3073 #3113) +#3287 := (implies #3114 #3286) +#3288 := (implies #3060 #3287) +#3289 := (implies true #3288) +#3290 := (implies #3060 #3289) +#3291 := (implies true #3290) +#3292 := (implies #3060 #3291) +#3293 := (implies true #3292) +#3302 := (and #3293 #3301) +#3303 := (implies #3060 #3302) +#3070 := (< uf_293 uf_272) +#3069 := (= #3068 uf_292) +#3071 := (and #3069 #3070) +#3304 := (implies #3071 #3303) +#3063 := (<= #3040 uf_292) +#3062 := (< #161 uf_294) +#3064 := (implies #3062 #3063) +#3065 := (implies #645 #3064) +#3066 := (forall (vars (?x774 int)) #3065) +#3305 := (implies #3066 #3304) +#3061 := (<= uf_294 uf_272) +#3306 := (implies #3061 #3305) +#3307 := (implies #3060 #3306) +#3057 := (<= 0::int uf_294) +#3056 := (<= uf_294 uf_76) +#3058 := (and #3056 #3057) +#3308 := (implies #3058 #3307) +#3052 := (<= uf_293 uf_76) +#3054 := (and #3052 #3053) +#3309 := (implies #3054 #3308) +#3049 := (<= 0::int uf_292) +#3048 := (<= uf_292 uf_78) +#3050 := (and #3048 #3049) +#3310 := (implies #3050 #3309) +#3311 := (implies true #3310) +#3312 := (implies #3046 #3311) +#3313 := (and #3312 #3046) +#3314 := (implies #3044 #3313) +#3315 := (and #3314 #3044) +#3316 := (implies #3037 #3315) +#3317 := (and #3316 #3037) +#3033 := (<= 1::int 1::int) +#3034 := (and #3033 #3033) +#3032 := (<= 0::int 0::int) +#3035 := (and #3032 #3034) +#3036 := (and #3032 #3035) +#3318 := (implies #3036 #3317) +#3319 := (implies #3031 #3318) +#3320 := (implies #3028 #3319) +#3321 := (implies #3025 #3320) +#3322 := (implies #3022 #3321) +#3323 := (implies #3019 #3322) +#3324 := (and #3323 #3019) +#3325 := (implies #3016 #3324) +#3326 := (and #3325 #3016) +#3327 := (implies #3010 #3326) +#3328 := (and #3327 #3010) +#3006 := (<= 0::int uf_272) +#3005 := (<= uf_272 uf_76) +#3007 := (and #3005 #3006) +#3329 := (implies #3007 #3328) +#3002 := (= #3000 uf_9) +#3003 := (iff #3002 false) +#3004 := (forall (vars (?x772 T5)) (:pat #3001) #3003) +#3330 := (implies #3004 #3329) +#3331 := (implies #2999 #3330) +#3332 := (implies #2996 #3331) +#3333 := (implies #2994 #3332) +#2988 := (< #2985 uf_277) +#2989 := (forall (vars (?x771 T19)) (:pat #2986) #2988) +#3334 := (implies #2989 #3333) +#2982 := (= #2981 uf_9) +#2983 := (and #2979 #2982) +#3335 := (implies #2983 #3334) +#2977 := (= #2976 uf_9) +#3336 := (implies #2977 #3335) +#3337 := (implies true #3336) +#2970 := (= #2969 uf_9) +#2968 := (= #2967 uf_26) +#2971 := (and #2968 #2970) +#2966 := (= #2965 uf_9) +#2972 := (and #2966 #2971) +#2964 := (= #2963 uf_9) +#2973 := (and #2964 #2972) +#2956 := (= #2955 uf_16) +#2957 := (not #2956) +#2974 := (and #2957 #2973) +#2954 := (= #2953 uf_9) +#2975 := (and #2954 #2974) +#3338 := (implies #2975 #3337) +#3339 := (implies #2951 #3338) +#2950 := (< uf_272 1099511627776::int) +#3340 := (implies #2950 #3339) +#2947 := (<= 0::int uf_271) +#2946 := (<= uf_271 uf_76) +#2948 := (and #2946 #2947) +#3341 := (implies #2948 #3340) +#2943 := (<= 0::int uf_270) +#2942 := (<= uf_270 uf_76) +#2944 := (and #2942 #2943) +#3342 := (implies #2944 #3341) +#2939 := (<= 0::int uf_269) +#2938 := (<= uf_269 uf_78) +#2940 := (and #2938 #2939) +#3343 := (implies #2940 #3342) +#3344 := (implies true #3343) +#3345 := (not #3344) +#14774 := (iff #3345 #14771) +#12868 := (and #3211 #12862) +#12847 := (not #3203) +#12848 := (or #12847 #3204) +#5718 := (and #285 #410) +#5727 := (not #5718) +#12854 := (or #5727 #12848) +#12859 := (forall (vars (?x785 int)) #12854) +#12892 := (not #12859) +#12893 := (or #12892 #12868) +#12901 := (and #12859 #12893) +#12909 := (not #3202) +#12910 := (or #12909 #12901) +#12918 := (and #3202 #12910) +#12933 := (not #3201) +#12934 := (or #12933 #12918) +#12943 := (or #12942 #12934) +#12832 := (+ 1::int uf_294) +#12844 := (= uf_302 #12832) +#12951 := (not #12844) +#12952 := (or #12951 #12943) +#12838 := (<= 0::int #12832) +#12835 := (<= #12832 uf_76) +#12841 := (and #12835 #12838) +#12960 := (not #12841) +#12961 := (or #12960 #12952) +#12969 := (and #12841 #12961) +#12829 := (and #3059 #3190) +#12977 := (not #12829) +#12978 := (or #12977 #12969) +#13082 := (or #12978 #13081) +#13091 := (or #13090 #13082) +#13106 := (not #3059) +#13107 := (or #13106 #13091) +#13116 := (or #13115 #13107) +#13125 := (or #13124 #13116) +#13134 := (or #13133 #13125) +#13143 := (or #13142 #13134) +#13151 := (and #12818 #13143) +#13160 := (or #13159 #13151) +#13168 := (and #12809 #13160) +#12687 := (not #3060) +#13176 := (or #12687 #13168) +#13191 := (or #12687 #13176) +#13199 := (not #3238) +#13200 := (or #13199 #13191) +#13208 := (or #12687 #13200) +#12994 := (or #12993 #12978) +#13003 := (or #13002 #12994) +#13018 := (or #12687 #13003) +#13026 := (or #12687 #13018) +#13041 := (or #12687 #13026) +#13049 := (not #3185) +#13050 := (or #13049 #13041) +#13058 := (or #12687 #13050) +#13220 := (and #13058 #13208) +#13226 := (or #12687 #13220) +#13234 := (or #13142 #13226) +#13242 := (and #12818 #13234) +#13250 := (or #13159 #13242) +#13258 := (and #12809 #13250) +#13266 := (or #12687 #13258) +#13281 := (or #12687 #13266) +#13289 := (not #3174) +#13290 := (or #13289 #13281) +#13298 := (or #12687 #13290) +#12594 := (and #3140 #5718) +#12597 := (and #3145 #12594) +#12600 := (exists (vars (?x782 int)) #12597) +#12579 := (not #3140) +#12580 := (or #12579 #3141) +#12586 := (or #5727 #12580) +#12591 := (forall (vars (?x781 int)) #12586) +#12620 := (not #12591) +#12621 := (or #12620 #12600) +#12629 := (and #12591 #12621) +#12645 := (or #12644 #12629) +#12654 := (or #12653 #12645) +#12663 := (or #12662 #12654) +#12672 := (or #12671 #12663) +#12688 := (or #12687 #12672) +#12696 := (or #12687 #12688) +#12711 := (or #12687 #12696) +#12720 := (or #12719 #12711) +#12728 := (and up_216 #12720) +#12736 := (or #12687 #12728) +#12751 := (or #12687 #12736) +#12759 := (or #12687 #12751) +#12774 := (or #12687 #12759) +#12782 := (not #3131) +#12783 := (or #12782 #12774) +#12791 := (or #12687 #12783) +#13310 := (and #12791 #13298) +#13316 := (or #12687 #13310) +#13332 := (or #13331 #13316) +#13341 := (or #13340 #13332) +#13350 := (or #13349 #13341) +#13359 := (or #13358 #13350) +#13368 := (or #13367 #13359) +#13377 := (or #13376 #13368) +#12544 := (and #3115 #12437) +#13385 := (not #12544) +#13386 := (or #13385 #13377) +#13394 := (or #13385 #13386) +#13402 := (or #12687 #13394) +#13417 := (or #12687 #13402) +#13432 := (or #12687 #13417) +#13508 := (or #12687 #13432) +#12432 := (and #3070 #12426) +#13516 := (not #12432) +#13517 := (or #13516 #13508) +#12411 := (not #3062) +#12412 := (or #12411 #3063) +#12418 := (or #5727 #12412) +#12423 := (forall (vars (?x774 int)) #12418) +#13525 := (not #12423) +#13526 := (or #13525 #13517) +#13534 := (not #3061) +#13535 := (or #13534 #13526) +#13543 := (or #12687 #13535) +#13551 := (not #3058) +#13552 := (or #13551 #13543) +#13560 := (not #3054) +#13561 := (or #13560 #13552) +#13569 := (not #3050) +#13570 := (or #13569 #13561) +#12406 := (and #2951 #3022) +#13585 := (not #12406) +#13586 := (or #13585 #13570) +#13594 := (and #12406 #13586) +#12386 := (not #3038) +#12387 := (or #12386 #3041) +#12393 := (or #5727 #12387) +#12398 := (forall (vars (?x773 int)) #12393) +#13602 := (not #12398) +#13603 := (or #13602 #13594) +#13611 := (and #12398 #13603) +#13619 := (not #3037) +#13620 := (or #13619 #13611) +#13628 := (and #3037 #13620) +#12380 := (and #3032 #3033) +#12383 := (and #3032 #12380) +#13636 := (not #12383) +#13637 := (or #13636 #13628) +#13646 := (or #13645 #13637) +#13655 := (or #13654 #13646) +#13664 := (or #13663 #13655) +#13673 := (or #13672 #13664) +#13682 := (or #13681 #13673) +#13690 := (and #12373 #13682) +#13699 := (or #13698 #13690) +#13707 := (and #12364 #13699) +#13716 := (or #13715 #13707) +#13724 := (and #12355 #13716) +#13732 := (not #3007) +#13733 := (or #13732 #13724) +#13742 := (or #13741 #13733) +#13751 := (or #13750 #13742) +#13760 := (or #13759 #13751) +#13769 := (or #13768 #13760) +#13777 := (not #2989) +#13778 := (or #13777 #13769) +#13787 := (or #13786 #13778) +#13796 := (or #13795 #13787) +#12314 := (and #12308 #12311) +#12317 := (and #12305 #12314) +#12320 := (and #12302 #12317) +#12323 := (and #12299 #12320) +#12326 := (and #12293 #12323) +#13811 := (not #12326) +#13812 := (or #13811 #13796) +#13820 := (not #2951) +#13821 := (or #13820 #13812) +#13829 := (not #2950) +#13830 := (or #13829 #13821) +#13838 := (not #2948) +#13839 := (or #13838 #13830) +#13847 := (not #2944) +#13848 := (or #13847 #13839) +#13856 := (not #2940) +#13857 := (or #13856 #13848) +#13869 := (not #13857) +#14772 := (iff #13869 #14771) +#14769 := (iff #13857 #14766) +#14724 := (or #13883 #14643) +#14727 := (or #13741 #14724) +#14730 := (or #13750 #14727) +#14733 := (or #13759 #14730) +#14736 := (or #13768 #14733) +#14739 := (or #14655 #14736) +#14742 := (or #13786 #14739) +#14745 := (or #13795 #14742) +#14748 := (or #14661 #14745) +#14751 := (or #13942 #14748) +#14754 := (or #14671 #14751) +#14757 := (or #14693 #14754) +#14760 := (or #14707 #14757) +#14763 := (or #14721 #14760) +#14767 := (iff #14763 #14766) +#14768 := [rewrite]: #14767 +#14764 := (iff #13857 #14763) +#14761 := (iff #13848 #14760) +#14758 := (iff #13839 #14757) +#14755 := (iff #13830 #14754) +#14752 := (iff #13821 #14751) +#14749 := (iff #13812 #14748) +#14746 := (iff #13796 #14745) +#14743 := (iff #13787 #14742) +#14740 := (iff #13778 #14739) +#14737 := (iff #13769 #14736) +#14734 := (iff #13760 #14733) +#14731 := (iff #13751 #14730) +#14728 := (iff #13742 #14727) +#14725 := (iff #13733 #14724) +#14644 := (iff #13724 #14643) +#14641 := (iff #13716 #14640) +#14638 := (iff #13707 #14635) +#14632 := (and #12364 #14629) +#14636 := (iff #14632 #14635) +#14637 := [rewrite]: #14636 +#14633 := (iff #13707 #14632) +#14630 := (iff #13699 #14629) +#14627 := (iff #13690 #14624) +#14621 := (and #12373 #14616) +#14625 := (iff #14621 #14624) +#14626 := [rewrite]: #14625 +#14622 := (iff #13690 #14621) +#14619 := (iff #13682 #14616) +#14598 := (or false #14595) +#14601 := (or #13645 #14598) +#14604 := (or #13654 #14601) +#14607 := (or #13663 #14604) +#14610 := (or #13672 #14607) +#14613 := (or #13681 #14610) +#14617 := (iff #14613 #14616) +#14618 := [rewrite]: #14617 +#14614 := (iff #13682 #14613) +#14611 := (iff #13673 #14610) +#14608 := (iff #13664 #14607) +#14605 := (iff #13655 #14604) +#14602 := (iff #13646 #14601) +#14599 := (iff #13637 #14598) +#14596 := (iff #13628 #14595) +#14593 := (iff #13620 #14592) +#14590 := (iff #13611 #14589) +#14587 := (iff #13603 #14586) +#14584 := (iff #13594 #14581) +#14499 := (and #13943 #3022) +#14578 := (and #14499 #14573) +#14582 := (iff #14578 #14581) +#14583 := [rewrite]: #14582 +#14579 := (iff #13594 #14578) +#14576 := (iff #13586 #14573) +#14510 := (or #13955 #14376) +#14513 := (or #13331 #14510) +#14516 := (or #13340 #14513) +#14519 := (or #13349 #14516) +#14522 := (or #13358 #14519) +#14525 := (or #13367 #14522) +#14528 := (or #13376 #14525) +#14531 := (or #14399 #14528) +#14534 := (or #14399 #14531) +#14537 := (or #13955 #14534) +#14540 := (or #13955 #14537) +#14543 := (or #13955 #14540) +#14546 := (or #13955 #14543) +#14549 := (or #14416 #14546) +#14552 := (or #14450 #14549) +#14555 := (or #14456 #14552) +#14558 := (or #13955 #14555) +#14561 := (or #14468 #14558) +#14564 := (or #14483 #14561) +#14567 := (or #14496 #14564) +#14570 := (or #14507 #14567) +#14574 := (iff #14570 #14573) +#14575 := [rewrite]: #14574 +#14571 := (iff #13586 #14570) +#14568 := (iff #13570 #14567) +#14565 := (iff #13561 #14564) +#14562 := (iff #13552 #14561) +#14559 := (iff #13543 #14558) +#14556 := (iff #13535 #14555) +#14553 := (iff #13526 #14552) +#14550 := (iff #13517 #14549) +#14547 := (iff #13508 #14546) +#14544 := (iff #13432 #14543) +#14541 := (iff #13417 #14540) +#14538 := (iff #13402 #14537) +#14535 := (iff #13394 #14534) +#14532 := (iff #13386 #14531) +#14529 := (iff #13377 #14528) +#14526 := (iff #13368 #14525) +#14523 := (iff #13359 #14522) +#14520 := (iff #13350 #14519) +#14517 := (iff #13341 #14516) +#14514 := (iff #13332 #14513) +#14511 := (iff #13316 #14510) +#14377 := (iff #13310 #14376) +#14374 := (iff #13298 #14371) +#14359 := (or #13955 #14345) +#14362 := (or #13955 #14359) +#14365 := (or #14046 #14362) +#14368 := (or #13955 #14365) +#14372 := (iff #14368 #14371) +#14373 := [rewrite]: #14372 +#14369 := (iff #13298 #14368) +#14366 := (iff #13290 #14365) +#14363 := (iff #13281 #14362) +#14360 := (iff #13266 #14359) +#14348 := (iff #13258 #14345) +#14342 := (and #12809 #14339) +#14346 := (iff #14342 #14345) +#14347 := [rewrite]: #14346 +#14343 := (iff #13258 #14342) +#14340 := (iff #13250 #14339) +#14337 := (iff #13242 #14334) +#14331 := (and #12818 #14326) +#14335 := (iff #14331 #14334) +#14336 := [rewrite]: #14335 +#14332 := (iff #13242 #14331) +#14329 := (iff #13234 #14326) +#14320 := (or #13955 #14317) +#14323 := (or #13142 #14320) +#14327 := (iff #14323 #14326) +#14328 := [rewrite]: #14327 +#14324 := (iff #13234 #14323) +#14321 := (iff #13226 #14320) +#14318 := (iff #13220 #14317) +#14315 := (iff #13208 #14312) +#14300 := (or #13955 #14286) +#14303 := (or #13955 #14300) +#14306 := (or #14207 #14303) +#14309 := (or #13955 #14306) +#14313 := (iff #14309 #14312) +#14314 := [rewrite]: #14313 +#14310 := (iff #13208 #14309) +#14307 := (iff #13200 #14306) +#14304 := (iff #13191 #14303) +#14301 := (iff #13176 #14300) +#14289 := (iff #13168 #14286) +#14283 := (and #12809 #14280) +#14287 := (iff #14283 #14286) +#14288 := [rewrite]: #14287 +#14284 := (iff #13168 #14283) +#14281 := (iff #13160 #14280) +#14278 := (iff #13151 #14275) +#14272 := (and #12818 #14267) +#14276 := (iff #14272 #14275) +#14277 := [rewrite]: #14276 +#14273 := (iff #13151 #14272) +#14270 := (iff #13143 #14267) +#14214 := (or #14081 #14201) +#14246 := (or #14214 #13081) +#14249 := (or #13090 #14246) +#14252 := (or #14243 #14249) +#14255 := (or #13115 #14252) +#14258 := (or #13124 #14255) +#14261 := (or #13133 #14258) +#14264 := (or #13142 #14261) +#14268 := (iff #14264 #14267) +#14269 := [rewrite]: #14268 +#14265 := (iff #13143 #14264) +#14262 := (iff #13134 #14261) +#14259 := (iff #13125 #14258) +#14256 := (iff #13116 #14255) +#14253 := (iff #13107 #14252) +#14250 := (iff #13091 #14249) +#14247 := (iff #13082 #14246) +#14215 := (iff #12978 #14214) +#14204 := (iff #12969 #14201) +#14198 := (and #14175 #14193) +#14202 := (iff #14198 #14201) +#14203 := [rewrite]: #14202 +#14199 := (iff #12969 #14198) +#14196 := (iff #12961 #14193) +#14181 := (or #14097 #14165) +#14184 := (or #12942 #14181) +#14187 := (or #14172 #14184) +#14190 := (or #14178 #14187) +#14194 := (iff #14190 #14193) +#14195 := [rewrite]: #14194 +#14191 := (iff #12961 #14190) +#14188 := (iff #12952 #14187) +#14185 := (iff #12943 #14184) +#14182 := (iff #12934 #14181) +#14166 := (iff #12918 #14165) +#14163 := (iff #12910 #14162) +#14160 := (iff #12901 #14159) +#14157 := (iff #12893 #14156) +#14154 := (iff #12868 #14151) +#14148 := (and #14145 #12862) +#14152 := (iff #14148 #14151) +#14153 := [rewrite]: #14152 +#14149 := (iff #12868 #14148) +#14146 := (iff #3211 #14145) +#14147 := [rewrite]: #14146 +#14150 := [monotonicity #14147]: #14149 +#14155 := [trans #14150 #14153]: #14154 +#14140 := (iff #12892 #14139) +#14137 := (iff #12859 #14136) +#14134 := (iff #12854 #14131) +#14125 := (or #14109 #14122) +#14128 := (or #5739 #14125) +#14132 := (iff #14128 #14131) +#14133 := [rewrite]: #14132 +#14129 := (iff #12854 #14128) +#14126 := (iff #12848 #14125) +#14123 := (iff #3204 #14122) +#14124 := [rewrite]: #14123 +#14118 := (iff #12847 #14109) +#14108 := (not #14109) +#14113 := (not #14108) +#14116 := (iff #14113 #14109) +#14117 := [rewrite]: #14116 +#14114 := (iff #12847 #14113) +#14111 := (iff #3203 #14108) +#14112 := [rewrite]: #14111 +#14115 := [monotonicity #14112]: #14114 +#14119 := [trans #14115 #14117]: #14118 +#14127 := [monotonicity #14119 #14124]: #14126 +#5740 := (iff #5727 #5739) +#5737 := (iff #5718 #5736) +#4420 := (iff #410 #4419) +#4421 := [rewrite]: #4420 +#4083 := (iff #285 #4084) +#4085 := [rewrite]: #4083 +#5738 := [monotonicity #4085 #4421]: #5737 +#5741 := [monotonicity #5738]: #5740 +#14130 := [monotonicity #5741 #14127]: #14129 +#14135 := [trans #14130 #14133]: #14134 +#14138 := [quant-intro #14135]: #14137 +#14141 := [monotonicity #14138]: #14140 +#14158 := [monotonicity #14141 #14155]: #14157 +#14161 := [monotonicity #14138 #14158]: #14160 +#14106 := (iff #12909 #14105) +#14103 := (iff #3202 #14100) +#14104 := [rewrite]: #14103 +#14107 := [monotonicity #14104]: #14106 +#14164 := [monotonicity #14107 #14161]: #14163 +#14167 := [monotonicity #14104 #14164]: #14166 +#14098 := (iff #12933 #14097) +#14095 := (iff #3201 #14094) +#14091 := (iff #3200 #14092) +#14093 := [rewrite]: #14091 +#14075 := (iff #3190 #14076) +#14077 := [rewrite]: #14075 +#14096 := [monotonicity #14077 #14093]: #14095 +#14099 := [monotonicity #14096]: #14098 +#14183 := [monotonicity #14099 #14167]: #14182 +#14186 := [monotonicity #14183]: #14185 +#14173 := (iff #12951 #14172) +#14170 := (iff #12844 #14168) +#14171 := [rewrite]: #14170 +#14174 := [monotonicity #14171]: #14173 +#14189 := [monotonicity #14174 #14186]: #14188 +#14179 := (iff #12960 #14178) +#14176 := (iff #12841 #14175) +#14089 := (iff #12838 #14088) +#14090 := [rewrite]: #14089 +#14086 := (iff #12835 #14084) +#14087 := [rewrite]: #14086 +#14177 := [monotonicity #14087 #14090]: #14176 +#14180 := [monotonicity #14177]: #14179 +#14192 := [monotonicity #14180 #14189]: #14191 +#14197 := [trans #14192 #14195]: #14196 +#14200 := [monotonicity #14177 #14197]: #14199 +#14205 := [trans #14200 #14203]: #14204 +#14082 := (iff #12977 #14081) +#14079 := (iff #12829 #14078) +#13949 := (iff #3059 #13950) +#13951 := [rewrite]: #13949 +#14080 := [monotonicity #13951 #14077]: #14079 +#14083 := [monotonicity #14080]: #14082 +#14216 := [monotonicity #14083 #14205]: #14215 +#14248 := [monotonicity #14216]: #14247 +#14251 := [monotonicity #14248]: #14250 +#14244 := (iff #13106 #14243) +#14245 := [monotonicity #13951]: #14244 +#14254 := [monotonicity #14245 #14251]: #14253 +#14257 := [monotonicity #14254]: #14256 +#14260 := [monotonicity #14257]: #14259 +#14263 := [monotonicity #14260]: #14262 +#14266 := [monotonicity #14263]: #14265 +#14271 := [trans #14266 #14269]: #14270 +#14274 := [monotonicity #14271]: #14273 +#14279 := [trans #14274 #14277]: #14278 +#14282 := [monotonicity #14279]: #14281 +#14285 := [monotonicity #14282]: #14284 +#14290 := [trans #14285 #14288]: #14289 +#13956 := (iff #12687 #13955) +#13953 := (iff #3060 #13952) +#13946 := (iff #3053 #13947) +#13948 := [rewrite]: #13946 +#13954 := [monotonicity #13948 #13951]: #13953 +#13957 := [monotonicity #13954]: #13956 +#14302 := [monotonicity #13957 #14290]: #14301 +#14305 := [monotonicity #13957 #14302]: #14304 +#14298 := (iff #13199 #14207) +#14293 := (not #14211) +#14296 := (iff #14293 #14207) +#14297 := [rewrite]: #14296 +#14294 := (iff #13199 #14293) +#14291 := (iff #3238 #14211) +#14292 := [rewrite]: #14291 +#14295 := [monotonicity #14292]: #14294 +#14299 := [trans #14295 #14297]: #14298 +#14308 := [monotonicity #14299 #14305]: #14307 +#14311 := [monotonicity #13957 #14308]: #14310 +#14316 := [trans #14311 #14314]: #14315 +#14241 := (iff #13058 #14238) +#14217 := (or #12993 #14214) +#14220 := (or #13002 #14217) +#14223 := (or #13955 #14220) +#14226 := (or #13955 #14223) +#14229 := (or #13955 #14226) +#14232 := (or #14211 #14229) +#14235 := (or #13955 #14232) +#14239 := (iff #14235 #14238) +#14240 := [rewrite]: #14239 +#14236 := (iff #13058 #14235) +#14233 := (iff #13050 #14232) +#14230 := (iff #13041 #14229) +#14227 := (iff #13026 #14226) +#14224 := (iff #13018 #14223) +#14221 := (iff #13003 #14220) +#14218 := (iff #12994 #14217) +#14219 := [monotonicity #14216]: #14218 +#14222 := [monotonicity #14219]: #14221 +#14225 := [monotonicity #13957 #14222]: #14224 +#14228 := [monotonicity #13957 #14225]: #14227 +#14231 := [monotonicity #13957 #14228]: #14230 +#14212 := (iff #13049 #14211) +#14206 := (iff #3185 #14207) +#14210 := [rewrite]: #14206 +#14213 := [monotonicity #14210]: #14212 +#14234 := [monotonicity #14213 #14231]: #14233 +#14237 := [monotonicity #13957 #14234]: #14236 +#14242 := [trans #14237 #14240]: #14241 +#14319 := [monotonicity #14242 #14316]: #14318 +#14322 := [monotonicity #13957 #14319]: #14321 +#14325 := [monotonicity #14322]: #14324 +#14330 := [trans #14325 #14328]: #14329 +#14333 := [monotonicity #14330]: #14332 +#14338 := [trans #14333 #14336]: #14337 +#14341 := [monotonicity #14338]: #14340 +#14344 := [monotonicity #14341]: #14343 +#14349 := [trans #14344 #14347]: #14348 +#14361 := [monotonicity #13957 #14349]: #14360 +#14364 := [monotonicity #13957 #14361]: #14363 +#14357 := (iff #13289 #14046) +#14352 := (not #14049) +#14355 := (iff #14352 #14046) +#14356 := [rewrite]: #14355 +#14353 := (iff #13289 #14352) +#14350 := (iff #3174 #14049) +#14351 := [rewrite]: #14350 +#14354 := [monotonicity #14351]: #14353 +#14358 := [trans #14354 #14356]: #14357 +#14367 := [monotonicity #14358 #14364]: #14366 +#14370 := [monotonicity #13957 #14367]: #14369 +#14375 := [trans #14370 #14373]: #14374 +#14073 := (iff #12791 #14070) +#14052 := (or #13955 #14041) +#14055 := (or #13955 #14052) +#14058 := (or #13955 #14055) +#14061 := (or #13955 #14058) +#14064 := (or #14049 #14061) +#14067 := (or #13955 #14064) +#14071 := (iff #14067 #14070) +#14072 := [rewrite]: #14071 +#14068 := (iff #12791 #14067) +#14065 := (iff #12783 #14064) +#14062 := (iff #12774 #14061) +#14059 := (iff #12759 #14058) +#14056 := (iff #12751 #14055) +#14053 := (iff #12736 #14052) +#14042 := (iff #12728 #14041) +#14039 := (iff #12720 #14036) +#14012 := (or #12644 #14009) +#14015 := (or #12653 #14012) +#14018 := (or #12662 #14015) +#14021 := (or #12671 #14018) +#14024 := (or #13955 #14021) +#14027 := (or #13955 #14024) +#14030 := (or #13955 #14027) +#14033 := (or #12719 #14030) +#14037 := (iff #14033 #14036) +#14038 := [rewrite]: #14037 +#14034 := (iff #12720 #14033) +#14031 := (iff #12711 #14030) +#14028 := (iff #12696 #14027) +#14025 := (iff #12688 #14024) +#14022 := (iff #12672 #14021) +#14019 := (iff #12663 #14018) +#14016 := (iff #12654 #14015) +#14013 := (iff #12645 #14012) +#14010 := (iff #12629 #14009) +#14007 := (iff #12621 #14006) +#14004 := (iff #12600 #14003) +#14001 := (iff #12597 #13998) +#13992 := (and #13958 #5736) +#13995 := (and #3145 #13992) +#13999 := (iff #13995 #13998) +#14000 := [rewrite]: #13999 +#13996 := (iff #12597 #13995) +#13993 := (iff #12594 #13992) +#13961 := (iff #3140 #13958) +#13962 := [rewrite]: #13961 +#13994 := [monotonicity #13962 #5738]: #13993 +#13997 := [monotonicity #13994]: #13996 +#14002 := [trans #13997 #14000]: #14001 +#14005 := [quant-intro #14002]: #14004 +#13990 := (iff #12620 #13989) +#13987 := (iff #12591 #13986) +#13984 := (iff #12586 #13981) +#13975 := (or #13959 #13972) +#13978 := (or #5739 #13975) +#13982 := (iff #13978 #13981) +#13983 := [rewrite]: #13982 +#13979 := (iff #12586 #13978) +#13976 := (iff #12580 #13975) +#13973 := (iff #3141 #13972) +#13974 := [rewrite]: #13973 +#13968 := (iff #12579 #13959) +#13963 := (not #13958) +#13966 := (iff #13963 #13959) +#13967 := [rewrite]: #13966 +#13964 := (iff #12579 #13963) +#13965 := [monotonicity #13962]: #13964 +#13969 := [trans #13965 #13967]: #13968 +#13977 := [monotonicity #13969 #13974]: #13976 +#13980 := [monotonicity #5741 #13977]: #13979 +#13985 := [trans #13980 #13983]: #13984 +#13988 := [quant-intro #13985]: #13987 +#13991 := [monotonicity #13988]: #13990 +#14008 := [monotonicity #13991 #14005]: #14007 +#14011 := [monotonicity #13988 #14008]: #14010 +#14014 := [monotonicity #14011]: #14013 +#14017 := [monotonicity #14014]: #14016 +#14020 := [monotonicity #14017]: #14019 +#14023 := [monotonicity #14020]: #14022 +#14026 := [monotonicity #13957 #14023]: #14025 +#14029 := [monotonicity #13957 #14026]: #14028 +#14032 := [monotonicity #13957 #14029]: #14031 +#14035 := [monotonicity #14032]: #14034 +#14040 := [trans #14035 #14038]: #14039 +#14043 := [monotonicity #14040]: #14042 +#14054 := [monotonicity #13957 #14043]: #14053 +#14057 := [monotonicity #13957 #14054]: #14056 +#14060 := [monotonicity #13957 #14057]: #14059 +#14063 := [monotonicity #13957 #14060]: #14062 +#14050 := (iff #12782 #14049) +#14047 := (iff #3131 #14046) +#14048 := [rewrite]: #14047 +#14051 := [monotonicity #14048]: #14050 +#14066 := [monotonicity #14051 #14063]: #14065 +#14069 := [monotonicity #13957 #14066]: #14068 +#14074 := [trans #14069 #14072]: #14073 +#14378 := [monotonicity #14074 #14375]: #14377 +#14512 := [monotonicity #13957 #14378]: #14511 +#14515 := [monotonicity #14512]: #14514 +#14518 := [monotonicity #14515]: #14517 +#14521 := [monotonicity #14518]: #14520 +#14524 := [monotonicity #14521]: #14523 +#14527 := [monotonicity #14524]: #14526 +#14530 := [monotonicity #14527]: #14529 +#14400 := (iff #13385 #14399) +#14397 := (iff #12544 #12437) +#12517 := (and true true) +#14392 := (and #12517 #12437) +#14395 := (iff #14392 #12437) +#14396 := [rewrite]: #14395 +#14393 := (iff #12544 #14392) +#14390 := (iff #3115 #12517) +#14388 := (iff #3078 true) +#14389 := [rewrite]: #14388 +#14386 := (iff #3077 true) +#14381 := (forall (vars (?x775 T5)) (:pat #3075) true) +#14384 := (iff #14381 true) +#14385 := [elim-unused]: #14384 +#14382 := (iff #3077 #14381) +#14379 := (iff #3076 true) +#14380 := [rewrite]: #14379 +#14383 := [quant-intro #14380]: #14382 +#14387 := [trans #14383 #14385]: #14386 +#14391 := [monotonicity #14387 #14389]: #14390 +#14394 := [monotonicity #14391]: #14393 +#14398 := [trans #14394 #14396]: #14397 +#14401 := [monotonicity #14398]: #14400 +#14533 := [monotonicity #14401 #14530]: #14532 +#14536 := [monotonicity #14401 #14533]: #14535 +#14539 := [monotonicity #13957 #14536]: #14538 +#14542 := [monotonicity #13957 #14539]: #14541 +#14545 := [monotonicity #13957 #14542]: #14544 +#14548 := [monotonicity #13957 #14545]: #14547 +#14417 := (iff #13516 #14416) +#14414 := (iff #12432 #14411) +#14408 := (and #14405 #12426) +#14412 := (iff #14408 #14411) +#14413 := [rewrite]: #14412 +#14409 := (iff #12432 #14408) +#14406 := (iff #3070 #14405) +#14407 := [rewrite]: #14406 +#14410 := [monotonicity #14407]: #14409 +#14415 := [trans #14410 #14413]: #14414 +#14418 := [monotonicity #14415]: #14417 +#14551 := [monotonicity #14418 #14548]: #14550 +#14451 := (iff #13525 #14450) +#14448 := (iff #12423 #14447) +#14445 := (iff #12418 #14442) +#14436 := (or #14420 #14433) +#14439 := (or #5739 #14436) +#14443 := (iff #14439 #14442) +#14444 := [rewrite]: #14443 +#14440 := (iff #12418 #14439) +#14437 := (iff #12412 #14436) +#14434 := (iff #3063 #14433) +#14435 := [rewrite]: #14434 +#14429 := (iff #12411 #14420) +#14419 := (not #14420) +#14424 := (not #14419) +#14427 := (iff #14424 #14420) +#14428 := [rewrite]: #14427 +#14425 := (iff #12411 #14424) +#14422 := (iff #3062 #14419) +#14423 := [rewrite]: #14422 +#14426 := [monotonicity #14423]: #14425 +#14430 := [trans #14426 #14428]: #14429 +#14438 := [monotonicity #14430 #14435]: #14437 +#14441 := [monotonicity #5741 #14438]: #14440 +#14446 := [trans #14441 #14444]: #14445 +#14449 := [quant-intro #14446]: #14448 +#14452 := [monotonicity #14449]: #14451 +#14554 := [monotonicity #14452 #14551]: #14553 +#14457 := (iff #13534 #14456) +#14454 := (iff #3061 #14453) +#14455 := [rewrite]: #14454 +#14458 := [monotonicity #14455]: #14457 +#14557 := [monotonicity #14458 #14554]: #14556 +#14560 := [monotonicity #13957 #14557]: #14559 +#14469 := (iff #13551 #14468) +#14466 := (iff #3058 #14465) +#14463 := (iff #3057 #14462) +#14464 := [rewrite]: #14463 +#14460 := (iff #3056 #14459) +#14461 := [rewrite]: #14460 +#14467 := [monotonicity #14461 #14464]: #14466 +#14470 := [monotonicity #14467]: #14469 +#14563 := [monotonicity #14470 #14560]: #14562 +#14484 := (iff #13560 #14483) +#14481 := (iff #3054 #14478) +#14475 := (and #14471 #13947) +#14479 := (iff #14475 #14478) +#14480 := [rewrite]: #14479 +#14476 := (iff #3054 #14475) +#14473 := (iff #3052 #14471) +#14474 := [rewrite]: #14473 +#14477 := [monotonicity #14474 #13948]: #14476 +#14482 := [trans #14477 #14480]: #14481 +#14485 := [monotonicity #14482]: #14484 +#14566 := [monotonicity #14485 #14563]: #14565 +#14497 := (iff #13569 #14496) +#14494 := (iff #3050 #14493) +#14491 := (iff #3049 #14490) +#14492 := [rewrite]: #14491 +#14488 := (iff #3048 #14486) +#14489 := [rewrite]: #14488 +#14495 := [monotonicity #14489 #14492]: #14494 +#14498 := [monotonicity #14495]: #14497 +#14569 := [monotonicity #14498 #14566]: #14568 +#14508 := (iff #13585 #14507) +#14505 := (iff #12406 #14502) +#14503 := (iff #14499 #14502) +#14504 := [rewrite]: #14503 +#14500 := (iff #12406 #14499) +#13944 := (iff #2951 #13943) +#13945 := [rewrite]: #13944 +#14501 := [monotonicity #13945]: #14500 +#14506 := [trans #14501 #14504]: #14505 +#14509 := [monotonicity #14506]: #14508 +#14572 := [monotonicity #14509 #14569]: #14571 +#14577 := [trans #14572 #14575]: #14576 +#14580 := [monotonicity #14501 #14577]: #14579 +#14585 := [trans #14580 #14583]: #14584 +#13940 := (iff #13602 #13939) +#13937 := (iff #12398 #13936) +#13934 := (iff #12393 #13931) +#13925 := (or #13910 #13921) +#13928 := (or #5739 #13925) +#13932 := (iff #13928 #13931) +#13933 := [rewrite]: #13932 +#13929 := (iff #12393 #13928) +#13926 := (iff #12387 #13925) +#13920 := (iff #3041 #13921) +#13924 := [rewrite]: #13920 +#13918 := (iff #12386 #13910) +#13909 := (not #13910) +#13913 := (not #13909) +#13916 := (iff #13913 #13910) +#13917 := [rewrite]: #13916 +#13914 := (iff #12386 #13913) +#13911 := (iff #3038 #13909) +#13912 := [rewrite]: #13911 +#13915 := [monotonicity #13912]: #13914 +#13919 := [trans #13915 #13917]: #13918 +#13927 := [monotonicity #13919 #13924]: #13926 +#13930 := [monotonicity #5741 #13927]: #13929 +#13935 := [trans #13930 #13933]: #13934 +#13938 := [quant-intro #13935]: #13937 +#13941 := [monotonicity #13938]: #13940 +#14588 := [monotonicity #13941 #14585]: #14587 +#14591 := [monotonicity #13938 #14588]: #14590 +#13907 := (iff #13619 #13906) +#13904 := (iff #3037 #13903) +#13905 := [rewrite]: #13904 +#13908 := [monotonicity #13905]: #13907 +#14594 := [monotonicity #13908 #14591]: #14593 +#14597 := [monotonicity #13905 #14594]: #14596 +#13901 := (iff #13636 false) +#13444 := (iff #3294 false) +#13445 := [rewrite]: #13444 +#13899 := (iff #13636 #3294) +#13897 := (iff #12383 true) +#13892 := (and true #12517) +#13895 := (iff #13892 true) +#13896 := [rewrite]: #13895 +#13893 := (iff #12383 #13892) +#13890 := (iff #12380 #12517) +#13888 := (iff #3033 true) +#13889 := [rewrite]: #13888 +#13886 := (iff #3032 true) +#13887 := [rewrite]: #13886 +#13891 := [monotonicity #13887 #13889]: #13890 +#13894 := [monotonicity #13887 #13891]: #13893 +#13898 := [trans #13894 #13896]: #13897 +#13900 := [monotonicity #13898]: #13899 +#13902 := [trans #13900 #13445]: #13901 +#14600 := [monotonicity #13902 #14597]: #14599 +#14603 := [monotonicity #14600]: #14602 +#14606 := [monotonicity #14603]: #14605 +#14609 := [monotonicity #14606]: #14608 +#14612 := [monotonicity #14609]: #14611 +#14615 := [monotonicity #14612]: #14614 +#14620 := [trans #14615 #14618]: #14619 +#14623 := [monotonicity #14620]: #14622 +#14628 := [trans #14623 #14626]: #14627 +#14631 := [monotonicity #14628]: #14630 +#14634 := [monotonicity #14631]: #14633 +#14639 := [trans #14634 #14637]: #14638 +#14642 := [monotonicity #14639]: #14641 +#14645 := [monotonicity #14642]: #14644 +#13884 := (iff #13732 #13883) +#13881 := (iff #3007 #13880) +#13878 := (iff #3006 #13877) +#13879 := [rewrite]: #13878 +#13875 := (iff #3005 #13872) +#13876 := [rewrite]: #13875 +#13882 := [monotonicity #13876 #13879]: #13881 +#13885 := [monotonicity #13882]: #13884 +#14726 := [monotonicity #13885 #14645]: #14725 +#14729 := [monotonicity #14726]: #14728 +#14732 := [monotonicity #14729]: #14731 +#14735 := [monotonicity #14732]: #14734 +#14738 := [monotonicity #14735]: #14737 +#14656 := (iff #13777 #14655) +#14653 := (iff #2989 #14652) +#14650 := (iff #2988 #14646) +#14651 := [rewrite]: #14650 +#14654 := [quant-intro #14651]: #14653 +#14657 := [monotonicity #14654]: #14656 +#14741 := [monotonicity #14657 #14738]: #14740 +#14744 := [monotonicity #14741]: #14743 +#14747 := [monotonicity #14744]: #14746 +#14662 := (iff #13811 #14661) +#14659 := (iff #12326 #14658) +#14660 := [rewrite]: #14659 +#14663 := [monotonicity #14660]: #14662 +#14750 := [monotonicity #14663 #14747]: #14749 +#14669 := (iff #13820 #13942) +#14664 := (not #13943) +#14667 := (iff #14664 #13942) +#14668 := [rewrite]: #14667 +#14665 := (iff #13820 #14664) +#14666 := [monotonicity #13945]: #14665 +#14670 := [trans #14666 #14668]: #14669 +#14753 := [monotonicity #14670 #14750]: #14752 +#14680 := (iff #13829 #14671) +#14672 := (not #14671) +#14675 := (not #14672) +#14678 := (iff #14675 #14671) +#14679 := [rewrite]: #14678 +#14676 := (iff #13829 #14675) +#14673 := (iff #2950 #14672) +#14674 := [rewrite]: #14673 +#14677 := [monotonicity #14674]: #14676 +#14681 := [trans #14677 #14679]: #14680 +#14756 := [monotonicity #14681 #14753]: #14755 +#14694 := (iff #13838 #14693) +#14691 := (iff #2948 #14690) +#14688 := (iff #2947 #14687) +#14689 := [rewrite]: #14688 +#14685 := (iff #2946 #14682) +#14686 := [rewrite]: #14685 +#14692 := [monotonicity #14686 #14689]: #14691 +#14695 := [monotonicity #14692]: #14694 +#14759 := [monotonicity #14695 #14756]: #14758 +#14708 := (iff #13847 #14707) +#14705 := (iff #2944 #14704) +#14702 := (iff #2943 #14701) +#14703 := [rewrite]: #14702 +#14699 := (iff #2942 #14696) +#14700 := [rewrite]: #14699 +#14706 := [monotonicity #14700 #14703]: #14705 +#14709 := [monotonicity #14706]: #14708 +#14762 := [monotonicity #14709 #14759]: #14761 +#14722 := (iff #13856 #14721) +#14719 := (iff #2940 #14718) +#14716 := (iff #2939 #14715) +#14717 := [rewrite]: #14716 +#14713 := (iff #2938 #14710) +#14714 := [rewrite]: #14713 +#14720 := [monotonicity #14714 #14717]: #14719 +#14723 := [monotonicity #14720]: #14722 +#14765 := [monotonicity #14723 #14762]: #14764 +#14770 := [trans #14765 #14768]: #14769 +#14773 := [monotonicity #14770]: #14772 +#13870 := (iff #3345 #13869) +#13867 := (iff #3344 #13857) +#13862 := (implies true #13857) +#13865 := (iff #13862 #13857) +#13866 := [rewrite]: #13865 +#13863 := (iff #3344 #13862) +#13860 := (iff #3343 #13857) +#13853 := (implies #2940 #13848) +#13858 := (iff #13853 #13857) +#13859 := [rewrite]: #13858 +#13854 := (iff #3343 #13853) +#13851 := (iff #3342 #13848) +#13844 := (implies #2944 #13839) +#13849 := (iff #13844 #13848) +#13850 := [rewrite]: #13849 +#13845 := (iff #3342 #13844) +#13842 := (iff #3341 #13839) +#13835 := (implies #2948 #13830) +#13840 := (iff #13835 #13839) +#13841 := [rewrite]: #13840 +#13836 := (iff #3341 #13835) +#13833 := (iff #3340 #13830) +#13826 := (implies #2950 #13821) +#13831 := (iff #13826 #13830) +#13832 := [rewrite]: #13831 +#13827 := (iff #3340 #13826) +#13824 := (iff #3339 #13821) +#13817 := (implies #2951 #13812) +#13822 := (iff #13817 #13821) +#13823 := [rewrite]: #13822 +#13818 := (iff #3339 #13817) +#13815 := (iff #3338 #13812) +#13808 := (implies #12326 #13796) +#13813 := (iff #13808 #13812) +#13814 := [rewrite]: #13813 +#13809 := (iff #3338 #13808) +#13806 := (iff #3337 #13796) +#13801 := (implies true #13796) +#13804 := (iff #13801 #13796) +#13805 := [rewrite]: #13804 +#13802 := (iff #3337 #13801) +#13799 := (iff #3336 #13796) +#13792 := (implies #12329 #13787) +#13797 := (iff #13792 #13796) +#13798 := [rewrite]: #13797 +#13793 := (iff #3336 #13792) +#13790 := (iff #3335 #13787) +#13783 := (implies #12338 #13778) +#13788 := (iff #13783 #13787) +#13789 := [rewrite]: #13788 +#13784 := (iff #3335 #13783) +#13781 := (iff #3334 #13778) +#13774 := (implies #2989 #13769) +#13779 := (iff #13774 #13778) +#13780 := [rewrite]: #13779 +#13775 := (iff #3334 #13774) +#13772 := (iff #3333 #13769) +#13765 := (implies #2994 #13760) +#13770 := (iff #13765 #13769) +#13771 := [rewrite]: #13770 +#13766 := (iff #3333 #13765) +#13763 := (iff #3332 #13760) +#13756 := (implies #2996 #13751) +#13761 := (iff #13756 #13760) +#13762 := [rewrite]: #13761 +#13757 := (iff #3332 #13756) +#13754 := (iff #3331 #13751) +#13747 := (implies #2999 #13742) +#13752 := (iff #13747 #13751) +#13753 := [rewrite]: #13752 +#13748 := (iff #3331 #13747) +#13745 := (iff #3330 #13742) +#13738 := (implies #12352 #13733) +#13743 := (iff #13738 #13742) +#13744 := [rewrite]: #13743 +#13739 := (iff #3330 #13738) +#13736 := (iff #3329 #13733) +#13729 := (implies #3007 #13724) +#13734 := (iff #13729 #13733) +#13735 := [rewrite]: #13734 +#13730 := (iff #3329 #13729) +#13727 := (iff #3328 #13724) +#13721 := (and #13716 #12355) +#13725 := (iff #13721 #13724) +#13726 := [rewrite]: #13725 +#13722 := (iff #3328 #13721) +#12356 := (iff #3010 #12355) +#12357 := [rewrite]: #12356 +#13719 := (iff #3327 #13716) +#13712 := (implies #12355 #13707) +#13717 := (iff #13712 #13716) +#13718 := [rewrite]: #13717 +#13713 := (iff #3327 #13712) +#13710 := (iff #3326 #13707) +#13704 := (and #13699 #12364) +#13708 := (iff #13704 #13707) +#13709 := [rewrite]: #13708 +#13705 := (iff #3326 #13704) +#12365 := (iff #3016 #12364) +#12362 := (iff #3015 #12361) +#12363 := [rewrite]: #12362 +#12359 := (iff #3013 #12358) +#12360 := [rewrite]: #12359 +#12366 := [monotonicity #12360 #12363]: #12365 +#13702 := (iff #3325 #13699) +#13695 := (implies #12364 #13690) +#13700 := (iff #13695 #13699) +#13701 := [rewrite]: #13700 +#13696 := (iff #3325 #13695) +#13693 := (iff #3324 #13690) +#13687 := (and #13682 #12373) +#13691 := (iff #13687 #13690) +#13692 := [rewrite]: #13691 +#13688 := (iff #3324 #13687) +#12376 := (iff #3019 #12373) +#12370 := (and #12367 #12361) +#12374 := (iff #12370 #12373) +#12375 := [rewrite]: #12374 +#12371 := (iff #3019 #12370) +#12368 := (iff #3018 #12367) +#12369 := [rewrite]: #12368 +#12372 := [monotonicity #12369 #12363]: #12371 +#12377 := [trans #12372 #12375]: #12376 +#13685 := (iff #3323 #13682) +#13678 := (implies #12373 #13673) +#13683 := (iff #13678 #13682) +#13684 := [rewrite]: #13683 +#13679 := (iff #3323 #13678) +#13676 := (iff #3322 #13673) +#13669 := (implies #3022 #13664) +#13674 := (iff #13669 #13673) +#13675 := [rewrite]: #13674 +#13670 := (iff #3322 #13669) +#13667 := (iff #3321 #13664) +#13660 := (implies #3025 #13655) +#13665 := (iff #13660 #13664) +#13666 := [rewrite]: #13665 +#13661 := (iff #3321 #13660) +#13658 := (iff #3320 #13655) +#13651 := (implies #3028 #13646) +#13656 := (iff #13651 #13655) +#13657 := [rewrite]: #13656 +#13652 := (iff #3320 #13651) +#13649 := (iff #3319 #13646) +#13642 := (implies #3031 #13637) +#13647 := (iff #13642 #13646) +#13648 := [rewrite]: #13647 +#13643 := (iff #3319 #13642) +#13640 := (iff #3318 #13637) +#13633 := (implies #12383 #13628) +#13638 := (iff #13633 #13637) +#13639 := [rewrite]: #13638 +#13634 := (iff #3318 #13633) +#13631 := (iff #3317 #13628) +#13625 := (and #13620 #3037) +#13629 := (iff #13625 #13628) +#13630 := [rewrite]: #13629 +#13626 := (iff #3317 #13625) +#13623 := (iff #3316 #13620) +#13616 := (implies #3037 #13611) +#13621 := (iff #13616 #13620) +#13622 := [rewrite]: #13621 +#13617 := (iff #3316 #13616) +#13614 := (iff #3315 #13611) +#13608 := (and #13603 #12398) +#13612 := (iff #13608 #13611) +#13613 := [rewrite]: #13612 +#13609 := (iff #3315 #13608) +#12399 := (iff #3044 #12398) +#12396 := (iff #3043 #12393) +#12390 := (implies #5718 #12387) +#12394 := (iff #12390 #12393) +#12395 := [rewrite]: #12394 +#12391 := (iff #3043 #12390) +#12388 := (iff #3042 #12387) +#12389 := [rewrite]: #12388 +#5719 := (iff #645 #5718) +#5720 := [rewrite]: #5719 +#12392 := [monotonicity #5720 #12389]: #12391 +#12397 := [trans #12392 #12395]: #12396 +#12400 := [quant-intro #12397]: #12399 +#13606 := (iff #3314 #13603) +#13599 := (implies #12398 #13594) +#13604 := (iff #13599 #13603) +#13605 := [rewrite]: #13604 +#13600 := (iff #3314 #13599) +#13597 := (iff #3313 #13594) +#13591 := (and #13586 #12406) +#13595 := (iff #13591 #13594) +#13596 := [rewrite]: #13595 +#13592 := (iff #3313 #13591) +#12409 := (iff #3046 #12406) +#12403 := (and #3022 #2951) +#12407 := (iff #12403 #12406) +#12408 := [rewrite]: #12407 +#12404 := (iff #3046 #12403) +#12401 := (iff #3045 #3022) +#12402 := [rewrite]: #12401 +#12405 := [monotonicity #12402]: #12404 +#12410 := [trans #12405 #12408]: #12409 +#13589 := (iff #3312 #13586) +#13582 := (implies #12406 #13570) +#13587 := (iff #13582 #13586) +#13588 := [rewrite]: #13587 +#13583 := (iff #3312 #13582) +#13580 := (iff #3311 #13570) +#13575 := (implies true #13570) +#13578 := (iff #13575 #13570) +#13579 := [rewrite]: #13578 +#13576 := (iff #3311 #13575) +#13573 := (iff #3310 #13570) +#13566 := (implies #3050 #13561) +#13571 := (iff #13566 #13570) +#13572 := [rewrite]: #13571 +#13567 := (iff #3310 #13566) +#13564 := (iff #3309 #13561) +#13557 := (implies #3054 #13552) +#13562 := (iff #13557 #13561) +#13563 := [rewrite]: #13562 +#13558 := (iff #3309 #13557) +#13555 := (iff #3308 #13552) +#13548 := (implies #3058 #13543) +#13553 := (iff #13548 #13552) +#13554 := [rewrite]: #13553 +#13549 := (iff #3308 #13548) +#13546 := (iff #3307 #13543) +#13540 := (implies #3060 #13535) +#13544 := (iff #13540 #13543) +#13545 := [rewrite]: #13544 +#13541 := (iff #3307 #13540) +#13538 := (iff #3306 #13535) +#13531 := (implies #3061 #13526) +#13536 := (iff #13531 #13535) +#13537 := [rewrite]: #13536 +#13532 := (iff #3306 #13531) +#13529 := (iff #3305 #13526) +#13522 := (implies #12423 #13517) +#13527 := (iff #13522 #13526) +#13528 := [rewrite]: #13527 +#13523 := (iff #3305 #13522) +#13520 := (iff #3304 #13517) +#13513 := (implies #12432 #13508) +#13518 := (iff #13513 #13517) +#13519 := [rewrite]: #13518 +#13514 := (iff #3304 #13513) +#13511 := (iff #3303 #13508) +#13505 := (implies #3060 #13432) +#13509 := (iff #13505 #13508) +#13510 := [rewrite]: #13509 +#13506 := (iff #3303 #13505) +#13503 := (iff #3302 #13432) +#13498 := (and #13432 true) +#13501 := (iff #13498 #13432) +#13502 := [rewrite]: #13501 +#13499 := (iff #3302 #13498) +#13496 := (iff #3301 true) +#13491 := (implies true true) +#13494 := (iff #13491 true) +#13495 := [rewrite]: #13494 +#13492 := (iff #3301 #13491) +#13489 := (iff #3300 true) +#13484 := (implies #3060 true) +#13487 := (iff #13484 true) +#13488 := [rewrite]: #13487 +#13485 := (iff #3300 #13484) +#13482 := (iff #3299 true) +#13449 := (or #13376 #12751) +#13457 := (or #12687 #13449) +#13472 := (or #12687 #13457) +#13477 := (implies false #13472) +#13480 := (iff #13477 true) +#13481 := [rewrite]: #13480 +#13478 := (iff #3299 #13477) +#13475 := (iff #3298 #13472) +#13469 := (implies #3060 #13457) +#13473 := (iff #13469 #13472) +#13474 := [rewrite]: #13473 +#13470 := (iff #3298 #13469) +#13467 := (iff #3297 #13457) +#13462 := (implies true #13457) +#13465 := (iff #13462 #13457) +#13466 := [rewrite]: #13465 +#13463 := (iff #3297 #13462) +#13460 := (iff #3296 #13457) +#13454 := (implies #3060 #13449) +#13458 := (iff #13454 #13457) +#13459 := [rewrite]: #13458 +#13455 := (iff #3296 #13454) +#13452 := (iff #3295 #13449) +#13446 := (implies #12556 #12751) +#13450 := (iff #13446 #13449) +#13451 := [rewrite]: #13450 +#13447 := (iff #3295 #13446) +#12754 := (iff #3167 #12751) +#12748 := (implies #3060 #12736) +#12752 := (iff #12748 #12751) +#12753 := [rewrite]: #12752 +#12749 := (iff #3167 #12748) +#12746 := (iff #3166 #12736) +#12741 := (implies true #12736) +#12744 := (iff #12741 #12736) +#12745 := [rewrite]: #12744 +#12742 := (iff #3166 #12741) +#12739 := (iff #3165 #12736) +#12733 := (implies #3060 #12728) +#12737 := (iff #12733 #12736) +#12738 := [rewrite]: #12737 +#12734 := (iff #3165 #12733) +#12731 := (iff #3164 #12728) +#12725 := (and #12720 up_216) +#12729 := (iff #12725 #12728) +#12730 := [rewrite]: #12729 +#12726 := (iff #3164 #12725) +#12723 := (iff #3163 #12720) +#12716 := (implies up_216 #12711) +#12721 := (iff #12716 #12720) +#12722 := [rewrite]: #12721 +#12717 := (iff #3163 #12716) +#12714 := (iff #3162 #12711) +#12708 := (implies #3060 #12696) +#12712 := (iff #12708 #12711) +#12713 := [rewrite]: #12712 +#12709 := (iff #3162 #12708) +#12706 := (iff #3161 #12696) +#12701 := (implies true #12696) +#12704 := (iff #12701 #12696) +#12705 := [rewrite]: #12704 +#12702 := (iff #3161 #12701) +#12699 := (iff #3160 #12696) +#12693 := (implies #3060 #12688) +#12697 := (iff #12693 #12696) +#12698 := [rewrite]: #12697 +#12694 := (iff #3160 #12693) +#12691 := (iff #3159 #12688) +#12684 := (implies #3060 #12672) +#12689 := (iff #12684 #12688) +#12690 := [rewrite]: #12689 +#12685 := (iff #3159 #12684) +#12682 := (iff #3158 #12672) +#12677 := (implies true #12672) +#12680 := (iff #12677 #12672) +#12681 := [rewrite]: #12680 +#12678 := (iff #3158 #12677) +#12675 := (iff #3157 #12672) +#12668 := (implies #12567 #12663) +#12673 := (iff #12668 #12672) +#12674 := [rewrite]: #12673 +#12669 := (iff #3157 #12668) +#12666 := (iff #3156 #12663) +#12659 := (implies #12570 #12654) +#12664 := (iff #12659 #12663) +#12665 := [rewrite]: #12664 +#12660 := (iff #3156 #12659) +#12657 := (iff #3155 #12654) +#12650 := (implies #12573 #12645) +#12655 := (iff #12650 #12654) +#12656 := [rewrite]: #12655 +#12651 := (iff #3155 #12650) +#12648 := (iff #3154 #12645) +#12641 := (implies #12576 #12629) +#12646 := (iff #12641 #12645) +#12647 := [rewrite]: #12646 +#12642 := (iff #3154 #12641) +#12639 := (iff #3153 #12629) +#12634 := (implies true #12629) +#12637 := (iff #12634 #12629) +#12638 := [rewrite]: #12637 +#12635 := (iff #3153 #12634) +#12632 := (iff #3152 #12629) +#12626 := (and #12621 #12591) +#12630 := (iff #12626 #12629) +#12631 := [rewrite]: #12630 +#12627 := (iff #3152 #12626) +#12592 := (iff #3144 #12591) +#12589 := (iff #3143 #12586) +#12583 := (implies #5718 #12580) +#12587 := (iff #12583 #12586) +#12588 := [rewrite]: #12587 +#12584 := (iff #3143 #12583) +#12581 := (iff #3142 #12580) +#12582 := [rewrite]: #12581 +#12585 := [monotonicity #5720 #12582]: #12584 +#12590 := [trans #12585 #12588]: #12589 +#12593 := [quant-intro #12590]: #12592 +#12624 := (iff #3151 #12621) +#12617 := (implies #12591 #12600) +#12622 := (iff #12617 #12621) +#12623 := [rewrite]: #12622 +#12618 := (iff #3151 #12617) +#12615 := (iff #3150 #12600) +#12610 := (and true #12600) +#12613 := (iff #12610 #12600) +#12614 := [rewrite]: #12613 +#12611 := (iff #3150 #12610) +#12601 := (iff #3148 #12600) +#12598 := (iff #3147 #12597) +#12595 := (iff #3146 #12594) +#12596 := [monotonicity #5720]: #12595 +#12599 := [monotonicity #12596]: #12598 +#12602 := [quant-intro #12599]: #12601 +#12608 := (iff #3149 true) +#12603 := (implies #12600 true) +#12606 := (iff #12603 true) +#12607 := [rewrite]: #12606 +#12604 := (iff #3149 #12603) +#12605 := [monotonicity #12602]: #12604 +#12609 := [trans #12605 #12607]: #12608 +#12612 := [monotonicity #12609 #12602]: #12611 +#12616 := [trans #12612 #12614]: #12615 +#12619 := [monotonicity #12593 #12616]: #12618 +#12625 := [trans #12619 #12623]: #12624 +#12628 := [monotonicity #12625 #12593]: #12627 +#12633 := [trans #12628 #12631]: #12632 +#12636 := [monotonicity #12633]: #12635 +#12640 := [trans #12636 #12638]: #12639 +#12577 := (iff #3139 #12576) +#12578 := [rewrite]: #12577 +#12643 := [monotonicity #12578 #12640]: #12642 +#12649 := [trans #12643 #12647]: #12648 +#12574 := (iff #3137 #12573) +#12575 := [rewrite]: #12574 +#12652 := [monotonicity #12575 #12649]: #12651 +#12658 := [trans #12652 #12656]: #12657 +#12571 := (iff #3135 #12570) +#12572 := [rewrite]: #12571 +#12661 := [monotonicity #12572 #12658]: #12660 +#12667 := [trans #12661 #12665]: #12666 +#12568 := (iff #3133 #12567) +#12569 := [rewrite]: #12568 +#12670 := [monotonicity #12569 #12667]: #12669 +#12676 := [trans #12670 #12674]: #12675 +#12679 := [monotonicity #12676]: #12678 +#12683 := [trans #12679 #12681]: #12682 +#12686 := [monotonicity #12683]: #12685 +#12692 := [trans #12686 #12690]: #12691 +#12695 := [monotonicity #12692]: #12694 +#12700 := [trans #12695 #12698]: #12699 +#12703 := [monotonicity #12700]: #12702 +#12707 := [trans #12703 #12705]: #12706 +#12710 := [monotonicity #12707]: #12709 +#12715 := [trans #12710 #12713]: #12714 +#12718 := [monotonicity #12715]: #12717 +#12724 := [trans #12718 #12722]: #12723 +#12727 := [monotonicity #12724]: #12726 +#12732 := [trans #12727 #12730]: #12731 +#12735 := [monotonicity #12732]: #12734 +#12740 := [trans #12735 #12738]: #12739 +#12743 := [monotonicity #12740]: #12742 +#12747 := [trans #12743 #12745]: #12746 +#12750 := [monotonicity #12747]: #12749 +#12755 := [trans #12750 #12753]: #12754 +#12557 := (iff #3120 #12556) +#12554 := (iff #3119 #12553) +#12555 := [rewrite]: #12554 +#12333 := (iff #2979 #12332) +#12334 := [rewrite]: #12333 +#12558 := [monotonicity #12334 #12555]: #12557 +#13448 := [monotonicity #12558 #12755]: #13447 +#13453 := [trans #13448 #13451]: #13452 +#13456 := [monotonicity #13453]: #13455 +#13461 := [trans #13456 #13459]: #13460 +#13464 := [monotonicity #13461]: #13463 +#13468 := [trans #13464 #13466]: #13467 +#13471 := [monotonicity #13468]: #13470 +#13476 := [trans #13471 #13474]: #13475 +#13479 := [monotonicity #13445 #13476]: #13478 +#13483 := [trans #13479 #13481]: #13482 +#13486 := [monotonicity #13483]: #13485 +#13490 := [trans #13486 #13488]: #13489 +#13493 := [monotonicity #13490]: #13492 +#13497 := [trans #13493 #13495]: #13496 +#13442 := (iff #3293 #13432) +#13437 := (implies true #13432) +#13440 := (iff #13437 #13432) +#13441 := [rewrite]: #13440 +#13438 := (iff #3293 #13437) +#13435 := (iff #3292 #13432) +#13429 := (implies #3060 #13417) +#13433 := (iff #13429 #13432) +#13434 := [rewrite]: #13433 +#13430 := (iff #3292 #13429) +#13427 := (iff #3291 #13417) +#13422 := (implies true #13417) +#13425 := (iff #13422 #13417) +#13426 := [rewrite]: #13425 +#13423 := (iff #3291 #13422) +#13420 := (iff #3290 #13417) +#13414 := (implies #3060 #13402) +#13418 := (iff #13414 #13417) +#13419 := [rewrite]: #13418 +#13415 := (iff #3290 #13414) +#13412 := (iff #3289 #13402) +#13407 := (implies true #13402) +#13410 := (iff #13407 #13402) +#13411 := [rewrite]: #13410 +#13408 := (iff #3289 #13407) +#13405 := (iff #3288 #13402) +#13399 := (implies #3060 #13394) +#13403 := (iff #13399 #13402) +#13404 := [rewrite]: #13403 +#13400 := (iff #3288 #13399) +#13397 := (iff #3287 #13394) +#13391 := (implies #12544 #13386) +#13395 := (iff #13391 #13394) +#13396 := [rewrite]: #13395 +#13392 := (iff #3287 #13391) +#13389 := (iff #3286 #13386) +#13382 := (implies #12544 #13377) +#13387 := (iff #13382 #13386) +#13388 := [rewrite]: #13387 +#13383 := (iff #3286 #13382) +#13380 := (iff #3285 #13377) +#13373 := (implies #12556 #13368) +#13378 := (iff #13373 #13377) +#13379 := [rewrite]: #13378 +#13374 := (iff #3285 #13373) +#13371 := (iff #3284 #13368) +#13364 := (implies #3121 #13359) +#13369 := (iff #13364 #13368) +#13370 := [rewrite]: #13369 +#13365 := (iff #3284 #13364) +#13362 := (iff #3283 #13359) +#13355 := (implies #3122 #13350) +#13360 := (iff #13355 #13359) +#13361 := [rewrite]: #13360 +#13356 := (iff #3283 #13355) +#13353 := (iff #3282 #13350) +#13346 := (implies #3123 #13341) +#13351 := (iff #13346 #13350) +#13352 := [rewrite]: #13351 +#13347 := (iff #3282 #13346) +#13344 := (iff #3281 #13341) +#13337 := (implies #3124 #13332) +#13342 := (iff #13337 #13341) +#13343 := [rewrite]: #13342 +#13338 := (iff #3281 #13337) +#13335 := (iff #3280 #13332) +#13328 := (implies #3127 #13316) +#13333 := (iff #13328 #13332) +#13334 := [rewrite]: #13333 +#13329 := (iff #3280 #13328) +#13326 := (iff #3279 #13316) +#13321 := (implies true #13316) +#13324 := (iff #13321 #13316) +#13325 := [rewrite]: #13324 +#13322 := (iff #3279 #13321) +#13319 := (iff #3278 #13316) +#13313 := (implies #3060 #13310) +#13317 := (iff #13313 #13316) +#13318 := [rewrite]: #13317 +#13314 := (iff #3278 #13313) +#13311 := (iff #3277 #13310) +#13308 := (iff #3276 #13298) +#13303 := (implies true #13298) +#13306 := (iff #13303 #13298) +#13307 := [rewrite]: #13306 +#13304 := (iff #3276 #13303) +#13301 := (iff #3275 #13298) +#13295 := (implies #3060 #13290) +#13299 := (iff #13295 #13298) +#13300 := [rewrite]: #13299 +#13296 := (iff #3275 #13295) +#13293 := (iff #3274 #13290) +#13286 := (implies #3174 #13281) +#13291 := (iff #13286 #13290) +#13292 := [rewrite]: #13291 +#13287 := (iff #3274 #13286) +#13284 := (iff #3273 #13281) +#13278 := (implies #3060 #13266) +#13282 := (iff #13278 #13281) +#13283 := [rewrite]: #13282 +#13279 := (iff #3273 #13278) +#13276 := (iff #3272 #13266) +#13271 := (implies true #13266) +#13274 := (iff #13271 #13266) +#13275 := [rewrite]: #13274 +#13272 := (iff #3272 #13271) +#13269 := (iff #3271 #13266) +#13263 := (implies #3060 #13258) +#13267 := (iff #13263 #13266) +#13268 := [rewrite]: #13267 +#13264 := (iff #3271 #13263) +#13261 := (iff #3270 #13258) +#13255 := (and #13250 #12809) +#13259 := (iff #13255 #13258) +#13260 := [rewrite]: #13259 +#13256 := (iff #3270 #13255) +#12810 := (iff #3180 #12809) +#12807 := (iff #3179 #12806) +#12808 := [rewrite]: #12807 +#12804 := (iff #3177 #12803) +#12805 := [rewrite]: #12804 +#12811 := [monotonicity #12805 #12808]: #12810 +#13253 := (iff #3269 #13250) +#13247 := (implies #12809 #13242) +#13251 := (iff #13247 #13250) +#13252 := [rewrite]: #13251 +#13248 := (iff #3269 #13247) +#13245 := (iff #3268 #13242) +#13239 := (and #13234 #12818) +#13243 := (iff #13239 #13242) +#13244 := [rewrite]: #13243 +#13240 := (iff #3268 #13239) +#12821 := (iff #3183 #12818) +#12815 := (and #12812 #12806) +#12819 := (iff #12815 #12818) +#12820 := [rewrite]: #12819 +#12816 := (iff #3183 #12815) +#12813 := (iff #3182 #12812) +#12814 := [rewrite]: #12813 +#12817 := [monotonicity #12814 #12808]: #12816 +#12822 := [trans #12817 #12820]: #12821 +#13237 := (iff #3267 #13234) +#13231 := (implies #12818 #13226) +#13235 := (iff #13231 #13234) +#13236 := [rewrite]: #13235 +#13232 := (iff #3267 #13231) +#13229 := (iff #3266 #13226) +#13223 := (implies #3060 #13220) +#13227 := (iff #13223 #13226) +#13228 := [rewrite]: #13227 +#13224 := (iff #3266 #13223) +#13221 := (iff #3265 #13220) +#13218 := (iff #3264 #13208) +#13213 := (implies true #13208) +#13216 := (iff #13213 #13208) +#13217 := [rewrite]: #13216 +#13214 := (iff #3264 #13213) +#13211 := (iff #3263 #13208) +#13205 := (implies #3060 #13200) +#13209 := (iff #13205 #13208) +#13210 := [rewrite]: #13209 +#13206 := (iff #3263 #13205) +#13203 := (iff #3262 #13200) +#13196 := (implies #3238 #13191) +#13201 := (iff #13196 #13200) +#13202 := [rewrite]: #13201 +#13197 := (iff #3262 #13196) +#13194 := (iff #3261 #13191) +#13188 := (implies #3060 #13176) +#13192 := (iff #13188 #13191) +#13193 := [rewrite]: #13192 +#13189 := (iff #3261 #13188) +#13186 := (iff #3260 #13176) +#13181 := (implies true #13176) +#13184 := (iff #13181 #13176) +#13185 := [rewrite]: #13184 +#13182 := (iff #3260 #13181) +#13179 := (iff #3259 #13176) +#13173 := (implies #3060 #13168) +#13177 := (iff #13173 #13176) +#13178 := [rewrite]: #13177 +#13174 := (iff #3259 #13173) +#13171 := (iff #3258 #13168) +#13165 := (and #13160 #12809) +#13169 := (iff #13165 #13168) +#13170 := [rewrite]: #13169 +#13166 := (iff #3258 #13165) +#13163 := (iff #3257 #13160) +#13156 := (implies #12809 #13151) +#13161 := (iff #13156 #13160) +#13162 := [rewrite]: #13161 +#13157 := (iff #3257 #13156) +#13154 := (iff #3256 #13151) +#13148 := (and #13143 #12818) +#13152 := (iff #13148 #13151) +#13153 := [rewrite]: #13152 +#13149 := (iff #3256 #13148) +#13146 := (iff #3255 #13143) +#13139 := (implies #12818 #13134) +#13144 := (iff #13139 #13143) +#13145 := [rewrite]: #13144 +#13140 := (iff #3255 #13139) +#13137 := (iff #3254 #13134) +#13130 := (implies #13070 #13125) +#13135 := (iff #13130 #13134) +#13136 := [rewrite]: #13135 +#13131 := (iff #3254 #13130) +#13128 := (iff #3253 #13125) +#13121 := (implies #3242 #13116) +#13126 := (iff #13121 #13125) +#13127 := [rewrite]: #13126 +#13122 := (iff #3253 #13121) +#13119 := (iff #3252 #13116) +#13112 := (implies #3244 #13107) +#13117 := (iff #13112 #13116) +#13118 := [rewrite]: #13117 +#13113 := (iff #3252 #13112) +#13110 := (iff #3251 #13107) +#13103 := (implies #3059 #13091) +#13108 := (iff #13103 #13107) +#13109 := [rewrite]: #13108 +#13104 := (iff #3251 #13103) +#13101 := (iff #3250 #13091) +#13096 := (implies true #13091) +#13099 := (iff #13096 #13091) +#13100 := [rewrite]: #13099 +#13097 := (iff #3250 #13096) +#13094 := (iff #3249 #13091) +#13087 := (implies #3246 #13082) +#13092 := (iff #13087 #13091) +#13093 := [rewrite]: #13092 +#13088 := (iff #3249 #13087) +#13085 := (iff #3248 #13082) +#13078 := (implies #13075 #12978) +#13083 := (iff #13078 #13082) +#13084 := [rewrite]: #13083 +#13079 := (iff #3248 #13078) +#12988 := (iff #3227 #12978) +#12983 := (implies true #12978) +#12986 := (iff #12983 #12978) +#12987 := [rewrite]: #12986 +#12984 := (iff #3227 #12983) +#12981 := (iff #3226 #12978) +#12974 := (implies #12829 #12969) +#12979 := (iff #12974 #12978) +#12980 := [rewrite]: #12979 +#12975 := (iff #3226 #12974) +#12972 := (iff #3225 #12969) +#12966 := (and #12961 #12841) +#12970 := (iff #12966 #12969) +#12971 := [rewrite]: #12970 +#12967 := (iff #3225 #12966) +#12842 := (iff #3195 #12841) +#12839 := (iff #3194 #12838) +#12833 := (= #3192 #12832) +#12834 := [rewrite]: #12833 +#12840 := [monotonicity #12834]: #12839 +#12836 := (iff #3193 #12835) +#12837 := [monotonicity #12834]: #12836 +#12843 := [monotonicity #12837 #12840]: #12842 +#12964 := (iff #3224 #12961) +#12957 := (implies #12841 #12952) +#12962 := (iff #12957 #12961) +#12963 := [rewrite]: #12962 +#12958 := (iff #3224 #12957) +#12955 := (iff #3223 #12952) +#12948 := (implies #12844 #12943) +#12953 := (iff #12948 #12952) +#12954 := [rewrite]: #12953 +#12949 := (iff #3223 #12948) +#12946 := (iff #3222 #12943) +#12939 := (implies #3199 #12934) +#12944 := (iff #12939 #12943) +#12945 := [rewrite]: #12944 +#12940 := (iff #3222 #12939) +#12937 := (iff #3221 #12934) +#12930 := (implies #3201 #12918) +#12935 := (iff #12930 #12934) +#12936 := [rewrite]: #12935 +#12931 := (iff #3221 #12930) +#12928 := (iff #3220 #12918) +#12923 := (implies true #12918) +#12926 := (iff #12923 #12918) +#12927 := [rewrite]: #12926 +#12924 := (iff #3220 #12923) +#12921 := (iff #3219 #12918) +#12915 := (and #12910 #3202) +#12919 := (iff #12915 #12918) +#12920 := [rewrite]: #12919 +#12916 := (iff #3219 #12915) +#12913 := (iff #3218 #12910) +#12906 := (implies #3202 #12901) +#12911 := (iff #12906 #12910) +#12912 := [rewrite]: #12911 +#12907 := (iff #3218 #12906) +#12904 := (iff #3217 #12901) +#12898 := (and #12893 #12859) +#12902 := (iff #12898 #12901) +#12903 := [rewrite]: #12902 +#12899 := (iff #3217 #12898) +#12860 := (iff #3207 #12859) +#12857 := (iff #3206 #12854) +#12851 := (implies #5718 #12848) +#12855 := (iff #12851 #12854) +#12856 := [rewrite]: #12855 +#12852 := (iff #3206 #12851) +#12849 := (iff #3205 #12848) +#12850 := [rewrite]: #12849 +#12853 := [monotonicity #5720 #12850]: #12852 +#12858 := [trans #12853 #12856]: #12857 +#12861 := [quant-intro #12858]: #12860 +#12896 := (iff #3216 #12893) +#12889 := (implies #12859 #12868) +#12894 := (iff #12889 #12893) +#12895 := [rewrite]: #12894 +#12890 := (iff #3216 #12889) +#12887 := (iff #3215 #12868) +#12882 := (and true #12868) +#12885 := (iff #12882 #12868) +#12886 := [rewrite]: #12885 +#12883 := (iff #3215 #12882) +#12871 := (iff #3212 #12868) +#12865 := (and #12862 #3211) +#12869 := (iff #12865 #12868) +#12870 := [rewrite]: #12869 +#12866 := (iff #3212 #12865) +#12863 := (iff #3210 #12862) +#12864 := [rewrite]: #12863 +#12867 := [monotonicity #12864]: #12866 +#12872 := [trans #12867 #12870]: #12871 +#12880 := (iff #3214 true) +#12875 := (implies #12868 true) +#12878 := (iff #12875 true) +#12879 := [rewrite]: #12878 +#12876 := (iff #3214 #12875) +#12873 := (iff #3213 true) +#12874 := [rewrite]: #12873 +#12877 := [monotonicity #12872 #12874]: #12876 +#12881 := [trans #12877 #12879]: #12880 +#12884 := [monotonicity #12881 #12872]: #12883 +#12888 := [trans #12884 #12886]: #12887 +#12891 := [monotonicity #12861 #12888]: #12890 +#12897 := [trans #12891 #12895]: #12896 +#12900 := [monotonicity #12897 #12861]: #12899 +#12905 := [trans #12900 #12903]: #12904 +#12908 := [monotonicity #12905]: #12907 +#12914 := [trans #12908 #12912]: #12913 +#12917 := [monotonicity #12914]: #12916 +#12922 := [trans #12917 #12920]: #12921 +#12925 := [monotonicity #12922]: #12924 +#12929 := [trans #12925 #12927]: #12928 +#12932 := [monotonicity #12929]: #12931 +#12938 := [trans #12932 #12936]: #12937 +#12941 := [monotonicity #12938]: #12940 +#12947 := [trans #12941 #12945]: #12946 +#12845 := (iff #3197 #12844) +#12846 := [monotonicity #12834]: #12845 +#12950 := [monotonicity #12846 #12947]: #12949 +#12956 := [trans #12950 #12954]: #12955 +#12959 := [monotonicity #12843 #12956]: #12958 +#12965 := [trans #12959 #12963]: #12964 +#12968 := [monotonicity #12965 #12843]: #12967 +#12973 := [trans #12968 #12971]: #12972 +#12830 := (iff #3191 #12829) +#12831 := [rewrite]: #12830 +#12976 := [monotonicity #12831 #12973]: #12975 +#12982 := [trans #12976 #12980]: #12981 +#12985 := [monotonicity #12982]: #12984 +#12989 := [trans #12985 #12987]: #12988 +#13076 := (iff #3247 #13075) +#13077 := [rewrite]: #13076 +#13080 := [monotonicity #13077 #12989]: #13079 +#13086 := [trans #13080 #13084]: #13085 +#13089 := [monotonicity #13086]: #13088 +#13095 := [trans #13089 #13093]: #13094 +#13098 := [monotonicity #13095]: #13097 +#13102 := [trans #13098 #13100]: #13101 +#13073 := (iff #3245 #3059) +#13074 := [rewrite]: #13073 +#13105 := [monotonicity #13074 #13102]: #13104 +#13111 := [trans #13105 #13109]: #13110 +#13114 := [monotonicity #13111]: #13113 +#13120 := [trans #13114 #13118]: #13119 +#13123 := [monotonicity #13120]: #13122 +#13129 := [trans #13123 #13127]: #13128 +#13071 := (iff #3240 #13070) +#13072 := [rewrite]: #13071 +#13132 := [monotonicity #13072 #13129]: #13131 +#13138 := [trans #13132 #13136]: #13137 +#13141 := [monotonicity #12822 #13138]: #13140 +#13147 := [trans #13141 #13145]: #13146 +#13150 := [monotonicity #13147 #12822]: #13149 +#13155 := [trans #13150 #13153]: #13154 +#13158 := [monotonicity #12811 #13155]: #13157 +#13164 := [trans #13158 #13162]: #13163 +#13167 := [monotonicity #13164 #12811]: #13166 +#13172 := [trans #13167 #13170]: #13171 +#13175 := [monotonicity #13172]: #13174 +#13180 := [trans #13175 #13178]: #13179 +#13183 := [monotonicity #13180]: #13182 +#13187 := [trans #13183 #13185]: #13186 +#13190 := [monotonicity #13187]: #13189 +#13195 := [trans #13190 #13193]: #13194 +#13198 := [monotonicity #13195]: #13197 +#13204 := [trans #13198 #13202]: #13203 +#13207 := [monotonicity #13204]: #13206 +#13212 := [trans #13207 #13210]: #13211 +#13215 := [monotonicity #13212]: #13214 +#13219 := [trans #13215 #13217]: #13218 +#13068 := (iff #3237 #13058) +#13063 := (implies true #13058) +#13066 := (iff #13063 #13058) +#13067 := [rewrite]: #13066 +#13064 := (iff #3237 #13063) +#13061 := (iff #3236 #13058) +#13055 := (implies #3060 #13050) +#13059 := (iff #13055 #13058) +#13060 := [rewrite]: #13059 +#13056 := (iff #3236 #13055) +#13053 := (iff #3235 #13050) +#13046 := (implies #3185 #13041) +#13051 := (iff #13046 #13050) +#13052 := [rewrite]: #13051 +#13047 := (iff #3235 #13046) +#13044 := (iff #3234 #13041) +#13038 := (implies #3060 #13026) +#13042 := (iff #13038 #13041) +#13043 := [rewrite]: #13042 +#13039 := (iff #3234 #13038) +#13036 := (iff #3233 #13026) +#13031 := (implies true #13026) +#13034 := (iff #13031 #13026) +#13035 := [rewrite]: #13034 +#13032 := (iff #3233 #13031) +#13029 := (iff #3232 #13026) +#13023 := (implies #3060 #13018) +#13027 := (iff #13023 #13026) +#13028 := [rewrite]: #13027 +#13024 := (iff #3232 #13023) +#13021 := (iff #3231 #13018) +#13015 := (implies #3060 #13003) +#13019 := (iff #13015 #13018) +#13020 := [rewrite]: #13019 +#13016 := (iff #3231 #13015) +#13013 := (iff #3230 #13003) +#13008 := (implies true #13003) +#13011 := (iff #13008 #13003) +#13012 := [rewrite]: #13011 +#13009 := (iff #3230 #13008) +#13006 := (iff #3229 #13003) +#12999 := (implies #12823 #12994) +#13004 := (iff #12999 #13003) +#13005 := [rewrite]: #13004 +#13000 := (iff #3229 #12999) +#12997 := (iff #3228 #12994) +#12990 := (implies #12826 #12978) +#12995 := (iff #12990 #12994) +#12996 := [rewrite]: #12995 +#12991 := (iff #3228 #12990) +#12827 := (iff #3189 #12826) +#12828 := [rewrite]: #12827 +#12992 := [monotonicity #12828 #12989]: #12991 +#12998 := [trans #12992 #12996]: #12997 +#12824 := (iff #3187 #12823) +#12825 := [rewrite]: #12824 +#13001 := [monotonicity #12825 #12998]: #13000 +#13007 := [trans #13001 #13005]: #13006 +#13010 := [monotonicity #13007]: #13009 +#13014 := [trans #13010 #13012]: #13013 +#13017 := [monotonicity #13014]: #13016 +#13022 := [trans #13017 #13020]: #13021 +#13025 := [monotonicity #13022]: #13024 +#13030 := [trans #13025 #13028]: #13029 +#13033 := [monotonicity #13030]: #13032 +#13037 := [trans #13033 #13035]: #13036 +#13040 := [monotonicity #13037]: #13039 +#13045 := [trans #13040 #13043]: #13044 +#13048 := [monotonicity #13045]: #13047 +#13054 := [trans #13048 #13052]: #13053 +#13057 := [monotonicity #13054]: #13056 +#13062 := [trans #13057 #13060]: #13061 +#13065 := [monotonicity #13062]: #13064 +#13069 := [trans #13065 #13067]: #13068 +#13222 := [monotonicity #13069 #13219]: #13221 +#13225 := [monotonicity #13222]: #13224 +#13230 := [trans #13225 #13228]: #13229 +#13233 := [monotonicity #12822 #13230]: #13232 +#13238 := [trans #13233 #13236]: #13237 +#13241 := [monotonicity #13238 #12822]: #13240 +#13246 := [trans #13241 #13244]: #13245 +#13249 := [monotonicity #12811 #13246]: #13248 +#13254 := [trans #13249 #13252]: #13253 +#13257 := [monotonicity #13254 #12811]: #13256 +#13262 := [trans #13257 #13260]: #13261 +#13265 := [monotonicity #13262]: #13264 +#13270 := [trans #13265 #13268]: #13269 +#13273 := [monotonicity #13270]: #13272 +#13277 := [trans #13273 #13275]: #13276 +#13280 := [monotonicity #13277]: #13279 +#13285 := [trans #13280 #13283]: #13284 +#13288 := [monotonicity #13285]: #13287 +#13294 := [trans #13288 #13292]: #13293 +#13297 := [monotonicity #13294]: #13296 +#13302 := [trans #13297 #13300]: #13301 +#13305 := [monotonicity #13302]: #13304 +#13309 := [trans #13305 #13307]: #13308 +#12801 := (iff #3173 #12791) +#12796 := (implies true #12791) +#12799 := (iff #12796 #12791) +#12800 := [rewrite]: #12799 +#12797 := (iff #3173 #12796) +#12794 := (iff #3172 #12791) +#12788 := (implies #3060 #12783) +#12792 := (iff #12788 #12791) +#12793 := [rewrite]: #12792 +#12789 := (iff #3172 #12788) +#12786 := (iff #3171 #12783) +#12779 := (implies #3131 #12774) +#12784 := (iff #12779 #12783) +#12785 := [rewrite]: #12784 +#12780 := (iff #3171 #12779) +#12777 := (iff #3170 #12774) +#12771 := (implies #3060 #12759) +#12775 := (iff #12771 #12774) +#12776 := [rewrite]: #12775 +#12772 := (iff #3170 #12771) +#12769 := (iff #3169 #12759) +#12764 := (implies true #12759) +#12767 := (iff #12764 #12759) +#12768 := [rewrite]: #12767 +#12765 := (iff #3169 #12764) +#12762 := (iff #3168 #12759) +#12756 := (implies #3060 #12751) +#12760 := (iff #12756 #12759) +#12761 := [rewrite]: #12760 +#12757 := (iff #3168 #12756) +#12758 := [monotonicity #12755]: #12757 +#12763 := [trans #12758 #12761]: #12762 +#12766 := [monotonicity #12763]: #12765 +#12770 := [trans #12766 #12768]: #12769 +#12773 := [monotonicity #12770]: #12772 +#12778 := [trans #12773 #12776]: #12777 +#12781 := [monotonicity #12778]: #12780 +#12787 := [trans #12781 #12785]: #12786 +#12790 := [monotonicity #12787]: #12789 +#12795 := [trans #12790 #12793]: #12794 +#12798 := [monotonicity #12795]: #12797 +#12802 := [trans #12798 #12800]: #12801 +#13312 := [monotonicity #12802 #13309]: #13311 +#13315 := [monotonicity #13312]: #13314 +#13320 := [trans #13315 #13318]: #13319 +#12565 := (iff #3130 true) +#12520 := (iff #12517 true) +#12521 := [rewrite]: #12520 +#12563 := (iff #3130 #12517) +#12561 := (iff #3129 true) +#12562 := [rewrite]: #12561 +#12559 := (iff #3128 true) +#12560 := [rewrite]: #12559 +#12564 := [monotonicity #12560 #12562]: #12563 +#12566 := [trans #12564 #12521]: #12565 +#13323 := [monotonicity #12566 #13320]: #13322 +#13327 := [trans #13323 #13325]: #13326 +#13330 := [monotonicity #13327]: #13329 +#13336 := [trans #13330 #13334]: #13335 +#13339 := [monotonicity #13336]: #13338 +#13345 := [trans #13339 #13343]: #13344 +#13348 := [monotonicity #13345]: #13347 +#13354 := [trans #13348 #13352]: #13353 +#13357 := [monotonicity #13354]: #13356 +#13363 := [trans #13357 #13361]: #13362 +#13366 := [monotonicity #13363]: #13365 +#13372 := [trans #13366 #13370]: #13371 +#13375 := [monotonicity #12558 #13372]: #13374 +#13381 := [trans #13375 #13379]: #13380 +#12551 := (iff #3116 #12544) +#12541 := (and #12437 #3115) +#12545 := (iff #12541 #12544) +#12546 := [rewrite]: #12545 +#12549 := (iff #3116 #12541) +#12438 := (iff #3073 #12437) +#12439 := [rewrite]: #12438 +#12550 := [monotonicity #12439]: #12549 +#12552 := [trans #12550 #12546]: #12551 +#13384 := [monotonicity #12552 #13381]: #13383 +#13390 := [trans #13384 #13388]: #13389 +#12547 := (iff #3114 #12544) +#12542 := (iff #3114 #12541) +#12539 := (iff #3113 #3115) +#12537 := (iff #3112 #3078) +#12532 := (and #3078 true) +#12535 := (iff #12532 #3078) +#12536 := [rewrite]: #12535 +#12533 := (iff #3112 #12532) +#12530 := (iff #3111 true) +#12528 := (iff #3111 #12517) +#12526 := (iff #3110 true) +#12524 := (iff #3110 #12517) +#12522 := (iff #3109 true) +#12518 := (iff #3109 #12517) +#12515 := (iff #3108 true) +#12476 := (forall (vars (?x777 T5)) (:pat #3090) true) +#12479 := (iff #12476 true) +#12480 := [elim-unused]: #12479 +#12513 := (iff #3108 #12476) +#12511 := (iff #3107 true) +#12500 := (= uf_261 #3104) +#12503 := (not #12500) +#12506 := (implies #12503 #12503) +#12509 := (iff #12506 true) +#12510 := [rewrite]: #12509 +#12507 := (iff #3107 #12506) +#12504 := (iff #3106 #12503) +#12501 := (iff #3105 #12500) +#12502 := [rewrite]: #12501 +#12505 := [monotonicity #12502]: #12504 +#12508 := [monotonicity #12505 #12505]: #12507 +#12512 := [trans #12508 #12510]: #12511 +#12514 := [quant-intro #12512]: #12513 +#12516 := [trans #12514 #12480]: #12515 +#12498 := (iff #3101 true) +#12493 := (forall (vars (?x778 T5)) (:pat #3097) true) +#12496 := (iff #12493 true) +#12497 := [elim-unused]: #12496 +#12494 := (iff #3101 #12493) +#12491 := (iff #3100 true) +#12440 := (= uf_9 #3082) +#12452 := (implies #12440 #12440) +#12455 := (iff #12452 true) +#12456 := [rewrite]: #12455 +#12489 := (iff #3100 #12452) +#12487 := (iff #3099 #12440) +#12445 := (and #12440 true) +#12448 := (iff #12445 #12440) +#12449 := [rewrite]: #12448 +#12485 := (iff #3099 #12445) +#12483 := (iff #3098 true) +#12484 := [rewrite]: #12483 +#12441 := (iff #3083 #12440) +#12442 := [rewrite]: #12441 +#12486 := [monotonicity #12442 #12484]: #12485 +#12488 := [trans #12486 #12449]: #12487 +#12490 := [monotonicity #12442 #12488]: #12489 +#12492 := [trans #12490 #12456]: #12491 +#12495 := [quant-intro #12492]: #12494 +#12499 := [trans #12495 #12497]: #12498 +#12519 := [monotonicity #12499 #12516]: #12518 +#12523 := [trans #12519 #12521]: #12522 +#12481 := (iff #3094 true) +#12477 := (iff #3094 #12476) +#12474 := (iff #3093 true) +#12472 := (iff #3093 #12452) +#12470 := (iff #3092 #12440) +#12468 := (iff #3092 #12445) +#12466 := (iff #3091 true) +#12467 := [rewrite]: #12466 +#12469 := [monotonicity #12442 #12467]: #12468 +#12471 := [trans #12469 #12449]: #12470 +#12473 := [monotonicity #12442 #12471]: #12472 +#12475 := [trans #12473 #12456]: #12474 +#12478 := [quant-intro #12475]: #12477 +#12482 := [trans #12478 #12480]: #12481 +#12525 := [monotonicity #12482 #12523]: #12524 +#12527 := [trans #12525 #12521]: #12526 +#12464 := (iff #3087 true) +#12459 := (forall (vars (?x776 T5)) (:pat #3081) true) +#12462 := (iff #12459 true) +#12463 := [elim-unused]: #12462 +#12460 := (iff #3087 #12459) +#12457 := (iff #3086 true) +#12453 := (iff #3086 #12452) +#12450 := (iff #3085 #12440) +#12446 := (iff #3085 #12445) +#12443 := (iff #3084 true) +#12444 := [rewrite]: #12443 +#12447 := [monotonicity #12442 #12444]: #12446 +#12451 := [trans #12447 #12449]: #12450 +#12454 := [monotonicity #12442 #12451]: #12453 +#12458 := [trans #12454 #12456]: #12457 +#12461 := [quant-intro #12458]: #12460 +#12465 := [trans #12461 #12463]: #12464 +#12529 := [monotonicity #12465 #12527]: #12528 +#12531 := [trans #12529 #12521]: #12530 +#12534 := [monotonicity #12531]: #12533 +#12538 := [trans #12534 #12536]: #12537 +#12540 := [monotonicity #12538]: #12539 +#12543 := [monotonicity #12439 #12540]: #12542 +#12548 := [trans #12543 #12546]: #12547 +#13393 := [monotonicity #12548 #13390]: #13392 +#13398 := [trans #13393 #13396]: #13397 +#13401 := [monotonicity #13398]: #13400 +#13406 := [trans #13401 #13404]: #13405 +#13409 := [monotonicity #13406]: #13408 +#13413 := [trans #13409 #13411]: #13412 +#13416 := [monotonicity #13413]: #13415 +#13421 := [trans #13416 #13419]: #13420 +#13424 := [monotonicity #13421]: #13423 +#13428 := [trans #13424 #13426]: #13427 +#13431 := [monotonicity #13428]: #13430 +#13436 := [trans #13431 #13434]: #13435 +#13439 := [monotonicity #13436]: #13438 +#13443 := [trans #13439 #13441]: #13442 +#13500 := [monotonicity #13443 #13497]: #13499 +#13504 := [trans #13500 #13502]: #13503 +#13507 := [monotonicity #13504]: #13506 +#13512 := [trans #13507 #13510]: #13511 +#12435 := (iff #3071 #12432) +#12429 := (and #12426 #3070) +#12433 := (iff #12429 #12432) +#12434 := [rewrite]: #12433 +#12430 := (iff #3071 #12429) +#12427 := (iff #3069 #12426) +#12428 := [rewrite]: #12427 +#12431 := [monotonicity #12428]: #12430 +#12436 := [trans #12431 #12434]: #12435 +#13515 := [monotonicity #12436 #13512]: #13514 +#13521 := [trans #13515 #13519]: #13520 +#12424 := (iff #3066 #12423) +#12421 := (iff #3065 #12418) +#12415 := (implies #5718 #12412) +#12419 := (iff #12415 #12418) +#12420 := [rewrite]: #12419 +#12416 := (iff #3065 #12415) +#12413 := (iff #3064 #12412) +#12414 := [rewrite]: #12413 +#12417 := [monotonicity #5720 #12414]: #12416 +#12422 := [trans #12417 #12420]: #12421 +#12425 := [quant-intro #12422]: #12424 +#13524 := [monotonicity #12425 #13521]: #13523 +#13530 := [trans #13524 #13528]: #13529 +#13533 := [monotonicity #13530]: #13532 +#13539 := [trans #13533 #13537]: #13538 +#13542 := [monotonicity #13539]: #13541 +#13547 := [trans #13542 #13545]: #13546 +#13550 := [monotonicity #13547]: #13549 +#13556 := [trans #13550 #13554]: #13555 +#13559 := [monotonicity #13556]: #13558 +#13565 := [trans #13559 #13563]: #13564 +#13568 := [monotonicity #13565]: #13567 +#13574 := [trans #13568 #13572]: #13573 +#13577 := [monotonicity #13574]: #13576 +#13581 := [trans #13577 #13579]: #13580 +#13584 := [monotonicity #12410 #13581]: #13583 +#13590 := [trans #13584 #13588]: #13589 +#13593 := [monotonicity #13590 #12410]: #13592 +#13598 := [trans #13593 #13596]: #13597 +#13601 := [monotonicity #12400 #13598]: #13600 +#13607 := [trans #13601 #13605]: #13606 +#13610 := [monotonicity #13607 #12400]: #13609 +#13615 := [trans #13610 #13613]: #13614 +#13618 := [monotonicity #13615]: #13617 +#13624 := [trans #13618 #13622]: #13623 +#13627 := [monotonicity #13624]: #13626 +#13632 := [trans #13627 #13630]: #13631 +#12384 := (iff #3036 #12383) +#12381 := (iff #3035 #12380) +#12378 := (iff #3034 #3033) +#12379 := [rewrite]: #12378 +#12382 := [monotonicity #12379]: #12381 +#12385 := [monotonicity #12382]: #12384 +#13635 := [monotonicity #12385 #13632]: #13634 +#13641 := [trans #13635 #13639]: #13640 +#13644 := [monotonicity #13641]: #13643 +#13650 := [trans #13644 #13648]: #13649 +#13653 := [monotonicity #13650]: #13652 +#13659 := [trans #13653 #13657]: #13658 +#13662 := [monotonicity #13659]: #13661 +#13668 := [trans #13662 #13666]: #13667 +#13671 := [monotonicity #13668]: #13670 +#13677 := [trans #13671 #13675]: #13676 +#13680 := [monotonicity #12377 #13677]: #13679 +#13686 := [trans #13680 #13684]: #13685 +#13689 := [monotonicity #13686 #12377]: #13688 +#13694 := [trans #13689 #13692]: #13693 +#13697 := [monotonicity #12366 #13694]: #13696 +#13703 := [trans #13697 #13701]: #13702 +#13706 := [monotonicity #13703 #12366]: #13705 +#13711 := [trans #13706 #13709]: #13710 +#13714 := [monotonicity #12357 #13711]: #13713 +#13720 := [trans #13714 #13718]: #13719 +#13723 := [monotonicity #13720 #12357]: #13722 +#13728 := [trans #13723 #13726]: #13727 +#13731 := [monotonicity #13728]: #13730 +#13737 := [trans #13731 #13735]: #13736 +#12353 := (iff #3004 #12352) +#12350 := (iff #3003 #12347) +#12344 := (iff #12341 false) +#12348 := (iff #12344 #12347) +#12349 := [rewrite]: #12348 +#12345 := (iff #3003 #12344) +#12342 := (iff #3002 #12341) +#12343 := [rewrite]: #12342 +#12346 := [monotonicity #12343]: #12345 +#12351 := [trans #12346 #12349]: #12350 +#12354 := [quant-intro #12351]: #12353 +#13740 := [monotonicity #12354 #13737]: #13739 +#13746 := [trans #13740 #13744]: #13745 +#13749 := [monotonicity #13746]: #13748 +#13755 := [trans #13749 #13753]: #13754 +#13758 := [monotonicity #13755]: #13757 +#13764 := [trans #13758 #13762]: #13763 +#13767 := [monotonicity #13764]: #13766 +#13773 := [trans #13767 #13771]: #13772 +#13776 := [monotonicity #13773]: #13775 +#13782 := [trans #13776 #13780]: #13781 +#12339 := (iff #2983 #12338) +#12336 := (iff #2982 #12335) +#12337 := [rewrite]: #12336 +#12340 := [monotonicity #12334 #12337]: #12339 +#13785 := [monotonicity #12340 #13782]: #13784 +#13791 := [trans #13785 #13789]: #13790 +#12330 := (iff #2977 #12329) +#12331 := [rewrite]: #12330 +#13794 := [monotonicity #12331 #13791]: #13793 +#13800 := [trans #13794 #13798]: #13799 +#13803 := [monotonicity #13800]: #13802 +#13807 := [trans #13803 #13805]: #13806 +#12327 := (iff #2975 #12326) +#12324 := (iff #2974 #12323) +#12321 := (iff #2973 #12320) +#12318 := (iff #2972 #12317) +#12315 := (iff #2971 #12314) +#12312 := (iff #2970 #12311) +#12313 := [rewrite]: #12312 +#12309 := (iff #2968 #12308) +#12310 := [rewrite]: #12309 +#12316 := [monotonicity #12310 #12313]: #12315 +#12306 := (iff #2966 #12305) +#12307 := [rewrite]: #12306 +#12319 := [monotonicity #12307 #12316]: #12318 +#12303 := (iff #2964 #12302) +#12304 := [rewrite]: #12303 +#12322 := [monotonicity #12304 #12319]: #12321 +#12300 := (iff #2957 #12299) +#12297 := (iff #2956 #12296) +#12298 := [rewrite]: #12297 +#12301 := [monotonicity #12298]: #12300 +#12325 := [monotonicity #12301 #12322]: #12324 +#12294 := (iff #2954 #12293) +#12295 := [rewrite]: #12294 +#12328 := [monotonicity #12295 #12325]: #12327 +#13810 := [monotonicity #12328 #13807]: #13809 +#13816 := [trans #13810 #13814]: #13815 +#13819 := [monotonicity #13816]: #13818 +#13825 := [trans #13819 #13823]: #13824 +#13828 := [monotonicity #13825]: #13827 +#13834 := [trans #13828 #13832]: #13833 +#13837 := [monotonicity #13834]: #13836 +#13843 := [trans #13837 #13841]: #13842 +#13846 := [monotonicity #13843]: #13845 +#13852 := [trans #13846 #13850]: #13851 +#13855 := [monotonicity #13852]: #13854 +#13861 := [trans #13855 #13859]: #13860 +#13864 := [monotonicity #13861]: #13863 +#13868 := [trans #13864 #13866]: #13867 +#13871 := [monotonicity #13868]: #13870 +#14775 := [trans #13871 #14773]: #14774 +#12292 := [asserted]: #3345 +#14776 := [mp #12292 #14775]: #14771 +#14794 := [not-or-elim #14776]: #14658 +#14798 := [and-elim #14794]: #12305 +#233 := (:var 0 T3) +#15 := (:var 1 T5) +#2661 := (uf_48 #15 #233) +#2662 := (pattern #2661) +#11594 := (= uf_9 #2661) +#11601 := (not #11594) +#1250 := (uf_116 #15) +#2664 := (uf_43 #233 #1250) +#2665 := (= #15 #2664) +#11602 := (or #2665 #11601) +#11607 := (forall (vars (?x710 T5) (?x711 T3)) (:pat #2662) #11602) +#18734 := (~ #11607 #11607) +#18732 := (~ #11602 #11602) +#18733 := [refl]: #18732 +#18735 := [nnf-pos #18733]: #18734 +#2663 := (= #2661 uf_9) +#2666 := (implies #2663 #2665) +#2667 := (forall (vars (?x710 T5) (?x711 T3)) (:pat #2662) #2666) +#11608 := (iff #2667 #11607) +#11605 := (iff #2666 #11602) +#11598 := (implies #11594 #2665) +#11603 := (iff #11598 #11602) +#11604 := [rewrite]: #11603 +#11599 := (iff #2666 #11598) +#11596 := (iff #2663 #11594) +#11597 := [rewrite]: #11596 +#11600 := [monotonicity #11597]: #11599 +#11606 := [trans #11600 #11604]: #11605 +#11609 := [quant-intro #11606]: #11608 +#11593 := [asserted]: #2667 +#11612 := [mp #11593 #11609]: #11607 +#18736 := [mp~ #11612 #18735]: #11607 +#25403 := (not #12305) +#25416 := (not #11607) +#25417 := (or #25416 #25403 #25411) +#25412 := (or #25411 #25403) +#25418 := (or #25416 #25412) +#25425 := (iff #25418 #25417) +#25413 := (or #25403 #25411) +#25420 := (or #25416 #25413) +#25423 := (iff #25420 #25417) +#25424 := [rewrite]: #25423 +#25421 := (iff #25418 #25420) +#25414 := (iff #25412 #25413) +#25415 := [rewrite]: #25414 +#25422 := [monotonicity #25415]: #25421 +#25426 := [trans #25422 #25424]: #25425 +#25419 := [quant-inst]: #25418 +#25427 := [mp #25419 #25426]: #25417 +#27939 := [unit-resolution #25427 #18736 #14798]: #25411 +#27941 := [symm #27939]: #27940 +#26337 := [monotonicity #27941]: #26336 +#26339 := [trans #26337 #28359]: #26338 +#26341 := [monotonicity #26339]: #26340 +#26306 := [monotonicity #26341]: #26342 +#26296 := [symm #26306]: #26293 +#26299 := [monotonicity #26296]: #26298 +#14796 := [and-elim #14794]: #12299 +#26307 := [mp #14796 #26299]: #26297 +decl uf_196 :: (-> T4 T5 T5 T2) +#25980 := (uf_196 uf_273 #25404 #25404) +#25981 := (= uf_9 #25980) +#26002 := (not #25981) +#25982 := (uf_200 uf_273 #25404 #25404 uf_284) +#25983 := (= uf_9 #25982) +#25985 := (iff #25981 #25983) +#2240 := (:var 0 T16) +#24 := (:var 2 T5) +#13 := (:var 3 T4) +#2251 := (uf_200 #13 #24 #15 #2240) +#2252 := (pattern #2251) +#2254 := (uf_196 #13 #24 #15) +#10555 := (= uf_9 #2254) +#10551 := (= uf_9 #2251) +#10558 := (iff #10551 #10555) +#10561 := (forall (vars (?x586 T4) (?x587 T5) (?x588 T5) (?x589 T16)) (:pat #2252) #10558) +#18376 := (~ #10561 #10561) +#18374 := (~ #10558 #10558) +#18375 := [refl]: #18374 +#18377 := [nnf-pos #18375]: #18376 +#2255 := (= #2254 uf_9) +#2253 := (= #2251 uf_9) +#2256 := (iff #2253 #2255) +#2257 := (forall (vars (?x586 T4) (?x587 T5) (?x588 T5) (?x589 T16)) (:pat #2252) #2256) +#10562 := (iff #2257 #10561) +#10559 := (iff #2256 #10558) +#10556 := (iff #2255 #10555) +#10557 := [rewrite]: #10556 +#10553 := (iff #2253 #10551) +#10554 := [rewrite]: #10553 +#10560 := [monotonicity #10554 #10557]: #10559 +#10563 := [quant-intro #10560]: #10562 +#10550 := [asserted]: #2257 +#10566 := [mp #10550 #10563]: #10561 +#18378 := [mp~ #10566 #18377]: #10561 +#25995 := (not #10561) +#25996 := (or #25995 #25985) +#25984 := (iff #25983 #25981) +#25997 := (or #25995 #25984) +#26025 := (iff #25997 #25996) +#26077 := (iff #25996 #25996) +#26078 := [rewrite]: #26077 +#25986 := (iff #25984 #25985) +#25987 := [rewrite]: #25986 +#26076 := [monotonicity #25987]: #26025 +#26079 := [trans #26076 #26078]: #26025 +#26023 := [quant-inst]: #25997 +#26015 := [mp #26023 #26079]: #25996 +#27937 := [unit-resolution #26015 #18378]: #25985 +#25999 := (not #25983) +#26332 := (iff #13715 #25999) +#26334 := (iff #12355 #25983) +#26301 := (iff #25983 #12355) +#27942 := (= #25982 #3009) +#27943 := [monotonicity #27941 #27941]: #27942 +#26333 := [monotonicity #27943]: #26301 +#26335 := [symm #26333]: #26334 +#26349 := [monotonicity #26335]: #26332 +#26300 := [hypothesis]: #13715 +#26350 := [mp #26300 #26349]: #25999 +#26022 := (not #25985) +#25991 := (or #26022 #26002 #25983) +#25989 := [def-axiom]: #25991 +#26348 := [unit-resolution #25989 #26350 #27937]: #26002 +#26086 := (uf_48 #25404 #25815) +#26087 := (= uf_9 #26086) +#26398 := (= #2965 #26086) +#26351 := (= #26086 #2965) +#26352 := [monotonicity #27941 #26339]: #26351 +#26399 := [symm #26352]: #26398 +#26400 := [trans #14798 #26399]: #26087 +#26089 := (uf_27 uf_273 #25404) +#26090 := (= uf_9 #26089) +#26324 := (= #2963 #26089) +#26323 := (= #26089 #2963) +#26325 := [monotonicity #27941]: #26323 +#26327 := [symm #26325]: #26324 +#14797 := [and-elim #14794]: #12302 +#26322 := [trans #14797 #26327]: #26090 +#26091 := (not #26090) +#26088 := (not #26087) +#26490 := (or #25981 #26088 #26091 #26095) +#25827 := (uf_25 uf_273 #25404) +#26084 := (= uf_26 #25827) +#26331 := (= #2967 #25827) +#26328 := (= #25827 #2967) +#26329 := [monotonicity #27941]: #26328 +#26401 := [symm #26329]: #26331 +#14799 := [and-elim #14794]: #12308 +#26403 := [trans #14799 #26401]: #26084 +#25853 := (uf_24 uf_273 #25404) +#25854 := (= uf_9 #25853) +#26391 := (= #2969 #25853) +#26388 := (= #25853 #2969) +#26402 := [monotonicity #27941]: #26388 +#26389 := [symm #26402]: #26391 +#14800 := [and-elim #14794]: #12311 +#26392 := [trans #14800 #26389]: #25854 +#25816 := (uf_22 #25815) +#25823 := (= uf_9 #25816) +#26413 := (= #2953 #25816) +#26393 := (= #25816 #2953) +#26394 := [monotonicity #26339]: #26393 +#26414 := [symm #26394]: #26413 +#14795 := [and-elim #14794]: #12293 +#26488 := [trans #14795 #26414]: #25823 +#14783 := [not-or-elim #14776]: #12338 +#14784 := [and-elim #14783]: #12332 +#47 := (:var 1 T4) +#2213 := (uf_196 #47 #26 #26) +#2214 := (pattern #2213) +#10431 := (= uf_9 #2213) +#227 := (uf_55 #47) +#3939 := (= uf_9 #227) +#19933 := (not #3939) +#150 := (uf_25 #47 #26) +#3656 := (= uf_26 #150) +#19808 := (not #3656) +#33 := (uf_15 #26) +#148 := (uf_48 #26 #33) +#3653 := (= uf_9 #148) +#19807 := (not #3653) +#146 := (uf_27 #47 #26) +#3650 := (= uf_9 #146) +#11522 := (not #3650) +#135 := (uf_24 #47 #26) +#3635 := (= uf_9 #135) +#11145 := (not #3635) +#69 := (uf_22 #33) +#3470 := (= uf_9 #69) +#11200 := (not #3470) +#34 := (uf_14 #33) +#36 := (= #34 uf_16) +#22334 := (or #36 #11200 #11145 #11522 #19807 #19808 #19933 #10431) +#22339 := (forall (vars (?x572 T4) (?x573 T5)) (:pat #2214) #22334) +#52 := (not #36) +#10446 := (and #52 #3470 #3635 #3650 #3653 #3656 #3939) +#10449 := (not #10446) +#10455 := (or #10431 #10449) +#10460 := (forall (vars (?x572 T4) (?x573 T5)) (:pat #2214) #10455) +#22340 := (iff #10460 #22339) +#22337 := (iff #10455 #22334) +#22320 := (or #36 #11200 #11145 #11522 #19807 #19808 #19933) +#22331 := (or #10431 #22320) +#22335 := (iff #22331 #22334) +#22336 := [rewrite]: #22335 +#22332 := (iff #10455 #22331) +#22329 := (iff #10449 #22320) +#22321 := (not #22320) +#22324 := (not #22321) +#22327 := (iff #22324 #22320) +#22328 := [rewrite]: #22327 +#22325 := (iff #10449 #22324) +#22322 := (iff #10446 #22321) +#22323 := [rewrite]: #22322 +#22326 := [monotonicity #22323]: #22325 +#22330 := [trans #22326 #22328]: #22329 +#22333 := [monotonicity #22330]: #22332 +#22338 := [trans #22333 #22336]: #22337 +#22341 := [quant-intro #22338]: #22340 +#18344 := (~ #10460 #10460) +#18342 := (~ #10455 #10455) +#18343 := [refl]: #18342 +#18345 := [nnf-pos #18343]: #18344 +#2220 := (= #2213 uf_9) +#229 := (= #227 uf_9) +#136 := (= #135 uf_9) +#230 := (and #136 #229) +#151 := (= #150 uf_26) +#2215 := (and #151 #230) +#149 := (= #148 uf_9) +#2216 := (and #149 #2215) +#147 := (= #146 uf_9) +#2217 := (and #147 #2216) +#2218 := (and #52 #2217) +#70 := (= #69 uf_9) +#2219 := (and #70 #2218) +#2221 := (implies #2219 #2220) +#2222 := (forall (vars (?x572 T4) (?x573 T5)) (:pat #2214) #2221) +#10463 := (iff #2222 #10460) +#3943 := (and #3635 #3939) +#10415 := (and #3656 #3943) +#10419 := (and #3653 #10415) +#10422 := (and #3650 #10419) +#10425 := (and #52 #10422) +#10428 := (and #3470 #10425) +#10437 := (not #10428) +#10438 := (or #10437 #10431) +#10443 := (forall (vars (?x572 T4) (?x573 T5)) (:pat #2214) #10438) +#10461 := (iff #10443 #10460) +#10458 := (iff #10438 #10455) +#10452 := (or #10449 #10431) +#10456 := (iff #10452 #10455) +#10457 := [rewrite]: #10456 +#10453 := (iff #10438 #10452) +#10450 := (iff #10437 #10449) +#10447 := (iff #10428 #10446) +#10448 := [rewrite]: #10447 +#10451 := [monotonicity #10448]: #10450 +#10454 := [monotonicity #10451]: #10453 +#10459 := [trans #10454 #10457]: #10458 +#10462 := [quant-intro #10459]: #10461 +#10444 := (iff #2222 #10443) +#10441 := (iff #2221 #10438) +#10434 := (implies #10428 #10431) +#10439 := (iff #10434 #10438) +#10440 := [rewrite]: #10439 +#10435 := (iff #2221 #10434) +#10432 := (iff #2220 #10431) +#10433 := [rewrite]: #10432 +#10429 := (iff #2219 #10428) +#10426 := (iff #2218 #10425) +#10423 := (iff #2217 #10422) +#10420 := (iff #2216 #10419) +#10417 := (iff #2215 #10415) +#3944 := (iff #230 #3943) +#3941 := (iff #229 #3939) +#3942 := [rewrite]: #3941 +#3637 := (iff #136 #3635) +#3638 := [rewrite]: #3637 +#3945 := [monotonicity #3638 #3942]: #3944 +#3657 := (iff #151 #3656) +#3658 := [rewrite]: #3657 +#10418 := [monotonicity #3658 #3945]: #10417 +#3654 := (iff #149 #3653) +#3655 := [rewrite]: #3654 +#10421 := [monotonicity #3655 #10418]: #10420 +#3651 := (iff #147 #3650) +#3652 := [rewrite]: #3651 +#10424 := [monotonicity #3652 #10421]: #10423 +#10427 := [monotonicity #10424]: #10426 +#3471 := (iff #70 #3470) +#3472 := [rewrite]: #3471 +#10430 := [monotonicity #3472 #10427]: #10429 +#10436 := [monotonicity #10430 #10433]: #10435 +#10442 := [trans #10436 #10440]: #10441 +#10445 := [quant-intro #10442]: #10444 +#10464 := [trans #10445 #10462]: #10463 +#10414 := [asserted]: #2222 +#10465 := [mp #10414 #10464]: #10460 +#18346 := [mp~ #10465 #18345]: #10460 +#22342 := [mp #18346 #22341]: #22339 +#26085 := (not #26084) +#25880 := (not #25854) +#25824 := (not #25823) +#23209 := (not #12332) +#26081 := (not #22339) +#26110 := (or #26081 #23209 #25824 #25880 #25981 #26085 #26088 #26091 #26095) +#26093 := (= #26092 uf_16) +#26094 := (or #26093 #25824 #25880 #26091 #26088 #26085 #23209 #25981) +#26111 := (or #26081 #26094) +#26219 := (iff #26111 #26110) +#26101 := (or #23209 #25824 #25880 #25981 #26085 #26088 #26091 #26095) +#26107 := (or #26081 #26101) +#26215 := (iff #26107 #26110) +#26218 := [rewrite]: #26215 +#26113 := (iff #26111 #26107) +#26104 := (iff #26094 #26101) +#26098 := (or #26095 #25824 #25880 #26091 #26088 #26085 #23209 #25981) +#26102 := (iff #26098 #26101) +#26103 := [rewrite]: #26102 +#26099 := (iff #26094 #26098) +#26096 := (iff #26093 #26095) +#26097 := [rewrite]: #26096 +#26100 := [monotonicity #26097]: #26099 +#26105 := [trans #26100 #26103]: #26104 +#26150 := [monotonicity #26105]: #26113 +#26181 := [trans #26150 #26218]: #26219 +#26112 := [quant-inst]: #26111 +#26165 := [mp #26112 #26181]: #26110 +#26491 := [unit-resolution #26165 #22342 #14784 #26488 #26392 #26403]: #26490 +#26493 := [unit-resolution #26491 #26322 #26400 #26348 #26307]: false +#26494 := [lemma #26493]: #12355 +#23984 := (or #13715 #23981) +#22978 := (forall (vars (?x782 int)) #22967) +#22985 := (not #22978) +#22963 := (forall (vars (?x781 int)) #22958) +#22984 := (not #22963) +#22986 := (or #22984 #22985) +#22987 := (not #22986) +#23016 := (or #22987 #23013) +#23022 := (not #23016) +#23023 := (or #12671 #12662 #12653 #12644 #22873 #14243 #14049 #23022) +#23024 := (not #23023) +#22802 := (forall (vars (?x785 int)) #22797) +#22808 := (not #22802) +#22809 := (or #22784 #22808) +#22810 := (not #22809) +#22839 := (or #22810 #22836) +#22845 := (not #22839) +#22846 := (or #14105 #22845) +#22847 := (not #22846) +#22852 := (or #14105 #22847) +#22860 := (not #22852) +#22861 := (or #12942 #22858 #19034 #22859 #14172 #19037 #22860) +#22862 := (not #22861) +#22867 := (or #19034 #19037 #22862) +#22874 := (not #22867) +#22884 := (or #13124 #13115 #13090 #19011 #19017 #13133 #13081 #14243 #22858 #22874) +#22885 := (not #22884) +#22890 := (or #19011 #19017 #22885) +#22896 := (not #22890) +#22897 := (or #19008 #19011 #22896) +#22898 := (not #22897) +#22903 := (or #19008 #19011 #22898) +#22909 := (not #22903) +#22910 := (or #22873 #14243 #14207 #22909) +#22911 := (not #22910) +#22875 := (or #13002 #12993 #22873 #14243 #22858 #14211 #22874) +#22876 := (not #22875) +#22916 := (or #22876 #22911) +#22922 := (not #22916) +#22923 := (or #19011 #19017 #22873 #14243 #22922) +#22924 := (not #22923) +#22929 := (or #19011 #19017 #22924) +#22935 := (not #22929) +#22936 := (or #19008 #19011 #22935) +#22937 := (not #22936) +#22942 := (or #19008 #19011 #22937) +#22948 := (not #22942) +#22949 := (or #22873 #14243 #14046 #22948) +#22950 := (not #22949) +#23029 := (or #22950 #23024) +#23044 := (not #23029) +#22779 := (forall (vars (?x774 int)) #22774) +#23040 := (not #22779) +#23045 := (or #13672 #13367 #13358 #13349 #13340 #23035 #23036 #23037 #14399 #15709 #13942 #22873 #14243 #14404 #14456 #23038 #23039 #23041 #23042 #23043 #23040 #23044) +#23046 := (not #23045) +#23051 := (or #13672 #13942 #23046) +#23058 := (not #23051) +#22768 := (forall (vars (?x773 int)) #22763) +#23057 := (not #22768) +#23059 := (or #23057 #23058) +#23060 := (not #23059) +#23065 := (or #22757 #23060) +#23071 := (not #23065) +#23072 := (or #13906 #23071) +#23073 := (not #23072) +#23078 := (or #13906 #23073) +#23084 := (not #23078) +#23085 := (or #13672 #13663 #13654 #13645 #18900 #18906 #23084) +#23086 := (not #23085) +#23091 := (or #18900 #18906 #23086) +#23097 := (not #23091) +#23098 := (or #18897 #18900 #23097) +#23099 := (not #23098) +#23104 := (or #18897 #18900 #23099) +#23110 := (not #23104) +#23111 := (or #13715 #23110) +#23112 := (not #23111) +#23117 := (or #13715 #23112) +#23985 := (iff #23117 #23984) +#23982 := (iff #23112 #23981) +#23979 := (iff #23111 #23978) +#23976 := (iff #23110 #23975) +#23973 := (iff #23104 #23972) +#23970 := (iff #23099 #23969) +#23967 := (iff #23098 #23966) +#23964 := (iff #23097 #23963) +#23961 := (iff #23091 #23960) +#23958 := (iff #23086 #23957) +#23955 := (iff #23085 #23954) +#23952 := (iff #23084 #23951) +#23949 := (iff #23078 #23948) +#23946 := (iff #23073 #23945) +#23943 := (iff #23072 #23942) +#23940 := (iff #23071 #23939) +#23937 := (iff #23065 #23936) +#23934 := (iff #23060 #23933) +#23931 := (iff #23059 #23930) +#23928 := (iff #23058 #23927) +#23925 := (iff #23051 #23924) +#23922 := (iff #23046 #23921) +#23919 := (iff #23045 #23918) +#23916 := (iff #23044 #23915) +#23913 := (iff #23029 #23912) +#23910 := (iff #23024 #23909) +#23907 := (iff #23023 #23906) +#23904 := (iff #23022 #23903) +#23901 := (iff #23016 #23900) +#23898 := (iff #22987 #23897) +#23895 := (iff #22986 #23894) +#23892 := (iff #22985 #23891) +#23889 := (iff #22978 #23886) +#23887 := (iff #22967 #22967) +#23888 := [refl]: #23887 +#23890 := [quant-intro #23888]: #23889 +#23893 := [monotonicity #23890]: #23892 +#23884 := (iff #22984 #23883) +#23881 := (iff #22963 #23878) +#23879 := (iff #22958 #22958) +#23880 := [refl]: #23879 +#23882 := [quant-intro #23880]: #23881 +#23885 := [monotonicity #23882]: #23884 +#23896 := [monotonicity #23885 #23893]: #23895 +#23899 := [monotonicity #23896]: #23898 +#23902 := [monotonicity #23899]: #23901 +#23905 := [monotonicity #23902]: #23904 +#23908 := [monotonicity #23905]: #23907 +#23911 := [monotonicity #23908]: #23910 +#23876 := (iff #22950 #23875) +#23873 := (iff #22949 #23872) +#23870 := (iff #22948 #23869) +#23867 := (iff #22942 #23866) +#23864 := (iff #22937 #23863) +#23861 := (iff #22936 #23860) +#23858 := (iff #22935 #23857) +#23855 := (iff #22929 #23854) +#23852 := (iff #22924 #23851) +#23849 := (iff #22923 #23848) +#23846 := (iff #22922 #23845) +#23843 := (iff #22916 #23842) +#23840 := (iff #22911 #23839) +#23837 := (iff #22910 #23836) +#23834 := (iff #22909 #23833) +#23831 := (iff #22903 #23830) +#23828 := (iff #22898 #23827) +#23825 := (iff #22897 #23824) +#23822 := (iff #22896 #23821) +#23819 := (iff #22890 #23818) +#23816 := (iff #22885 #23815) +#23813 := (iff #22884 #23812) +#23804 := (iff #22874 #23803) +#23801 := (iff #22867 #23800) +#23798 := (iff #22862 #23797) +#23795 := (iff #22861 #23794) +#23792 := (iff #22860 #23791) +#23789 := (iff #22852 #23788) +#23786 := (iff #22847 #23785) +#23783 := (iff #22846 #23782) +#23780 := (iff #22845 #23779) +#23777 := (iff #22839 #23776) +#23774 := (iff #22810 #23773) +#23771 := (iff #22809 #23770) +#23768 := (iff #22808 #23767) +#23765 := (iff #22802 #23762) +#23763 := (iff #22797 #22797) +#23764 := [refl]: #23763 +#23766 := [quant-intro #23764]: #23765 +#23769 := [monotonicity #23766]: #23768 +#23772 := [monotonicity #23769]: #23771 +#23775 := [monotonicity #23772]: #23774 +#23778 := [monotonicity #23775]: #23777 +#23781 := [monotonicity #23778]: #23780 +#23784 := [monotonicity #23781]: #23783 +#23787 := [monotonicity #23784]: #23786 +#23790 := [monotonicity #23787]: #23789 +#23793 := [monotonicity #23790]: #23792 +#23796 := [monotonicity #23793]: #23795 +#23799 := [monotonicity #23796]: #23798 +#23802 := [monotonicity #23799]: #23801 +#23805 := [monotonicity #23802]: #23804 +#23814 := [monotonicity #23805]: #23813 +#23817 := [monotonicity #23814]: #23816 +#23820 := [monotonicity #23817]: #23819 +#23823 := [monotonicity #23820]: #23822 +#23826 := [monotonicity #23823]: #23825 +#23829 := [monotonicity #23826]: #23828 +#23832 := [monotonicity #23829]: #23831 +#23835 := [monotonicity #23832]: #23834 +#23838 := [monotonicity #23835]: #23837 +#23841 := [monotonicity #23838]: #23840 +#23810 := (iff #22876 #23809) +#23807 := (iff #22875 #23806) +#23808 := [monotonicity #23805]: #23807 +#23811 := [monotonicity #23808]: #23810 +#23844 := [monotonicity #23811 #23841]: #23843 +#23847 := [monotonicity #23844]: #23846 +#23850 := [monotonicity #23847]: #23849 +#23853 := [monotonicity #23850]: #23852 +#23856 := [monotonicity #23853]: #23855 +#23859 := [monotonicity #23856]: #23858 +#23862 := [monotonicity #23859]: #23861 +#23865 := [monotonicity #23862]: #23864 +#23868 := [monotonicity #23865]: #23867 +#23871 := [monotonicity #23868]: #23870 +#23874 := [monotonicity #23871]: #23873 +#23877 := [monotonicity #23874]: #23876 +#23914 := [monotonicity #23877 #23911]: #23913 +#23917 := [monotonicity #23914]: #23916 +#23760 := (iff #23040 #23759) +#23757 := (iff #22779 #23754) +#23755 := (iff #22774 #22774) +#23756 := [refl]: #23755 +#23758 := [quant-intro #23756]: #23757 +#23761 := [monotonicity #23758]: #23760 +#23920 := [monotonicity #23761 #23917]: #23919 +#23923 := [monotonicity #23920]: #23922 +#23926 := [monotonicity #23923]: #23925 +#23929 := [monotonicity #23926]: #23928 +#23752 := (iff #23057 #23751) +#23749 := (iff #22768 #23746) +#23747 := (iff #22763 #22763) +#23748 := [refl]: #23747 +#23750 := [quant-intro #23748]: #23749 +#23753 := [monotonicity #23750]: #23752 +#23932 := [monotonicity #23753 #23929]: #23931 +#23935 := [monotonicity #23932]: #23934 +#23938 := [monotonicity #23935]: #23937 +#23941 := [monotonicity #23938]: #23940 +#23944 := [monotonicity #23941]: #23943 +#23947 := [monotonicity #23944]: #23946 +#23950 := [monotonicity #23947]: #23949 +#23953 := [monotonicity #23950]: #23952 +#23956 := [monotonicity #23953]: #23955 +#23959 := [monotonicity #23956]: #23958 +#23962 := [monotonicity #23959]: #23961 +#23965 := [monotonicity #23962]: #23964 +#23968 := [monotonicity #23965]: #23967 +#23971 := [monotonicity #23968]: #23970 +#23974 := [monotonicity #23971]: #23973 +#23977 := [monotonicity #23974]: #23976 +#23980 := [monotonicity #23977]: #23979 +#23983 := [monotonicity #23980]: #23982 +#23986 := [monotonicity #23983]: #23985 +#19548 := (and #19191 #19192) +#19551 := (not #19548) +#19554 := (or #19530 #19543 #19551) +#19557 := (not #19554) +#16489 := (and #3145 #4084 #13958 #15606) +#19214 := (not #16489) +#19217 := (forall (vars (?x782 int)) #19214) +#14858 := (and #4084 #15606) +#14857 := (not #14858) +#16475 := (or #13959 #13972 #14857) +#16480 := (forall (vars (?x781 int)) #16475) +#19221 := (and #16480 #19217) +#19563 := (or #19221 #19557) +#19571 := (and #12567 #12570 #12573 #12576 #13947 #13950 #14046 #19563) +#19392 := (and #19055 #19056) +#19395 := (not #19392) +#19398 := (or #19374 #19387 #19395) +#19401 := (not #19398) +#16376 := (or #14109 #14122 #14857) +#16381 := (forall (vars (?x785 int)) #16376) +#19071 := (not #14151) +#19081 := (and #19071 #16381) +#19407 := (or #19081 #19401) +#19412 := (and #14100 #19407) +#19415 := (or #14105 #19412) +#19423 := (and #3199 #14076 #14088 #14092 #14168 #16368 #19415) +#19428 := (or #19034 #19037 #19423) +#19454 := (and #3242 #3244 #3246 #12806 #12812 #13070 #13075 #13950 #14076 #19428) +#19459 := (or #19011 #19017 #19454) +#19465 := (and #12803 #12806 #19459) +#19470 := (or #19008 #19011 #19465) +#19476 := (and #13947 #13950 #14211 #19470) +#19434 := (and #12823 #12826 #13947 #13950 #14076 #14207 #19428) +#19481 := (or #19434 #19476) +#19487 := (and #12806 #12812 #13947 #13950 #19481) +#19492 := (or #19011 #19017 #19487) +#19498 := (and #12803 #12806 #19492) +#19503 := (or #19008 #19011 #19498) +#19509 := (and #13947 #13950 #14049 #19503) +#19576 := (or #19509 #19571) +#16293 := (or #14420 #14433 #14857) +#16298 := (forall (vars (?x774 int)) #16293) +#19582 := (and #3022 #3121 #3122 #3123 #3124 #3125 #3126 #12426 #12437 #12553 #13943 #13947 #13950 #14405 #14453 #14462 #14490 #16298 #16310 #16332 #16349 #19576) +#19587 := (or #13672 #13942 #19582) +#16279 := (or #13910 #13921 #14857) +#16284 := (forall (vars (?x773 int)) #16279) +#19590 := (and #16284 #19587) +#19303 := (and #18930 #18931) +#19306 := (not #19303) +#19312 := (or #18938 #18939 #19306) +#19317 := (not #19312) +#19593 := (or #19317 #19590) +#19596 := (and #13903 #19593) +#19599 := (or #13906 #19596) +#19605 := (and #3022 #3025 #3028 #3031 #12361 #12367 #19599) +#19610 := (or #18900 #18906 #19605) +#19616 := (and #12358 #12361 #19610) +#19621 := (or #18897 #18900 #19616) +#19624 := (and #12355 #19621) +#19627 := (or #13715 #19624) +#23118 := (iff #19627 #23117) +#23115 := (iff #19624 #23112) +#23107 := (and #12355 #23104) +#23113 := (iff #23107 #23112) +#23114 := [rewrite]: #23113 +#23108 := (iff #19624 #23107) +#23105 := (iff #19621 #23104) +#23102 := (iff #19616 #23099) +#23094 := (and #12358 #12361 #23091) +#23100 := (iff #23094 #23099) +#23101 := [rewrite]: #23100 +#23095 := (iff #19616 #23094) +#23092 := (iff #19610 #23091) +#23089 := (iff #19605 #23086) +#23081 := (and #3022 #3025 #3028 #3031 #12361 #12367 #23078) +#23087 := (iff #23081 #23086) +#23088 := [rewrite]: #23087 +#23082 := (iff #19605 #23081) +#23079 := (iff #19599 #23078) +#23076 := (iff #19596 #23073) +#23068 := (and #13903 #23065) +#23074 := (iff #23068 #23073) +#23075 := [rewrite]: #23074 +#23069 := (iff #19596 #23068) +#23066 := (iff #19593 #23065) +#23063 := (iff #19590 #23060) +#23054 := (and #22768 #23051) +#23061 := (iff #23054 #23060) +#23062 := [rewrite]: #23061 +#23055 := (iff #19590 #23054) +#23052 := (iff #19587 #23051) +#23049 := (iff #19582 #23046) +#23032 := (and #3022 #3121 #3122 #3123 #3124 #3125 #3126 #12426 #12437 #12553 #13943 #13947 #13950 #14405 #14453 #14462 #14490 #22779 #16310 #16332 #16349 #23029) +#23047 := (iff #23032 #23046) +#23048 := [rewrite]: #23047 +#23033 := (iff #19582 #23032) +#23030 := (iff #19576 #23029) +#23027 := (iff #19571 #23024) +#23019 := (and #12567 #12570 #12573 #12576 #13947 #13950 #14046 #23016) +#23025 := (iff #23019 #23024) +#23026 := [rewrite]: #23025 +#23020 := (iff #19571 #23019) +#23017 := (iff #19563 #23016) +#23014 := (iff #19557 #23013) +#23011 := (iff #19554 #23008) +#22994 := (or #22992 #22993) +#23005 := (or #19530 #19543 #22994) +#23009 := (iff #23005 #23008) +#23010 := [rewrite]: #23009 +#23006 := (iff #19554 #23005) +#23003 := (iff #19551 #22994) +#22995 := (not #22994) +#22998 := (not #22995) +#23001 := (iff #22998 #22994) +#23002 := [rewrite]: #23001 +#22999 := (iff #19551 #22998) +#22996 := (iff #19548 #22995) +#22997 := [rewrite]: #22996 +#23000 := [monotonicity #22997]: #22999 +#23004 := [trans #23000 #23002]: #23003 +#23007 := [monotonicity #23004]: #23006 +#23012 := [trans #23007 #23010]: #23011 +#23015 := [monotonicity #23012]: #23014 +#22990 := (iff #19221 #22987) +#22981 := (and #22963 #22978) +#22988 := (iff #22981 #22987) +#22989 := [rewrite]: #22988 +#22982 := (iff #19221 #22981) +#22979 := (iff #19217 #22978) +#22976 := (iff #19214 #22967) +#22968 := (not #22967) +#22971 := (not #22968) +#22974 := (iff #22971 #22967) +#22975 := [rewrite]: #22974 +#22972 := (iff #19214 #22971) +#22969 := (iff #16489 #22968) +#22970 := [rewrite]: #22969 +#22973 := [monotonicity #22970]: #22972 +#22977 := [trans #22973 #22975]: #22976 +#22980 := [quant-intro #22977]: #22979 +#22964 := (iff #16480 #22963) +#22961 := (iff #16475 #22958) +#20695 := (or #5113 #20064) +#22955 := (or #13959 #13972 #20695) +#22959 := (iff #22955 #22958) +#22960 := [rewrite]: #22959 +#22956 := (iff #16475 #22955) +#20704 := (iff #14857 #20695) +#20696 := (not #20695) +#20699 := (not #20696) +#20702 := (iff #20699 #20695) +#20703 := [rewrite]: #20702 +#20700 := (iff #14857 #20699) +#20697 := (iff #14858 #20696) +#20698 := [rewrite]: #20697 +#20701 := [monotonicity #20698]: #20700 +#20705 := [trans #20701 #20703]: #20704 +#22957 := [monotonicity #20705]: #22956 +#22962 := [trans #22957 #22960]: #22961 +#22965 := [quant-intro #22962]: #22964 +#22983 := [monotonicity #22965 #22980]: #22982 +#22991 := [trans #22983 #22989]: #22990 +#23018 := [monotonicity #22991 #23015]: #23017 +#23021 := [monotonicity #23018]: #23020 +#23028 := [trans #23021 #23026]: #23027 +#22953 := (iff #19509 #22950) +#22945 := (and #13947 #13950 #14049 #22942) +#22951 := (iff #22945 #22950) +#22952 := [rewrite]: #22951 +#22946 := (iff #19509 #22945) +#22943 := (iff #19503 #22942) +#22940 := (iff #19498 #22937) +#22932 := (and #12803 #12806 #22929) +#22938 := (iff #22932 #22937) +#22939 := [rewrite]: #22938 +#22933 := (iff #19498 #22932) +#22930 := (iff #19492 #22929) +#22927 := (iff #19487 #22924) +#22919 := (and #12806 #12812 #13947 #13950 #22916) +#22925 := (iff #22919 #22924) +#22926 := [rewrite]: #22925 +#22920 := (iff #19487 #22919) +#22917 := (iff #19481 #22916) +#22914 := (iff #19476 #22911) +#22906 := (and #13947 #13950 #14211 #22903) +#22912 := (iff #22906 #22911) +#22913 := [rewrite]: #22912 +#22907 := (iff #19476 #22906) +#22904 := (iff #19470 #22903) +#22901 := (iff #19465 #22898) +#22893 := (and #12803 #12806 #22890) +#22899 := (iff #22893 #22898) +#22900 := [rewrite]: #22899 +#22894 := (iff #19465 #22893) +#22891 := (iff #19459 #22890) +#22888 := (iff #19454 #22885) +#22881 := (and #3242 #3244 #3246 #12806 #12812 #13070 #13075 #13950 #14076 #22867) +#22886 := (iff #22881 #22885) +#22887 := [rewrite]: #22886 +#22882 := (iff #19454 #22881) +#22868 := (iff #19428 #22867) +#22865 := (iff #19423 #22862) +#22855 := (and #3199 #14076 #14088 #14092 #14168 #16368 #22852) +#22863 := (iff #22855 #22862) +#22864 := [rewrite]: #22863 +#22856 := (iff #19423 #22855) +#22853 := (iff #19415 #22852) +#22850 := (iff #19412 #22847) +#22842 := (and #14100 #22839) +#22848 := (iff #22842 #22847) +#22849 := [rewrite]: #22848 +#22843 := (iff #19412 #22842) +#22840 := (iff #19407 #22839) +#22837 := (iff #19401 #22836) +#22834 := (iff #19398 #22831) +#22817 := (or #22815 #22816) +#22828 := (or #19374 #19387 #22817) +#22832 := (iff #22828 #22831) +#22833 := [rewrite]: #22832 +#22829 := (iff #19398 #22828) +#22826 := (iff #19395 #22817) +#22818 := (not #22817) +#22821 := (not #22818) +#22824 := (iff #22821 #22817) +#22825 := [rewrite]: #22824 +#22822 := (iff #19395 #22821) +#22819 := (iff #19392 #22818) +#22820 := [rewrite]: #22819 +#22823 := [monotonicity #22820]: #22822 +#22827 := [trans #22823 #22825]: #22826 +#22830 := [monotonicity #22827]: #22829 +#22835 := [trans #22830 #22833]: #22834 +#22838 := [monotonicity #22835]: #22837 +#22813 := (iff #19081 #22810) +#22805 := (and #22783 #22802) +#22811 := (iff #22805 #22810) +#22812 := [rewrite]: #22811 +#22806 := (iff #19081 #22805) +#22803 := (iff #16381 #22802) +#22800 := (iff #16376 #22797) +#22794 := (or #14109 #14122 #20695) +#22798 := (iff #22794 #22797) +#22799 := [rewrite]: #22798 +#22795 := (iff #16376 #22794) +#22796 := [monotonicity #20705]: #22795 +#22801 := [trans #22796 #22799]: #22800 +#22804 := [quant-intro #22801]: #22803 +#22792 := (iff #19071 #22783) +#22787 := (not #22784) +#22790 := (iff #22787 #22783) +#22791 := [rewrite]: #22790 +#22788 := (iff #19071 #22787) +#22785 := (iff #14151 #22784) +#22786 := [rewrite]: #22785 +#22789 := [monotonicity #22786]: #22788 +#22793 := [trans #22789 #22791]: #22792 +#22807 := [monotonicity #22793 #22804]: #22806 +#22814 := [trans #22807 #22812]: #22813 +#22841 := [monotonicity #22814 #22838]: #22840 +#22844 := [monotonicity #22841]: #22843 +#22851 := [trans #22844 #22849]: #22850 +#22854 := [monotonicity #22851]: #22853 +#22857 := [monotonicity #22854]: #22856 +#22866 := [trans #22857 #22864]: #22865 +#22869 := [monotonicity #22866]: #22868 +#22883 := [monotonicity #22869]: #22882 +#22889 := [trans #22883 #22887]: #22888 +#22892 := [monotonicity #22889]: #22891 +#22895 := [monotonicity #22892]: #22894 +#22902 := [trans #22895 #22900]: #22901 +#22905 := [monotonicity #22902]: #22904 +#22908 := [monotonicity #22905]: #22907 +#22915 := [trans #22908 #22913]: #22914 +#22879 := (iff #19434 #22876) +#22870 := (and #12823 #12826 #13947 #13950 #14076 #14207 #22867) +#22877 := (iff #22870 #22876) +#22878 := [rewrite]: #22877 +#22871 := (iff #19434 #22870) +#22872 := [monotonicity #22869]: #22871 +#22880 := [trans #22872 #22878]: #22879 +#22918 := [monotonicity #22880 #22915]: #22917 +#22921 := [monotonicity #22918]: #22920 +#22928 := [trans #22921 #22926]: #22927 +#22931 := [monotonicity #22928]: #22930 +#22934 := [monotonicity #22931]: #22933 +#22941 := [trans #22934 #22939]: #22940 +#22944 := [monotonicity #22941]: #22943 +#22947 := [monotonicity #22944]: #22946 +#22954 := [trans #22947 #22952]: #22953 +#23031 := [monotonicity #22954 #23028]: #23030 +#22780 := (iff #16298 #22779) +#22777 := (iff #16293 #22774) +#22771 := (or #14420 #14433 #20695) +#22775 := (iff #22771 #22774) +#22776 := [rewrite]: #22775 +#22772 := (iff #16293 #22771) +#22773 := [monotonicity #20705]: #22772 +#22778 := [trans #22773 #22776]: #22777 +#22781 := [quant-intro #22778]: #22780 +#23034 := [monotonicity #22781 #23031]: #23033 +#23050 := [trans #23034 #23048]: #23049 +#23053 := [monotonicity #23050]: #23052 +#22769 := (iff #16284 #22768) +#22766 := (iff #16279 #22763) +#22760 := (or #13910 #13921 #20695) +#22764 := (iff #22760 #22763) +#22765 := [rewrite]: #22764 +#22761 := (iff #16279 #22760) +#22762 := [monotonicity #20705]: #22761 +#22767 := [trans #22762 #22765]: #22766 +#22770 := [quant-intro #22767]: #22769 +#23056 := [monotonicity #22770 #23053]: #23055 +#23064 := [trans #23056 #23062]: #23063 +#22758 := (iff #19317 #22757) +#22755 := (iff #19312 #22752) +#22738 := (or #22736 #22737) +#22749 := (or #18938 #18939 #22738) +#22753 := (iff #22749 #22752) +#22754 := [rewrite]: #22753 +#22750 := (iff #19312 #22749) +#22747 := (iff #19306 #22738) +#22739 := (not #22738) +#22742 := (not #22739) +#22745 := (iff #22742 #22738) +#22746 := [rewrite]: #22745 +#22743 := (iff #19306 #22742) +#22740 := (iff #19303 #22739) +#22741 := [rewrite]: #22740 +#22744 := [monotonicity #22741]: #22743 +#22748 := [trans #22744 #22746]: #22747 +#22751 := [monotonicity #22748]: #22750 +#22756 := [trans #22751 #22754]: #22755 +#22759 := [monotonicity #22756]: #22758 +#23067 := [monotonicity #22759 #23064]: #23066 +#23070 := [monotonicity #23067]: #23069 +#23077 := [trans #23070 #23075]: #23076 +#23080 := [monotonicity #23077]: #23079 +#23083 := [monotonicity #23080]: #23082 +#23090 := [trans #23083 #23088]: #23089 +#23093 := [monotonicity #23090]: #23092 +#23096 := [monotonicity #23093]: #23095 +#23103 := [trans #23096 #23101]: #23102 +#23106 := [monotonicity #23103]: #23105 +#23109 := [monotonicity #23106]: #23108 +#23116 := [trans #23109 #23114]: #23115 +#23119 := [monotonicity #23116]: #23118 +#19193 := (and #19192 #19191) +#19194 := (not #19193) +#19197 := (+ #19196 #13970) +#19198 := (<= #19197 0::int) +#19199 := (+ ?x781!15 #13873) +#19200 := (>= #19199 0::int) +#19201 := (or #19200 #19198 #19194) +#19202 := (not #19201) +#19225 := (or #19202 #19221) +#18978 := (not #13955) +#19185 := (not #12644) +#19182 := (not #12653) +#19179 := (not #12662) +#19176 := (not #12671) +#19229 := (and #19176 #19179 #19182 #19185 #18978 #14352 #19225) +#16407 := (and #14088 #16368) +#16412 := (not #16407) +#19097 := (not #16412) +#19057 := (and #19056 #19055) +#19058 := (not #19057) +#19061 := (+ #19060 #14120) +#19062 := (<= #19061 0::int) +#19063 := (+ ?x785!14 #14101) +#19064 := (>= #19063 0::int) +#19065 := (or #19064 #19062 #19058) +#19066 := (not #19065) +#19085 := (or #19066 #19081) +#19051 := (not #14105) +#19089 := (and #19051 #19085) +#19093 := (or #14105 #19089) +#19046 := (not #14172) +#19043 := (not #14097) +#19040 := (not #12942) +#19100 := (and #19040 #19043 #19046 #19093 #19097) +#19104 := (or #19034 #19037 #19100) +#19029 := (not #14081) +#19129 := (not #14243) +#19126 := (not #13081) +#19123 := (not #13133) +#19020 := (not #13142) +#19120 := (not #13090) +#19117 := (not #13115) +#19114 := (not #13124) +#19132 := (and #19114 #19117 #19120 #19020 #19123 #19126 #19129 #19029 #19104) +#19136 := (or #19011 #19017 #19132) +#19014 := (not #13159) +#19140 := (and #19014 #19136) +#19144 := (or #19008 #19011 #19140) +#19148 := (and #18978 #14211 #19144) +#19026 := (not #12993) +#19023 := (not #13002) +#19108 := (and #19023 #19026 #18978 #19029 #14293 #19104) +#19152 := (or #19108 #19148) +#19156 := (and #19020 #18978 #19152) +#19160 := (or #19011 #19017 #19156) +#19164 := (and #19014 #19160) +#19168 := (or #19008 #19011 #19164) +#19172 := (and #18978 #14049 #19168) +#19233 := (or #19172 #19229) +#16357 := (and #14490 #16349) +#16362 := (not #16357) +#19003 := (not #16362) +#16337 := (and #13947 #16332) +#16340 := (not #16337) +#19000 := (not #16340) +#16318 := (and #14462 #16310) +#16323 := (not #16318) +#18997 := (not #16323) +#18987 := (not #14507) +#18984 := (not #14456) +#18981 := (not #14416) +#18975 := (not #15709) +#18972 := (not #14399) +#18969 := (not #13331) +#18966 := (not #13340) +#18963 := (not #13349) +#18960 := (not #13358) +#18957 := (not #13367) +#19237 := (and #18957 #18960 #18963 #18966 #18969 #18972 #18975 #18978 #18981 #18984 #18987 #16298 #18997 #19000 #19003 #19233) +#19241 := (or #13672 #14664 #19237) +#19245 := (and #16284 #19241) +#18932 := (and #18931 #18930) +#18933 := (not #18932) +#18940 := (or #18939 #18938 #18933) +#18941 := (not #18940) +#19249 := (or #18941 #19245) +#18926 := (not #13906) +#19253 := (and #18926 #19249) +#19257 := (or #13906 #19253) +#18921 := (not #13681) +#18918 := (not #13645) +#18915 := (not #13654) +#18912 := (not #13663) +#18909 := (not #13672) +#19261 := (and #18909 #18912 #18915 #18918 #18921 #19257) +#19265 := (or #18900 #18906 #19261) +#18903 := (not #13698) +#19269 := (and #18903 #19265) +#19273 := (or #18897 #18900 #19269) +#18894 := (not #13715) +#19277 := (and #18894 #19273) +#19281 := (or #13715 #19277) +#19628 := (iff #19281 #19627) +#19625 := (iff #19277 #19624) +#19622 := (iff #19273 #19621) +#19619 := (iff #19269 #19616) +#19613 := (and #12364 #19610) +#19617 := (iff #19613 #19616) +#19618 := [rewrite]: #19617 +#19614 := (iff #19269 #19613) +#19611 := (iff #19265 #19610) +#19608 := (iff #19261 #19605) +#19602 := (and #3022 #3025 #3028 #3031 #12373 #19599) +#19606 := (iff #19602 #19605) +#19607 := [rewrite]: #19606 +#19603 := (iff #19261 #19602) +#19600 := (iff #19257 #19599) +#19597 := (iff #19253 #19596) +#19594 := (iff #19249 #19593) +#19591 := (iff #19245 #19590) +#19588 := (iff #19241 #19587) +#19585 := (iff #19237 #19582) +#19579 := (and #3121 #3122 #3123 #3124 #3127 #12437 #12553 #13952 #14411 #14453 #14502 #16298 #16318 #16337 #16357 #19576) +#19583 := (iff #19579 #19582) +#19584 := [rewrite]: #19583 +#19580 := (iff #19237 #19579) +#19577 := (iff #19233 #19576) +#19574 := (iff #19229 #19571) +#19568 := (and #12567 #12570 #12573 #12576 #13952 #14046 #19563) +#19572 := (iff #19568 #19571) +#19573 := [rewrite]: #19572 +#19569 := (iff #19229 #19568) +#19566 := (iff #19225 #19563) +#19560 := (or #19557 #19221) +#19564 := (iff #19560 #19563) +#19565 := [rewrite]: #19564 +#19561 := (iff #19225 #19560) +#19558 := (iff #19202 #19557) +#19555 := (iff #19201 #19554) +#19552 := (iff #19194 #19551) +#19549 := (iff #19193 #19548) +#19550 := [rewrite]: #19549 +#19553 := [monotonicity #19550]: #19552 +#19546 := (iff #19198 #19543) +#19535 := (+ #13970 #19196) +#19538 := (<= #19535 0::int) +#19544 := (iff #19538 #19543) +#19545 := [rewrite]: #19544 +#19539 := (iff #19198 #19538) +#19536 := (= #19197 #19535) +#19537 := [rewrite]: #19536 +#19540 := [monotonicity #19537]: #19539 +#19547 := [trans #19540 #19545]: #19546 +#19533 := (iff #19200 #19530) +#19522 := (+ #13873 ?x781!15) +#19525 := (>= #19522 0::int) +#19531 := (iff #19525 #19530) +#19532 := [rewrite]: #19531 +#19526 := (iff #19200 #19525) +#19523 := (= #19199 #19522) +#19524 := [rewrite]: #19523 +#19527 := [monotonicity #19524]: #19526 +#19534 := [trans #19527 #19532]: #19533 +#19556 := [monotonicity #19534 #19547 #19553]: #19555 +#19559 := [monotonicity #19556]: #19558 +#19562 := [monotonicity #19559]: #19561 +#19567 := [trans #19562 #19565]: #19566 +#19334 := (iff #18978 #13952) +#19335 := [rewrite]: #19334 +#19520 := (iff #19185 #12576) +#19521 := [rewrite]: #19520 +#19518 := (iff #19182 #12573) +#19519 := [rewrite]: #19518 +#19516 := (iff #19179 #12570) +#19517 := [rewrite]: #19516 +#19514 := (iff #19176 #12567) +#19515 := [rewrite]: #19514 +#19570 := [monotonicity #19515 #19517 #19519 #19521 #19335 #14356 #19567]: #19569 +#19575 := [trans #19570 #19573]: #19574 +#19512 := (iff #19172 #19509) +#19506 := (and #13952 #14049 #19503) +#19510 := (iff #19506 #19509) +#19511 := [rewrite]: #19510 +#19507 := (iff #19172 #19506) +#19504 := (iff #19168 #19503) +#19501 := (iff #19164 #19498) +#19495 := (and #12809 #19492) +#19499 := (iff #19495 #19498) +#19500 := [rewrite]: #19499 +#19496 := (iff #19164 #19495) +#19493 := (iff #19160 #19492) +#19490 := (iff #19156 #19487) +#19484 := (and #12818 #13952 #19481) +#19488 := (iff #19484 #19487) +#19489 := [rewrite]: #19488 +#19485 := (iff #19156 #19484) +#19482 := (iff #19152 #19481) +#19479 := (iff #19148 #19476) +#19473 := (and #13952 #14211 #19470) +#19477 := (iff #19473 #19476) +#19478 := [rewrite]: #19477 +#19474 := (iff #19148 #19473) +#19471 := (iff #19144 #19470) +#19468 := (iff #19140 #19465) +#19462 := (and #12809 #19459) +#19466 := (iff #19462 #19465) +#19467 := [rewrite]: #19466 +#19463 := (iff #19140 #19462) +#19460 := (iff #19136 #19459) +#19457 := (iff #19132 #19454) +#19451 := (and #3242 #3244 #3246 #12818 #13070 #13075 #13950 #14078 #19428) +#19455 := (iff #19451 #19454) +#19456 := [rewrite]: #19455 +#19452 := (iff #19132 #19451) +#19429 := (iff #19104 #19428) +#19426 := (iff #19100 #19423) +#19420 := (and #3199 #14094 #14168 #19415 #16407) +#19424 := (iff #19420 #19423) +#19425 := [rewrite]: #19424 +#19421 := (iff #19100 #19420) +#19418 := (iff #19097 #16407) +#19419 := [rewrite]: #19418 +#19416 := (iff #19093 #19415) +#19413 := (iff #19089 #19412) +#19410 := (iff #19085 #19407) +#19404 := (or #19401 #19081) +#19408 := (iff #19404 #19407) +#19409 := [rewrite]: #19408 +#19405 := (iff #19085 #19404) +#19402 := (iff #19066 #19401) +#19399 := (iff #19065 #19398) +#19396 := (iff #19058 #19395) +#19393 := (iff #19057 #19392) +#19394 := [rewrite]: #19393 +#19397 := [monotonicity #19394]: #19396 +#19390 := (iff #19062 #19387) +#19379 := (+ #14120 #19060) +#19382 := (<= #19379 0::int) +#19388 := (iff #19382 #19387) +#19389 := [rewrite]: #19388 +#19383 := (iff #19062 #19382) +#19380 := (= #19061 #19379) +#19381 := [rewrite]: #19380 +#19384 := [monotonicity #19381]: #19383 +#19391 := [trans #19384 #19389]: #19390 +#19377 := (iff #19064 #19374) +#19366 := (+ #14101 ?x785!14) +#19369 := (>= #19366 0::int) +#19375 := (iff #19369 #19374) +#19376 := [rewrite]: #19375 +#19370 := (iff #19064 #19369) +#19367 := (= #19063 #19366) +#19368 := [rewrite]: #19367 +#19371 := [monotonicity #19368]: #19370 +#19378 := [trans #19371 #19376]: #19377 +#19400 := [monotonicity #19378 #19391 #19397]: #19399 +#19403 := [monotonicity #19400]: #19402 +#19406 := [monotonicity #19403]: #19405 +#19411 := [trans #19406 #19409]: #19410 +#19364 := (iff #19051 #14100) +#19365 := [rewrite]: #19364 +#19414 := [monotonicity #19365 #19411]: #19413 +#19417 := [monotonicity #19414]: #19416 +#19362 := (iff #19046 #14168) +#19363 := [rewrite]: #19362 +#19360 := (iff #19043 #14094) +#19361 := [rewrite]: #19360 +#19358 := (iff #19040 #3199) +#19359 := [rewrite]: #19358 +#19422 := [monotonicity #19359 #19361 #19363 #19417 #19419]: #19421 +#19427 := [trans #19422 #19425]: #19426 +#19430 := [monotonicity #19427]: #19429 +#19356 := (iff #19029 #14078) +#19357 := [rewrite]: #19356 +#19449 := (iff #19129 #13950) +#19450 := [rewrite]: #19449 +#19447 := (iff #19126 #13075) +#19448 := [rewrite]: #19447 +#19445 := (iff #19123 #13070) +#19446 := [rewrite]: #19445 +#19350 := (iff #19020 #12818) +#19351 := [rewrite]: #19350 +#19443 := (iff #19120 #3246) +#19444 := [rewrite]: #19443 +#19441 := (iff #19117 #3244) +#19442 := [rewrite]: #19441 +#19439 := (iff #19114 #3242) +#19440 := [rewrite]: #19439 +#19453 := [monotonicity #19440 #19442 #19444 #19351 #19446 #19448 #19450 #19357 #19430]: #19452 +#19458 := [trans #19453 #19456]: #19457 +#19461 := [monotonicity #19458]: #19460 +#19348 := (iff #19014 #12809) +#19349 := [rewrite]: #19348 +#19464 := [monotonicity #19349 #19461]: #19463 +#19469 := [trans #19464 #19467]: #19468 +#19472 := [monotonicity #19469]: #19471 +#19475 := [monotonicity #19335 #19472]: #19474 +#19480 := [trans #19475 #19478]: #19479 +#19437 := (iff #19108 #19434) +#19431 := (and #12823 #12826 #13952 #14078 #14207 #19428) +#19435 := (iff #19431 #19434) +#19436 := [rewrite]: #19435 +#19432 := (iff #19108 #19431) +#19354 := (iff #19026 #12826) +#19355 := [rewrite]: #19354 +#19352 := (iff #19023 #12823) +#19353 := [rewrite]: #19352 +#19433 := [monotonicity #19353 #19355 #19335 #19357 #14297 #19430]: #19432 +#19438 := [trans #19433 #19436]: #19437 +#19483 := [monotonicity #19438 #19480]: #19482 +#19486 := [monotonicity #19351 #19335 #19483]: #19485 +#19491 := [trans #19486 #19489]: #19490 +#19494 := [monotonicity #19491]: #19493 +#19497 := [monotonicity #19349 #19494]: #19496 +#19502 := [trans #19497 #19500]: #19501 +#19505 := [monotonicity #19502]: #19504 +#19508 := [monotonicity #19335 #19505]: #19507 +#19513 := [trans #19508 #19511]: #19512 +#19578 := [monotonicity #19513 #19575]: #19577 +#19346 := (iff #19003 #16357) +#19347 := [rewrite]: #19346 +#19344 := (iff #19000 #16337) +#19345 := [rewrite]: #19344 +#19342 := (iff #18997 #16318) +#19343 := [rewrite]: #19342 +#19340 := (iff #18987 #14502) +#19341 := [rewrite]: #19340 +#19338 := (iff #18984 #14453) +#19339 := [rewrite]: #19338 +#19336 := (iff #18981 #14411) +#19337 := [rewrite]: #19336 +#19332 := (iff #18975 #12553) +#19333 := [rewrite]: #19332 +#19330 := (iff #18972 #12437) +#19331 := [rewrite]: #19330 +#19328 := (iff #18969 #3127) +#19329 := [rewrite]: #19328 +#19326 := (iff #18966 #3124) +#19327 := [rewrite]: #19326 +#19324 := (iff #18963 #3123) +#19325 := [rewrite]: #19324 +#19322 := (iff #18960 #3122) +#19323 := [rewrite]: #19322 +#19320 := (iff #18957 #3121) +#19321 := [rewrite]: #19320 +#19581 := [monotonicity #19321 #19323 #19325 #19327 #19329 #19331 #19333 #19335 #19337 #19339 #19341 #19343 #19345 #19347 #19578]: #19580 +#19586 := [trans #19581 #19584]: #19585 +#19589 := [monotonicity #14668 #19586]: #19588 +#19592 := [monotonicity #19589]: #19591 +#19318 := (iff #18941 #19317) +#19315 := (iff #18940 #19312) +#19309 := (or #18939 #18938 #19306) +#19313 := (iff #19309 #19312) +#19314 := [rewrite]: #19313 +#19310 := (iff #18940 #19309) +#19307 := (iff #18933 #19306) +#19304 := (iff #18932 #19303) +#19305 := [rewrite]: #19304 +#19308 := [monotonicity #19305]: #19307 +#19311 := [monotonicity #19308]: #19310 +#19316 := [trans #19311 #19314]: #19315 +#19319 := [monotonicity #19316]: #19318 +#19595 := [monotonicity #19319 #19592]: #19594 +#19301 := (iff #18926 #13903) +#19302 := [rewrite]: #19301 +#19598 := [monotonicity #19302 #19595]: #19597 +#19601 := [monotonicity #19598]: #19600 +#19299 := (iff #18921 #12373) +#19300 := [rewrite]: #19299 +#19297 := (iff #18918 #3031) +#19298 := [rewrite]: #19297 +#19295 := (iff #18915 #3028) +#19296 := [rewrite]: #19295 +#19293 := (iff #18912 #3025) +#19294 := [rewrite]: #19293 +#19291 := (iff #18909 #3022) +#19292 := [rewrite]: #19291 +#19604 := [monotonicity #19292 #19294 #19296 #19298 #19300 #19601]: #19603 +#19609 := [trans #19604 #19607]: #19608 +#19612 := [monotonicity #19609]: #19611 +#19289 := (iff #18903 #12364) +#19290 := [rewrite]: #19289 +#19615 := [monotonicity #19290 #19612]: #19614 +#19620 := [trans #19615 #19618]: #19619 +#19623 := [monotonicity #19620]: #19622 +#19287 := (iff #18894 #12355) +#19288 := [rewrite]: #19287 +#19626 := [monotonicity #19288 #19623]: #19625 +#19629 := [monotonicity #19626]: #19628 +#16494 := (exists (vars (?x782 int)) #16489) +#16483 := (not #16480) +#16497 := (or #16483 #16494) +#16500 := (and #16480 #16497) +#16506 := (or #12671 #12662 #12653 #12644 #13955 #14049 #16500) +#16384 := (not #16381) +#16390 := (or #14151 #16384) +#16395 := (and #16381 #16390) +#16398 := (or #14105 #16395) +#16401 := (and #14100 #16398) +#16418 := (or #12942 #14097 #14172 #16401 #16412) +#16426 := (and #14088 #16368 #16418) +#16439 := (or #13124 #13115 #13090 #13142 #13133 #13081 #14243 #14081 #16426) +#16442 := (and #12806 #12812 #16439) +#16445 := (or #13159 #16442) +#16448 := (and #12803 #12806 #16445) +#16451 := (or #13955 #14207 #16448) +#16434 := (or #13002 #12993 #13955 #14081 #14211 #16426) +#16454 := (and #16434 #16451) +#16457 := (or #13142 #13955 #16454) +#16460 := (and #12806 #12812 #16457) +#16463 := (or #13159 #16460) +#16466 := (and #12803 #12806 #16463) +#16469 := (or #13955 #14046 #16466) +#16511 := (and #16469 #16506) +#16301 := (not #16298) +#16517 := (or #13367 #13358 #13349 #13340 #13331 #14399 #15709 #13955 #14416 #14456 #14507 #16301 #16323 #16340 #16362 #16511) +#16522 := (and #3022 #13943 #16517) +#16287 := (not #16284) +#16525 := (or #16287 #16522) +#16528 := (and #16284 #16525) +#16531 := (or #13906 #16528) +#16534 := (and #13903 #16531) +#16537 := (or #13672 #13663 #13654 #13645 #13681 #16534) +#16540 := (and #12361 #12367 #16537) +#16543 := (or #13698 #16540) +#16546 := (and #12358 #12361 #16543) +#16549 := (or #13715 #16546) +#16552 := (and #12355 #16549) +#16555 := (not #16552) +#19282 := (~ #16555 #19281) +#19278 := (not #16549) +#19279 := (~ #19278 #19277) +#19274 := (not #16546) +#19275 := (~ #19274 #19273) +#19270 := (not #16543) +#19271 := (~ #19270 #19269) +#19266 := (not #16540) +#19267 := (~ #19266 #19265) +#19262 := (not #16537) +#19263 := (~ #19262 #19261) +#19258 := (not #16534) +#19259 := (~ #19258 #19257) +#19254 := (not #16531) +#19255 := (~ #19254 #19253) +#19250 := (not #16528) +#19251 := (~ #19250 #19249) +#19246 := (not #16525) +#19247 := (~ #19246 #19245) +#19242 := (not #16522) +#19243 := (~ #19242 #19241) +#19238 := (not #16517) +#19239 := (~ #19238 #19237) +#19234 := (not #16511) +#19235 := (~ #19234 #19233) +#19230 := (not #16506) +#19231 := (~ #19230 #19229) +#19226 := (not #16500) +#19227 := (~ #19226 #19225) +#19222 := (not #16497) +#19223 := (~ #19222 #19221) +#19218 := (not #16494) +#19219 := (~ #19218 #19217) +#19215 := (~ #19214 #19214) +#19216 := [refl]: #19215 +#19220 := [nnf-neg #19216]: #19219 +#19211 := (not #16483) +#19212 := (~ #19211 #16480) +#19209 := (~ #16480 #16480) +#19207 := (~ #16475 #16475) +#19208 := [refl]: #19207 +#19210 := [nnf-pos #19208]: #19209 +#19213 := [nnf-neg #19210]: #19212 +#19224 := [nnf-neg #19213 #19220]: #19223 +#19203 := (~ #16483 #19202) +#19204 := [sk]: #19203 +#19228 := [nnf-neg #19204 #19224]: #19227 +#19188 := (~ #14352 #14352) +#19189 := [refl]: #19188 +#18979 := (~ #18978 #18978) +#18980 := [refl]: #18979 +#19186 := (~ #19185 #19185) +#19187 := [refl]: #19186 +#19183 := (~ #19182 #19182) +#19184 := [refl]: #19183 +#19180 := (~ #19179 #19179) +#19181 := [refl]: #19180 +#19177 := (~ #19176 #19176) +#19178 := [refl]: #19177 +#19232 := [nnf-neg #19178 #19181 #19184 #19187 #18980 #19189 #19228]: #19231 +#19173 := (not #16469) +#19174 := (~ #19173 #19172) +#19169 := (not #16466) +#19170 := (~ #19169 #19168) +#19165 := (not #16463) +#19166 := (~ #19165 #19164) +#19161 := (not #16460) +#19162 := (~ #19161 #19160) +#19157 := (not #16457) +#19158 := (~ #19157 #19156) +#19153 := (not #16454) +#19154 := (~ #19153 #19152) +#19149 := (not #16451) +#19150 := (~ #19149 #19148) +#19145 := (not #16448) +#19146 := (~ #19145 #19144) +#19141 := (not #16445) +#19142 := (~ #19141 #19140) +#19137 := (not #16442) +#19138 := (~ #19137 #19136) +#19133 := (not #16439) +#19134 := (~ #19133 #19132) +#19105 := (not #16426) +#19106 := (~ #19105 #19104) +#19101 := (not #16418) +#19102 := (~ #19101 #19100) +#19098 := (~ #19097 #19097) +#19099 := [refl]: #19098 +#19094 := (not #16401) +#19095 := (~ #19094 #19093) +#19090 := (not #16398) +#19091 := (~ #19090 #19089) +#19086 := (not #16395) +#19087 := (~ #19086 #19085) +#19082 := (not #16390) +#19083 := (~ #19082 #19081) +#19078 := (not #16384) +#19079 := (~ #19078 #16381) +#19076 := (~ #16381 #16381) +#19074 := (~ #16376 #16376) +#19075 := [refl]: #19074 +#19077 := [nnf-pos #19075]: #19076 +#19080 := [nnf-neg #19077]: #19079 +#19072 := (~ #19071 #19071) +#19073 := [refl]: #19072 +#19084 := [nnf-neg #19073 #19080]: #19083 +#19067 := (~ #16384 #19066) +#19068 := [sk]: #19067 +#19088 := [nnf-neg #19068 #19084]: #19087 +#19052 := (~ #19051 #19051) +#19053 := [refl]: #19052 +#19092 := [nnf-neg #19053 #19088]: #19091 +#19049 := (~ #14105 #14105) +#19050 := [refl]: #19049 +#19096 := [nnf-neg #19050 #19092]: #19095 +#19047 := (~ #19046 #19046) +#19048 := [refl]: #19047 +#19044 := (~ #19043 #19043) +#19045 := [refl]: #19044 +#19041 := (~ #19040 #19040) +#19042 := [refl]: #19041 +#19103 := [nnf-neg #19042 #19045 #19048 #19096 #19099]: #19102 +#19038 := (~ #19037 #19037) +#19039 := [refl]: #19038 +#19035 := (~ #19034 #19034) +#19036 := [refl]: #19035 +#19107 := [nnf-neg #19036 #19039 #19103]: #19106 +#19030 := (~ #19029 #19029) +#19031 := [refl]: #19030 +#19130 := (~ #19129 #19129) +#19131 := [refl]: #19130 +#19127 := (~ #19126 #19126) +#19128 := [refl]: #19127 +#19124 := (~ #19123 #19123) +#19125 := [refl]: #19124 +#19021 := (~ #19020 #19020) +#19022 := [refl]: #19021 +#19121 := (~ #19120 #19120) +#19122 := [refl]: #19121 +#19118 := (~ #19117 #19117) +#19119 := [refl]: #19118 +#19115 := (~ #19114 #19114) +#19116 := [refl]: #19115 +#19135 := [nnf-neg #19116 #19119 #19122 #19022 #19125 #19128 #19131 #19031 #19107]: #19134 +#19018 := (~ #19017 #19017) +#19019 := [refl]: #19018 +#19012 := (~ #19011 #19011) +#19013 := [refl]: #19012 +#19139 := [nnf-neg #19013 #19019 #19135]: #19138 +#19015 := (~ #19014 #19014) +#19016 := [refl]: #19015 +#19143 := [nnf-neg #19016 #19139]: #19142 +#19009 := (~ #19008 #19008) +#19010 := [refl]: #19009 +#19147 := [nnf-neg #19010 #19013 #19143]: #19146 +#19112 := (~ #14211 #14211) +#19113 := [refl]: #19112 +#19151 := [nnf-neg #18980 #19113 #19147]: #19150 +#19109 := (not #16434) +#19110 := (~ #19109 #19108) +#19032 := (~ #14293 #14293) +#19033 := [refl]: #19032 +#19027 := (~ #19026 #19026) +#19028 := [refl]: #19027 +#19024 := (~ #19023 #19023) +#19025 := [refl]: #19024 +#19111 := [nnf-neg #19025 #19028 #18980 #19031 #19033 #19107]: #19110 +#19155 := [nnf-neg #19111 #19151]: #19154 +#19159 := [nnf-neg #19022 #18980 #19155]: #19158 +#19163 := [nnf-neg #19013 #19019 #19159]: #19162 +#19167 := [nnf-neg #19016 #19163]: #19166 +#19171 := [nnf-neg #19010 #19013 #19167]: #19170 +#19006 := (~ #14049 #14049) +#19007 := [refl]: #19006 +#19175 := [nnf-neg #18980 #19007 #19171]: #19174 +#19236 := [nnf-neg #19175 #19232]: #19235 +#19004 := (~ #19003 #19003) +#19005 := [refl]: #19004 +#19001 := (~ #19000 #19000) +#19002 := [refl]: #19001 +#18998 := (~ #18997 #18997) +#18999 := [refl]: #18998 +#18994 := (not #16301) +#18995 := (~ #18994 #16298) +#18992 := (~ #16298 #16298) +#18990 := (~ #16293 #16293) +#18991 := [refl]: #18990 +#18993 := [nnf-pos #18991]: #18992 +#18996 := [nnf-neg #18993]: #18995 +#18988 := (~ #18987 #18987) +#18989 := [refl]: #18988 +#18985 := (~ #18984 #18984) +#18986 := [refl]: #18985 +#18982 := (~ #18981 #18981) +#18983 := [refl]: #18982 +#18976 := (~ #18975 #18975) +#18977 := [refl]: #18976 +#18973 := (~ #18972 #18972) +#18974 := [refl]: #18973 +#18970 := (~ #18969 #18969) +#18971 := [refl]: #18970 +#18967 := (~ #18966 #18966) +#18968 := [refl]: #18967 +#18964 := (~ #18963 #18963) +#18965 := [refl]: #18964 +#18961 := (~ #18960 #18960) +#18962 := [refl]: #18961 +#18958 := (~ #18957 #18957) +#18959 := [refl]: #18958 +#19240 := [nnf-neg #18959 #18962 #18965 #18968 #18971 #18974 #18977 #18980 #18983 #18986 #18989 #18996 #18999 #19002 #19005 #19236]: #19239 +#18955 := (~ #14664 #14664) +#18956 := [refl]: #18955 +#18953 := (~ #13672 #13672) +#18954 := [refl]: #18953 +#19244 := [nnf-neg #18954 #18956 #19240]: #19243 +#18950 := (not #16287) +#18951 := (~ #18950 #16284) +#18948 := (~ #16284 #16284) +#18946 := (~ #16279 #16279) +#18947 := [refl]: #18946 +#18949 := [nnf-pos #18947]: #18948 +#18952 := [nnf-neg #18949]: #18951 +#19248 := [nnf-neg #18952 #19244]: #19247 +#18942 := (~ #16287 #18941) +#18943 := [sk]: #18942 +#19252 := [nnf-neg #18943 #19248]: #19251 +#18927 := (~ #18926 #18926) +#18928 := [refl]: #18927 +#19256 := [nnf-neg #18928 #19252]: #19255 +#18924 := (~ #13906 #13906) +#18925 := [refl]: #18924 +#19260 := [nnf-neg #18925 #19256]: #19259 +#18922 := (~ #18921 #18921) +#18923 := [refl]: #18922 +#18919 := (~ #18918 #18918) +#18920 := [refl]: #18919 +#18916 := (~ #18915 #18915) +#18917 := [refl]: #18916 +#18913 := (~ #18912 #18912) +#18914 := [refl]: #18913 +#18910 := (~ #18909 #18909) +#18911 := [refl]: #18910 +#19264 := [nnf-neg #18911 #18914 #18917 #18920 #18923 #19260]: #19263 +#18907 := (~ #18906 #18906) +#18908 := [refl]: #18907 +#18901 := (~ #18900 #18900) +#18902 := [refl]: #18901 +#19268 := [nnf-neg #18902 #18908 #19264]: #19267 +#18904 := (~ #18903 #18903) +#18905 := [refl]: #18904 +#19272 := [nnf-neg #18905 #19268]: #19271 +#18898 := (~ #18897 #18897) +#18899 := [refl]: #18898 +#19276 := [nnf-neg #18899 #18902 #19272]: #19275 +#18895 := (~ #18894 #18894) +#18896 := [refl]: #18895 +#19280 := [nnf-neg #18896 #19276]: #19279 +#18892 := (~ #13715 #13715) +#18893 := [refl]: #18892 +#19283 := [nnf-neg #18893 #19280]: #19282 +#15734 := (or #12671 #12662 #12653 #12644 #13955 #14009 #14049) +#15742 := (and #14371 #15734) +#15750 := (or #13367 #13358 #13349 #13340 #13331 #14399 #15709 #13955 #14416 #14450 #14456 #14468 #14483 #14496 #14507 #15742) +#15755 := (and #3022 #13943 #15750) +#15758 := (or #13939 #15755) +#15761 := (and #13936 #15758) +#15764 := (or #13906 #15761) +#15767 := (and #13903 #15764) +#15770 := (or #13672 #13663 #13654 #13645 #13681 #15767) +#15773 := (and #12361 #12367 #15770) +#15776 := (or #13698 #15773) +#15779 := (and #12358 #12361 #15776) +#15782 := (or #13715 #15779) +#15785 := (and #12355 #15782) +#15788 := (not #15785) +#16556 := (iff #15788 #16555) +#16553 := (iff #15785 #16552) +#16550 := (iff #15782 #16549) +#16547 := (iff #15779 #16546) +#16544 := (iff #15776 #16543) +#16541 := (iff #15773 #16540) +#16538 := (iff #15770 #16537) +#16535 := (iff #15767 #16534) +#16532 := (iff #15764 #16531) +#16529 := (iff #15761 #16528) +#16526 := (iff #15758 #16525) +#16523 := (iff #15755 #16522) +#16520 := (iff #15750 #16517) +#16514 := (or #13367 #13358 #13349 #13340 #13331 #14399 #15709 #13955 #14416 #16301 #14456 #16323 #16340 #16362 #14507 #16511) +#16518 := (iff #16514 #16517) +#16519 := [rewrite]: #16518 +#16515 := (iff #15750 #16514) +#16512 := (iff #15742 #16511) +#16509 := (iff #15734 #16506) +#16503 := (or #12671 #12662 #12653 #12644 #13955 #16500 #14049) +#16507 := (iff #16503 #16506) +#16508 := [rewrite]: #16507 +#16504 := (iff #15734 #16503) +#16501 := (iff #14009 #16500) +#16498 := (iff #14006 #16497) +#16495 := (iff #14003 #16494) +#16492 := (iff #13998 #16489) +#16486 := (and #3145 #4084 #15606 #13958) +#16490 := (iff #16486 #16489) +#16491 := [rewrite]: #16490 +#16487 := (iff #13998 #16486) +#15605 := (iff #4419 #15606) +#15638 := -131073::int +#15614 := (+ -131073::int #161) +#15611 := (<= #15614 0::int) +#15607 := (iff #15611 #15606) +#15604 := [rewrite]: #15607 +#15608 := (iff #4419 #15611) +#15613 := (= #4418 #15614) +#15619 := (+ #161 -131073::int) +#15615 := (= #15619 #15614) +#15612 := [rewrite]: #15615 +#15616 := (= #4418 #15619) +#15637 := (= #4413 -131073::int) +#15643 := (* -1::int 131073::int) +#15639 := (= #15643 -131073::int) +#15636 := [rewrite]: #15639 +#15640 := (= #4413 #15643) +#7883 := (= uf_76 131073::int) +#1070 := 65536::int +#1313 := (+ 65536::int 65536::int) +#1318 := (+ #1313 1::int) +#1319 := (= uf_76 #1318) +#7884 := (iff #1319 #7883) +#7881 := (= #1318 131073::int) +#7874 := (+ 131072::int 1::int) +#7879 := (= #7874 131073::int) +#7880 := [rewrite]: #7879 +#7876 := (= #1318 #7874) +#7845 := (= #1313 131072::int) +#7846 := [rewrite]: #7845 +#7877 := [monotonicity #7846]: #7876 +#7882 := [trans #7877 #7880]: #7881 +#7885 := [monotonicity #7882]: #7884 +#7873 := [asserted]: #1319 +#7888 := [mp #7873 #7885]: #7883 +#15641 := [monotonicity #7888]: #15640 +#15634 := [trans #15641 #15636]: #15637 +#15617 := [monotonicity #15634]: #15616 +#15610 := [trans #15617 #15612]: #15613 +#15609 := [monotonicity #15610]: #15608 +#15602 := [trans #15609 #15604]: #15605 +#16488 := [monotonicity #15602]: #16487 +#16493 := [trans #16488 #16491]: #16492 +#16496 := [quant-intro #16493]: #16495 +#16484 := (iff #13989 #16483) +#16481 := (iff #13986 #16480) +#16478 := (iff #13981 #16475) +#16472 := (or #14857 #13959 #13972) +#16476 := (iff #16472 #16475) +#16477 := [rewrite]: #16476 +#16473 := (iff #13981 #16472) +#14854 := (iff #5739 #14857) +#14859 := (iff #5736 #14858) +#14856 := [monotonicity #15602]: #14859 +#14855 := [monotonicity #14856]: #14854 +#16474 := [monotonicity #14855]: #16473 +#16479 := [trans #16474 #16477]: #16478 +#16482 := [quant-intro #16479]: #16481 +#16485 := [monotonicity #16482]: #16484 +#16499 := [monotonicity #16485 #16496]: #16498 +#16502 := [monotonicity #16482 #16499]: #16501 +#16505 := [monotonicity #16502]: #16504 +#16510 := [trans #16505 #16508]: #16509 +#16470 := (iff #14371 #16469) +#16467 := (iff #14345 #16466) +#16464 := (iff #14339 #16463) +#16461 := (iff #14334 #16460) +#16458 := (iff #14326 #16457) +#16455 := (iff #14317 #16454) +#16452 := (iff #14312 #16451) +#16449 := (iff #14286 #16448) +#16446 := (iff #14280 #16445) +#16443 := (iff #14275 #16442) +#16440 := (iff #14267 #16439) +#16429 := (iff #14201 #16426) +#16423 := (and #16368 #14088 #16418) +#16427 := (iff #16423 #16426) +#16428 := [rewrite]: #16427 +#16424 := (iff #14201 #16423) +#16421 := (iff #14193 #16418) +#16415 := (or #12942 #14097 #16401 #14172 #16412) +#16419 := (iff #16415 #16418) +#16420 := [rewrite]: #16419 +#16416 := (iff #14193 #16415) +#16413 := (iff #14178 #16412) +#16410 := (iff #14175 #16407) +#16404 := (and #16368 #14088) +#16408 := (iff #16404 #16407) +#16409 := [rewrite]: #16408 +#16405 := (iff #14175 #16404) +#16371 := (iff #14084 #16368) +#16304 := (+ 131073::int #14044) +#16365 := (>= #16304 1::int) +#16369 := (iff #16365 #16368) +#16370 := [rewrite]: #16369 +#16366 := (iff #14084 #16365) +#16305 := (= #14085 #16304) +#16306 := [monotonicity #7888]: #16305 +#16367 := [monotonicity #16306]: #16366 +#16372 := [trans #16367 #16370]: #16371 +#16406 := [monotonicity #16372]: #16405 +#16411 := [trans #16406 #16409]: #16410 +#16414 := [monotonicity #16411]: #16413 +#16402 := (iff #14165 #16401) +#16399 := (iff #14162 #16398) +#16396 := (iff #14159 #16395) +#16393 := (iff #14156 #16390) +#16387 := (or #16384 #14151) +#16391 := (iff #16387 #16390) +#16392 := [rewrite]: #16391 +#16388 := (iff #14156 #16387) +#16385 := (iff #14139 #16384) +#16382 := (iff #14136 #16381) +#16379 := (iff #14131 #16376) +#16373 := (or #14857 #14109 #14122) +#16377 := (iff #16373 #16376) +#16378 := [rewrite]: #16377 +#16374 := (iff #14131 #16373) +#16375 := [monotonicity #14855]: #16374 +#16380 := [trans #16375 #16378]: #16379 +#16383 := [quant-intro #16380]: #16382 +#16386 := [monotonicity #16383]: #16385 +#16389 := [monotonicity #16386]: #16388 +#16394 := [trans #16389 #16392]: #16393 +#16397 := [monotonicity #16383 #16394]: #16396 +#16400 := [monotonicity #16397]: #16399 +#16403 := [monotonicity #16400]: #16402 +#16417 := [monotonicity #16403 #16414]: #16416 +#16422 := [trans #16417 #16420]: #16421 +#16425 := [monotonicity #16372 #16422]: #16424 +#16430 := [trans #16425 #16428]: #16429 +#16441 := [monotonicity #16430]: #16440 +#16444 := [monotonicity #16441]: #16443 +#16447 := [monotonicity #16444]: #16446 +#16450 := [monotonicity #16447]: #16449 +#16453 := [monotonicity #16450]: #16452 +#16437 := (iff #14238 #16434) +#16431 := (or #13002 #12993 #13955 #14081 #16426 #14211) +#16435 := (iff #16431 #16434) +#16436 := [rewrite]: #16435 +#16432 := (iff #14238 #16431) +#16433 := [monotonicity #16430]: #16432 +#16438 := [trans #16433 #16436]: #16437 +#16456 := [monotonicity #16438 #16453]: #16455 +#16459 := [monotonicity #16456]: #16458 +#16462 := [monotonicity #16459]: #16461 +#16465 := [monotonicity #16462]: #16464 +#16468 := [monotonicity #16465]: #16467 +#16471 := [monotonicity #16468]: #16470 +#16513 := [monotonicity #16471 #16510]: #16512 +#16363 := (iff #14496 #16362) +#16360 := (iff #14493 #16357) +#16354 := (and #16349 #14490) +#16358 := (iff #16354 #16357) +#16359 := [rewrite]: #16358 +#16355 := (iff #14493 #16354) +#16352 := (iff #14486 #16349) +#16343 := (+ 255::int #14431) +#16346 := (>= #16343 0::int) +#16350 := (iff #16346 #16349) +#16351 := [rewrite]: #16350 +#16347 := (iff #14486 #16346) +#16344 := (= #14487 #16343) +#1323 := (= uf_78 255::int) +#7887 := [asserted]: #1323 +#16345 := [monotonicity #7887]: #16344 +#16348 := [monotonicity #16345]: #16347 +#16353 := [trans #16348 #16351]: #16352 +#16356 := [monotonicity #16353]: #16355 +#16361 := [trans #16356 #16359]: #16360 +#16364 := [monotonicity #16361]: #16363 +#16341 := (iff #14483 #16340) +#16338 := (iff #14478 #16337) +#16335 := (iff #14471 #16332) +#16326 := (+ 131073::int #14402) +#16329 := (>= #16326 0::int) +#16333 := (iff #16329 #16332) +#16334 := [rewrite]: #16333 +#16330 := (iff #14471 #16329) +#16327 := (= #14472 #16326) +#16328 := [monotonicity #7888]: #16327 +#16331 := [monotonicity #16328]: #16330 +#16336 := [trans #16331 #16334]: #16335 +#16339 := [monotonicity #16336]: #16338 +#16342 := [monotonicity #16339]: #16341 +#16324 := (iff #14468 #16323) +#16321 := (iff #14465 #16318) +#16315 := (and #16310 #14462) +#16319 := (iff #16315 #16318) +#16320 := [rewrite]: #16319 +#16316 := (iff #14465 #16315) +#16313 := (iff #14459 #16310) +#16307 := (>= #16304 0::int) +#16311 := (iff #16307 #16310) +#16312 := [rewrite]: #16311 +#16308 := (iff #14459 #16307) +#16309 := [monotonicity #16306]: #16308 +#16314 := [trans #16309 #16312]: #16313 +#16317 := [monotonicity #16314]: #16316 +#16322 := [trans #16317 #16320]: #16321 +#16325 := [monotonicity #16322]: #16324 +#16302 := (iff #14450 #16301) +#16299 := (iff #14447 #16298) +#16296 := (iff #14442 #16293) +#16290 := (or #14857 #14420 #14433) +#16294 := (iff #16290 #16293) +#16295 := [rewrite]: #16294 +#16291 := (iff #14442 #16290) +#16292 := [monotonicity #14855]: #16291 +#16297 := [trans #16292 #16295]: #16296 +#16300 := [quant-intro #16297]: #16299 +#16303 := [monotonicity #16300]: #16302 +#16516 := [monotonicity #16303 #16325 #16342 #16364 #16513]: #16515 +#16521 := [trans #16516 #16519]: #16520 +#16524 := [monotonicity #16521]: #16523 +#16288 := (iff #13939 #16287) +#16285 := (iff #13936 #16284) +#16282 := (iff #13931 #16279) +#16276 := (or #14857 #13910 #13921) +#16280 := (iff #16276 #16279) +#16281 := [rewrite]: #16280 +#16277 := (iff #13931 #16276) +#16278 := [monotonicity #14855]: #16277 +#16283 := [trans #16278 #16281]: #16282 +#16286 := [quant-intro #16283]: #16285 +#16289 := [monotonicity #16286]: #16288 +#16527 := [monotonicity #16289 #16524]: #16526 +#16530 := [monotonicity #16286 #16527]: #16529 +#16533 := [monotonicity #16530]: #16532 +#16536 := [monotonicity #16533]: #16535 +#16539 := [monotonicity #16536]: #16538 +#16542 := [monotonicity #16539]: #16541 +#16545 := [monotonicity #16542]: #16544 +#16548 := [monotonicity #16545]: #16547 +#16551 := [monotonicity #16548]: #16550 +#16554 := [monotonicity #16551]: #16553 +#16557 := [monotonicity #16554]: #16556 +#14791 := (not #14643) +#15789 := (iff #14791 #15788) +#15786 := (iff #14643 #15785) +#15783 := (iff #14640 #15782) +#15780 := (iff #14635 #15779) +#15777 := (iff #14629 #15776) +#15774 := (iff #14624 #15773) +#15771 := (iff #14616 #15770) +#15768 := (iff #14595 #15767) +#15765 := (iff #14592 #15764) +#15762 := (iff #14589 #15761) +#15759 := (iff #14586 #15758) +#15756 := (iff #14581 #15755) +#15753 := (iff #14573 #15750) +#15747 := (or #13367 #13358 #13349 #13340 #13331 #14399 #15709 #13955 #15742 #14416 #14450 #14456 #14468 #14483 #14496 #14507) +#15751 := (iff #15747 #15750) +#15752 := [rewrite]: #15751 +#15748 := (iff #14573 #15747) +#15745 := (iff #14376 #15742) +#15739 := (and #15734 #14371) +#15743 := (iff #15739 #15742) +#15744 := [rewrite]: #15743 +#15740 := (iff #14376 #15739) +#15737 := (iff #14070 #15734) +#15719 := (or #12671 #12662 #12653 #12644 #13955 #14009) +#15731 := (or #13955 #15719 #14049) +#15735 := (iff #15731 #15734) +#15736 := [rewrite]: #15735 +#15732 := (iff #14070 #15731) +#15729 := (iff #14041 #15719) +#15724 := (and true #15719) +#15727 := (iff #15724 #15719) +#15728 := [rewrite]: #15727 +#15725 := (iff #14041 #15724) +#15722 := (iff #14036 #15719) +#15716 := (or false #12671 #12662 #12653 #12644 #13955 #14009) +#15720 := (iff #15716 #15719) +#15721 := [rewrite]: #15720 +#15717 := (iff #14036 #15716) +#15714 := (iff #12719 false) +#15712 := (iff #12719 #3294) +#15456 := (iff up_216 true) +#11194 := [asserted]: up_216 +#15457 := [iff-true #11194]: #15456 +#15713 := [monotonicity #15457]: #15712 +#15715 := [trans #15713 #13445]: #15714 +#15718 := [monotonicity #15715]: #15717 +#15723 := [trans #15718 #15721]: #15722 +#15726 := [monotonicity #15457 #15723]: #15725 +#15730 := [trans #15726 #15728]: #15729 +#15733 := [monotonicity #15730]: #15732 +#15738 := [trans #15733 #15736]: #15737 +#15741 := [monotonicity #15738]: #15740 +#15746 := [trans #15741 #15744]: #15745 +#15710 := (iff #13376 #15709) +#15707 := (iff #12556 #12553) +#15702 := (and true #12553) +#15705 := (iff #15702 #12553) +#15706 := [rewrite]: #15705 +#15703 := (iff #12556 #15702) +#15690 := (iff #12332 true) +#15691 := [iff-true #14784]: #15690 +#15704 := [monotonicity #15691]: #15703 +#15708 := [trans #15704 #15706]: #15707 +#15711 := [monotonicity #15708]: #15710 +#15749 := [monotonicity #15711 #15746]: #15748 +#15754 := [trans #15749 #15752]: #15753 +#15757 := [monotonicity #15754]: #15756 +#15760 := [monotonicity #15757]: #15759 +#15763 := [monotonicity #15760]: #15762 +#15766 := [monotonicity #15763]: #15765 +#15769 := [monotonicity #15766]: #15768 +#15772 := [monotonicity #15769]: #15771 +#15775 := [monotonicity #15772]: #15774 +#15778 := [monotonicity #15775]: #15777 +#15781 := [monotonicity #15778]: #15780 +#15784 := [monotonicity #15781]: #15783 +#15787 := [monotonicity #15784]: #15786 +#15790 := [monotonicity #15787]: #15789 +#14792 := [not-or-elim #14776]: #14791 +#15791 := [mp #14792 #15790]: #15788 +#16558 := [mp #15791 #16557]: #16555 +#19284 := [mp~ #16558 #19283]: #19281 +#19285 := [mp #19284 #19629]: #19627 +#23120 := [mp #19285 #23119]: #23117 +#23987 := [mp #23120 #23986]: #23984 +#28241 := [unit-resolution #23987 #26494]: #23981 +#28348 := (or #23978 #23957) +decl uf_136 :: (-> T14 T5) +#26312 := (uf_58 #3079 #3011) +#26553 := (uf_136 #26312) +#26565 := (uf_24 uf_273 #26553) +#26566 := (= uf_9 #26565) +#26600 := (not #26566) +decl uf_135 :: (-> T14 T2) +#26546 := (uf_135 #26312) +#26551 := (= uf_9 #26546) +#26552 := (not #26551) +#26788 := (or #26552 #26600) +#26791 := (not #26788) +decl uf_210 :: (-> T4 T5 T2) +#26631 := (uf_210 uf_273 #26553) +#26632 := (= uf_9 #26631) +#26630 := (uf_25 uf_273 #26553) +#26610 := (= uf_26 #26630) +#26753 := (or #26610 #26632) +#26766 := (not #26753) +#26287 := (uf_15 #3011) +#26634 := (uf_14 #26287) +#26726 := (= uf_16 #26634) +#26750 := (not #26726) +#26608 := (uf_15 #26553) +#26609 := (uf_14 #26608) +#26629 := (= uf_16 #26609) +#26796 := (or #26629 #26750 #26766 #26791) +#26807 := (not #26796) +#26557 := (uf_25 uf_273 #3011) +#26558 := (= uf_26 #26557) +#26555 := (uf_210 uf_273 #3011) +#26556 := (= uf_9 #26555) +#26756 := (or #26556 #26558) +#26759 := (not #26756) +#26745 := (or #26726 #26759) +#26748 := (not #26745) +#26809 := (or #26748 #26807) +#26812 := (not #26809) +#26819 := (or #18897 #26812) +#26823 := (not #26819) +#26851 := (iff #12367 #26823) +#2376 := (uf_67 #47 #26) +#2377 := (pattern #2376) +#281 := (uf_59 #47) +#2383 := (uf_58 #281 #26) +#2397 := (uf_135 #2383) +#10938 := (= uf_9 #2397) +#10941 := (not #10938) +#2384 := (uf_136 #2383) +#2394 := (uf_24 #47 #2384) +#10932 := (= uf_9 #2394) +#10935 := (not #10932) +#10944 := (or #10935 #10941) +#22490 := (not #10944) +#2390 := (uf_15 #2384) +#2391 := (uf_14 #2390) +#10926 := (= uf_16 #2391) +#2387 := (uf_25 #47 #2384) +#10920 := (= uf_26 #2387) +#2385 := (uf_210 #47 #2384) +#10917 := (= uf_9 #2385) +#10923 := (or #10917 #10920) +#22489 := (not #10923) +#22491 := (or #52 #22489 #10926 #22490) +#22492 := (not #22491) +#2379 := (uf_210 #47 #26) +#10898 := (= uf_9 #2379) +#10904 := (or #3656 #10898) +#22484 := (not #10904) +#22485 := (or #36 #22484) +#22486 := (not #22485) +#22495 := (or #22486 #22492) +#22501 := (not #22495) +#22502 := (or #11522 #22501) +#22503 := (not #22502) +#10894 := (= uf_9 #2376) +#22508 := (iff #10894 #22503) +#22511 := (forall (vars (?x632 T4) (?x633 T5)) (:pat #2377) #22508) +#10929 := (not #10926) +#10978 := (and #36 #10923 #10929 #10944) +#10912 := (and #52 #10904) +#10981 := (or #10912 #10978) +#10984 := (and #3650 #10981) +#10987 := (iff #10894 #10984) +#10990 := (forall (vars (?x632 T4) (?x633 T5)) (:pat #2377) #10987) +#22512 := (iff #10990 #22511) +#22509 := (iff #10987 #22508) +#22506 := (iff #10984 #22503) +#22498 := (and #3650 #22495) +#22504 := (iff #22498 #22503) +#22505 := [rewrite]: #22504 +#22499 := (iff #10984 #22498) +#22496 := (iff #10981 #22495) +#22493 := (iff #10978 #22492) +#22494 := [rewrite]: #22493 +#22487 := (iff #10912 #22486) +#22488 := [rewrite]: #22487 +#22497 := [monotonicity #22488 #22494]: #22496 +#22500 := [monotonicity #22497]: #22499 +#22507 := [trans #22500 #22505]: #22506 +#22510 := [monotonicity #22507]: #22509 +#22513 := [quant-intro #22510]: #22512 +#18466 := (~ #10990 #10990) +#18464 := (~ #10987 #10987) +#18465 := [refl]: #18464 +#18467 := [nnf-pos #18465]: #18466 +#2398 := (= #2397 uf_9) +#2399 := (not #2398) +#2395 := (= #2394 uf_9) +#2396 := (not #2395) +#2400 := (or #2396 #2399) +#2401 := (and #2400 #36) +#2392 := (= #2391 uf_16) +#2393 := (not #2392) +#2402 := (and #2393 #2401) +#2388 := (= #2387 uf_26) +#2386 := (= #2385 uf_9) +#2389 := (or #2386 #2388) +#2403 := (and #2389 #2402) +#2380 := (= #2379 uf_9) +#2381 := (or #2380 #151) +#2382 := (and #2381 #52) +#2404 := (or #2382 #2403) +#2405 := (and #2404 #147) +#2378 := (= #2376 uf_9) +#2406 := (iff #2378 #2405) +#2407 := (forall (vars (?x632 T4) (?x633 T5)) (:pat #2377) #2406) +#10993 := (iff #2407 #10990) +#10950 := (and #36 #10944) +#10955 := (and #10929 #10950) +#10958 := (and #10923 #10955) +#10961 := (or #10912 #10958) +#10967 := (and #3650 #10961) +#10972 := (iff #10894 #10967) +#10975 := (forall (vars (?x632 T4) (?x633 T5)) (:pat #2377) #10972) +#10991 := (iff #10975 #10990) +#10988 := (iff #10972 #10987) +#10985 := (iff #10967 #10984) +#10982 := (iff #10961 #10981) +#10979 := (iff #10958 #10978) +#10980 := [rewrite]: #10979 +#10983 := [monotonicity #10980]: #10982 +#10986 := [monotonicity #10983]: #10985 +#10989 := [monotonicity #10986]: #10988 +#10992 := [quant-intro #10989]: #10991 +#10976 := (iff #2407 #10975) +#10973 := (iff #2406 #10972) +#10970 := (iff #2405 #10967) +#10964 := (and #10961 #3650) +#10968 := (iff #10964 #10967) +#10969 := [rewrite]: #10968 +#10965 := (iff #2405 #10964) +#10962 := (iff #2404 #10961) +#10959 := (iff #2403 #10958) +#10956 := (iff #2402 #10955) +#10953 := (iff #2401 #10950) +#10947 := (and #10944 #36) +#10951 := (iff #10947 #10950) +#10952 := [rewrite]: #10951 +#10948 := (iff #2401 #10947) +#10945 := (iff #2400 #10944) +#10942 := (iff #2399 #10941) +#10939 := (iff #2398 #10938) +#10940 := [rewrite]: #10939 +#10943 := [monotonicity #10940]: #10942 +#10936 := (iff #2396 #10935) +#10933 := (iff #2395 #10932) +#10934 := [rewrite]: #10933 +#10937 := [monotonicity #10934]: #10936 +#10946 := [monotonicity #10937 #10943]: #10945 +#10949 := [monotonicity #10946]: #10948 +#10954 := [trans #10949 #10952]: #10953 +#10930 := (iff #2393 #10929) +#10927 := (iff #2392 #10926) +#10928 := [rewrite]: #10927 +#10931 := [monotonicity #10928]: #10930 +#10957 := [monotonicity #10931 #10954]: #10956 +#10924 := (iff #2389 #10923) +#10921 := (iff #2388 #10920) +#10922 := [rewrite]: #10921 +#10918 := (iff #2386 #10917) +#10919 := [rewrite]: #10918 +#10925 := [monotonicity #10919 #10922]: #10924 +#10960 := [monotonicity #10925 #10957]: #10959 +#10915 := (iff #2382 #10912) +#10909 := (and #10904 #52) +#10913 := (iff #10909 #10912) +#10914 := [rewrite]: #10913 +#10910 := (iff #2382 #10909) +#10907 := (iff #2381 #10904) +#10901 := (or #10898 #3656) +#10905 := (iff #10901 #10904) +#10906 := [rewrite]: #10905 +#10902 := (iff #2381 #10901) +#10899 := (iff #2380 #10898) +#10900 := [rewrite]: #10899 +#10903 := [monotonicity #10900 #3658]: #10902 +#10908 := [trans #10903 #10906]: #10907 +#10911 := [monotonicity #10908]: #10910 +#10916 := [trans #10911 #10914]: #10915 +#10963 := [monotonicity #10916 #10960]: #10962 +#10966 := [monotonicity #10963 #3652]: #10965 +#10971 := [trans #10966 #10969]: #10970 +#10896 := (iff #2378 #10894) +#10897 := [rewrite]: #10896 +#10974 := [monotonicity #10897 #10971]: #10973 +#10977 := [quant-intro #10974]: #10976 +#10994 := [trans #10977 #10992]: #10993 +#10893 := [asserted]: #2407 +#10995 := [mp #10893 #10994]: #10990 +#18468 := [mp~ #10995 #18467]: #10990 +#22514 := [mp #18468 #22513]: #22511 +#26854 := (not #22511) +#26855 := (or #26854 #26851) +#26606 := (or #26600 #26552) +#26607 := (not #26606) +#26633 := (or #26632 #26610) +#26628 := (not #26633) +#26635 := (= #26634 uf_16) +#26684 := (not #26635) +#26685 := (or #26684 #26628 #26629 #26607) +#26554 := (not #26685) +#26559 := (or #26558 #26556) +#26560 := (not #26559) +#26544 := (or #26635 #26560) +#26636 := (not #26544) +#26637 := (or #26636 #26554) +#26681 := (not #26637) +#26713 := (or #18897 #26681) +#26714 := (not #26713) +#26725 := (iff #12367 #26714) +#26840 := (or #26854 #26725) +#26842 := (iff #26840 #26855) +#26844 := (iff #26855 #26855) +#26839 := [rewrite]: #26844 +#26852 := (iff #26725 #26851) +#26824 := (iff #26714 #26823) +#26820 := (iff #26713 #26819) +#26813 := (iff #26681 #26812) +#26810 := (iff #26637 #26809) +#26808 := (iff #26554 #26807) +#26805 := (iff #26685 #26796) +#26793 := (or #26750 #26766 #26629 #26791) +#26802 := (iff #26793 #26796) +#26804 := [rewrite]: #26802 +#26794 := (iff #26685 #26793) +#26786 := (iff #26607 #26791) +#26789 := (iff #26606 #26788) +#26790 := [rewrite]: #26789 +#26792 := [monotonicity #26790]: #26786 +#26785 := (iff #26628 #26766) +#26764 := (iff #26633 #26753) +#26765 := [rewrite]: #26764 +#26787 := [monotonicity #26765]: #26785 +#26751 := (iff #26684 #26750) +#26754 := (iff #26635 #26726) +#26755 := [rewrite]: #26754 +#26752 := [monotonicity #26755]: #26751 +#26795 := [monotonicity #26752 #26787 #26792]: #26794 +#26806 := [trans #26795 #26804]: #26805 +#26803 := [monotonicity #26806]: #26808 +#26743 := (iff #26636 #26748) +#26746 := (iff #26544 #26745) +#26742 := (iff #26560 #26759) +#26757 := (iff #26559 #26756) +#26758 := [rewrite]: #26757 +#26744 := [monotonicity #26758]: #26742 +#26747 := [monotonicity #26755 #26744]: #26746 +#26749 := [monotonicity #26747]: #26743 +#26811 := [monotonicity #26749 #26803]: #26810 +#26818 := [monotonicity #26811]: #26813 +#26822 := [monotonicity #26818]: #26820 +#26850 := [monotonicity #26822]: #26824 +#26853 := [monotonicity #26850]: #26852 +#26843 := [monotonicity #26853]: #26842 +#26845 := [trans #26843 #26839]: #26842 +#26841 := [quant-inst]: #26840 +#26846 := [mp #26841 #26845]: #26855 +#27857 := [unit-resolution #26846 #22514]: #26851 +#27023 := (not #26851) +#27960 := (or #27023 #26819) +#27858 := [hypothesis]: #23954 +decl uf_144 :: (-> T3 T3) +#24114 := (uf_144 #2952) +#26288 := (= #24114 #26287) +#26263 := (uf_48 #3011 #24114) +#26264 := (= uf_9 #26263) +#26290 := (iff #26264 #26288) +#26074 := (not #26290) +#26175 := [hypothesis]: #26074 +#1381 := (uf_15 #15) +#9506 := (= #233 #1381) +#11615 := (iff #9506 #11594) +#23676 := (forall (vars (?x712 T5) (?x713 T3)) (:pat #2662) #11615) +#11620 := (forall (vars (?x712 T5) (?x713 T3)) #11615) +#23679 := (iff #11620 #23676) +#23677 := (iff #11615 #11615) +#23678 := [refl]: #23677 +#23680 := [quant-intro #23678]: #23679 +#18739 := (~ #11620 #11620) +#18737 := (~ #11615 #11615) +#18738 := [refl]: #18737 +#18740 := [nnf-pos #18738]: #18739 +#1882 := (= #1381 #233) +#2668 := (iff #2663 #1882) +#2669 := (forall (vars (?x712 T5) (?x713 T3)) #2668) +#11621 := (iff #2669 #11620) +#11618 := (iff #2668 #11615) +#11611 := (iff #11594 #9506) +#11616 := (iff #11611 #11615) +#11617 := [rewrite]: #11616 +#11613 := (iff #2668 #11611) +#9507 := (iff #1882 #9506) +#9508 := [rewrite]: #9507 +#11614 := [monotonicity #11597 #9508]: #11613 +#11619 := [trans #11614 #11617]: #11618 +#11622 := [quant-intro #11619]: #11621 +#11610 := [asserted]: #2669 +#11625 := [mp #11610 #11622]: #11620 +#18741 := [mp~ #11625 #18740]: #11620 +#23681 := [mp #18741 #23680]: #23676 +#25432 := (not #23676) +#26067 := (or #25432 #26290) +#26289 := (iff #26288 #26264) +#26068 := (or #25432 #26289) +#26069 := (iff #26068 #26067) +#26065 := (iff #26067 #26067) +#26071 := [rewrite]: #26065 +#26291 := (iff #26289 #26290) +#26292 := [rewrite]: #26291 +#26070 := [monotonicity #26292]: #26069 +#26072 := [trans #26070 #26071]: #26069 +#26066 := [quant-inst]: #26068 +#26073 := [mp #26066 #26072]: #26067 +#26176 := [unit-resolution #26073 #23681 #26175]: false +#26214 := [lemma #26176]: #26290 +#26294 := (or #26074 #12361) +#26357 := (uf_116 #23223) +decl uf_138 :: (-> T3 int) +#26356 := (uf_138 #24114) +#26365 := (+ #26356 #26357) +#26368 := (uf_43 #24114 #26365) +#26561 := (uf_15 #26368) +#26308 := (= #26561 #26287) +#26304 := (= #26287 #26561) +#26302 := (= #3011 #26368) +#26346 := (uf_66 #23223 0::int #24114) +#26371 := (= #26346 #26368) +#26374 := (not #26371) +decl uf_139 :: (-> T5 T5 T2) +#26347 := (uf_139 #26346 #23223) +#26354 := (= uf_9 #26347) +#26355 := (not #26354) +#26380 := (or #26355 #26374) +#26385 := (not #26380) +#247 := (:var 1 int) +#1568 := (uf_66 #24 #247 #233) +#1569 := (pattern #1568) +#1576 := (uf_139 #1568 #24) +#8688 := (= uf_9 #1576) +#21652 := (not #8688) +#1571 := (uf_138 #233) +#1570 := (uf_116 #24) +#8678 := (+ #1570 #1571) +#8679 := (+ #247 #8678) +#8682 := (uf_43 #233 #8679) +#8685 := (= #1568 #8682) +#21651 := (not #8685) +#21653 := (or #21651 #21652) +#21654 := (not #21653) +#21657 := (forall (vars (?x375 T5) (?x376 int) (?x377 T3)) (:pat #1569) #21654) +#8691 := (and #8685 #8688) +#8694 := (forall (vars (?x375 T5) (?x376 int) (?x377 T3)) (:pat #1569) #8691) +#21658 := (iff #8694 #21657) +#21655 := (iff #8691 #21654) +#21656 := [rewrite]: #21655 +#21659 := [quant-intro #21656]: #21658 +#17817 := (~ #8694 #8694) +#17815 := (~ #8691 #8691) +#17816 := [refl]: #17815 +#17818 := [nnf-pos #17816]: #17817 +#1577 := (= #1576 uf_9) +#1572 := (+ #247 #1571) +#1573 := (+ #1570 #1572) +#1574 := (uf_43 #233 #1573) +#1575 := (= #1568 #1574) +#1578 := (and #1575 #1577) +#1579 := (forall (vars (?x375 T5) (?x376 int) (?x377 T3)) (:pat #1569) #1578) +#8695 := (iff #1579 #8694) +#8692 := (iff #1578 #8691) +#8689 := (iff #1577 #8688) +#8690 := [rewrite]: #8689 +#8686 := (iff #1575 #8685) +#8683 := (= #1574 #8682) +#8680 := (= #1573 #8679) +#8681 := [rewrite]: #8680 +#8684 := [monotonicity #8681]: #8683 +#8687 := [monotonicity #8684]: #8686 +#8693 := [monotonicity #8687 #8690]: #8692 +#8696 := [quant-intro #8693]: #8695 +#8677 := [asserted]: #1579 +#8699 := [mp #8677 #8696]: #8694 +#17819 := [mp~ #8699 #17818]: #8694 +#21660 := [mp #17819 #21659]: #21657 +#26114 := (not #21657) +#26115 := (or #26114 #26385) +#26358 := (+ #26357 #26356) +#26359 := (+ 0::int #26358) +#26360 := (uf_43 #24114 #26359) +#26361 := (= #26346 #26360) +#26362 := (not #26361) +#26363 := (or #26362 #26355) +#26364 := (not #26363) +#26116 := (or #26114 #26364) +#26122 := (iff #26116 #26115) +#26125 := (iff #26115 #26115) +#26126 := [rewrite]: #26125 +#26386 := (iff #26364 #26385) +#26383 := (iff #26363 #26380) +#26377 := (or #26374 #26355) +#26381 := (iff #26377 #26380) +#26382 := [rewrite]: #26381 +#26378 := (iff #26363 #26377) +#26375 := (iff #26362 #26374) +#26372 := (iff #26361 #26371) +#26369 := (= #26360 #26368) +#26366 := (= #26359 #26365) +#26367 := [rewrite]: #26366 +#26370 := [monotonicity #26367]: #26369 +#26373 := [monotonicity #26370]: #26372 +#26376 := [monotonicity #26373]: #26375 +#26379 := [monotonicity #26376]: #26378 +#26384 := [trans #26379 #26382]: #26383 +#26387 := [monotonicity #26384]: #26386 +#26124 := [monotonicity #26387]: #26122 +#26127 := [trans #26124 #26126]: #26122 +#26117 := [quant-inst]: #26116 +#26128 := [mp #26117 #26127]: #26115 +#26282 := [unit-resolution #26128 #21660]: #26385 +#26130 := (or #26380 #26371) +#26131 := [def-axiom]: #26130 +#26283 := [unit-resolution #26131 #26282]: #26371 +#26285 := (= #3011 #26346) +#24115 := (= uf_7 #24114) +#1349 := (uf_124 #326 #161) +#1584 := (pattern #1349) +#1597 := (uf_144 #1349) +#8734 := (= #326 #1597) +#8738 := (forall (vars (?x388 T3) (?x389 int)) (:pat #1584) #8734) +#17847 := (~ #8738 #8738) +#17845 := (~ #8734 #8734) +#17846 := [refl]: #17845 +#17848 := [nnf-pos #17846]: #17847 +#1598 := (= #1597 #326) +#1599 := (forall (vars (?x388 T3) (?x389 int)) (:pat #1584) #1598) +#8739 := (iff #1599 #8738) +#8736 := (iff #1598 #8734) +#8737 := [rewrite]: #8736 +#8740 := [quant-intro #8737]: #8739 +#8733 := [asserted]: #1599 +#8743 := [mp #8733 #8740]: #8738 +#17849 := [mp~ #8743 #17848]: #8738 +#24118 := (not #8738) +#24119 := (or #24118 #24115) +#24120 := [quant-inst]: #24119 +#27681 := [unit-resolution #24120 #17849]: #24115 +#23226 := (= #2960 #23223) +#93 := (uf_29 #26) +#23593 := (pattern #93) +#94 := (uf_28 #93) +#3575 := (= #26 #94) +#23594 := (forall (vars (?x14 T5)) (:pat #23593) #3575) +#3578 := (forall (vars (?x14 T5)) #3575) +#23595 := (iff #3578 #23594) +#23597 := (iff #23594 #23594) +#23598 := [rewrite]: #23597 +#23596 := [rewrite]: #23595 +#23599 := [trans #23596 #23598]: #23595 +#16790 := (~ #3578 #3578) +#16780 := (~ #3575 #3575) +#16781 := [refl]: #16780 +#16851 := [nnf-pos #16781]: #16790 +#95 := (= #94 #26) +#96 := (forall (vars (?x14 T5)) #95) +#3579 := (iff #96 #3578) +#3576 := (iff #95 #3575) +#3577 := [rewrite]: #3576 +#3580 := [quant-intro #3577]: #3579 +#3574 := [asserted]: #96 +#3583 := [mp #3574 #3580]: #3578 +#16852 := [mp~ #3583 #16851]: #3578 +#23600 := [mp #16852 #23599]: #23594 +#23217 := (not #23594) +#23220 := (or #23217 #23226) +#23215 := [quant-inst]: #23220 +#26284 := [unit-resolution #23215 #23600]: #23226 +#26286 := [monotonicity #26284 #27681]: #26285 +#26303 := [trans #26286 #26283]: #26302 +#26305 := [monotonicity #26303]: #26304 +#26309 := [symm #26305]: #26308 +#26562 := (= #24114 #26561) +#26231 := (or #24181 #26562) +#26232 := [quant-inst]: #26231 +#26276 := [unit-resolution #26232 #23694]: #26562 +#26310 := [trans #26276 #26309]: #26288 +#26075 := (not #26288) +#26256 := [hypothesis]: #26290 +#26268 := (not #26264) +#26278 := (iff #18900 #26268) +#26267 := (iff #12361 #26264) +#26265 := (iff #26264 #12361) +#26258 := (= #26263 #3014) +#27682 := (= #24114 uf_7) +#27683 := [symm #27681]: #27682 +#26259 := [monotonicity #27683]: #26258 +#26266 := [monotonicity #26259]: #26265 +#26277 := [symm #26266]: #26267 +#26279 := [monotonicity #26277]: #26278 +#26257 := [hypothesis]: #18900 +#26280 := [mp #26257 #26279]: #26268 +#26106 := (or #26074 #26264 #26075) +#26108 := [def-axiom]: #26106 +#26281 := [unit-resolution #26108 #26280 #26256]: #26075 +#26311 := [unit-resolution #26281 #26310]: false +#26295 := [lemma #26311]: #26294 +#27925 := [unit-resolution #26295 #26214]: #12361 +#27926 := [hypothesis]: #23981 +#23241 := (or #23978 #23972) +#23222 := [def-axiom]: #23241 +#27936 := [unit-resolution #23222 #27926]: #23972 +decl uf_13 :: (-> T5 T6 T2) +decl uf_10 :: (-> T4 T5 T6) +#26039 := (uf_10 uf_273 #25404) +decl uf_143 :: (-> T3 int) +#24116 := (uf_143 #2952) +#26431 := (uf_124 #24114 #24116) +#26432 := (uf_43 #26431 #2961) +#26521 := (uf_13 #26432 #26039) +#26522 := (= uf_9 #26521) +#26040 := (uf_13 #25404 #26039) +#27955 := (= #26040 #26521) +#27949 := (= #26521 #26040) +#27947 := (= #26432 #25404) +#27934 := (= #26432 #2962) +#27932 := (= #26431 #2952) +#27923 := (= #24116 uf_272) +#24117 := (= uf_272 #24116) +#1594 := (uf_143 #1349) +#8727 := (= #161 #1594) +#8730 := (forall (vars (?x386 T3) (?x387 int)) (:pat #1584) #8727) +#17842 := (~ #8730 #8730) +#17840 := (~ #8727 #8727) +#17841 := [refl]: #17840 +#17843 := [nnf-pos #17841]: #17842 +#1595 := (= #1594 #161) +#1596 := (forall (vars (?x386 T3) (?x387 int)) (:pat #1584) #1595) +#8731 := (iff #1596 #8730) +#8728 := (iff #1595 #8727) +#8729 := [rewrite]: #8728 +#8732 := [quant-intro #8729]: #8731 +#8726 := [asserted]: #1596 +#8735 := [mp #8726 #8732]: #8730 +#17844 := [mp~ #8735 #17843]: #8730 +#24123 := (not #8730) +#24124 := (or #24123 #24117) +#24125 := [quant-inst]: #24124 +#27703 := [unit-resolution #24125 #17844]: #24117 +#27931 := [symm #27703]: #27923 +#27933 := [monotonicity #27683 #27931]: #27932 +#27935 := [monotonicity #27933]: #27934 +#27948 := [trans #27935 #27939]: #27947 +#27950 := [monotonicity #27948]: #27949 +#27953 := [symm #27950]: #27955 +#26041 := (= uf_9 #26040) +decl uf_53 :: (-> T4 T5 T6) +#26030 := (uf_53 uf_273 #25404) +#26031 := (uf_13 #26 #26030) +#26036 := (pattern #26031) +decl up_197 :: (-> T3 bool) +#26034 := (up_197 #25815) +#26032 := (= uf_9 #26031) +#26033 := (not #26032) +decl uf_147 :: (-> T5 T6 T2) +decl uf_192 :: (-> T7 T6) +decl uf_12 :: (-> T4 T5 T7) +#26026 := (uf_12 uf_273 #25404) +#26027 := (uf_192 #26026) +#26028 := (uf_147 #26 #26027) +#26029 := (= uf_9 #26028) +#26046 := (or #26029 #26033 #26034) +#26049 := (forall (vars (?x577 T5)) (:pat #26036) #26046) +#26052 := (not #26049) +#26042 := (not #26041) +#26055 := (or #25880 #26042 #26052) +#26058 := (not #26055) +#27945 := (= #3009 #25982) +#27946 := [symm #27943]: #27945 +#23240 := (or #23978 #12355) +#23229 := [def-axiom]: #23240 +#27938 := [unit-resolution #23229 #27926]: #12355 +#27924 := [trans #27938 #27946]: #25983 +#25988 := (or #26022 #25981 #25999) +#26021 := [def-axiom]: #25988 +#27927 := [unit-resolution #26021 #27924 #27937]: #25981 +#26061 := (or #26002 #26058) +#14 := (:var 2 T4) +#2162 := (uf_196 #14 #15 #26) +#2223 := (pattern #2162) +#2224 := (uf_53 #13 #24) +#2225 := (uf_13 #26 #2224) +#2226 := (pattern #2225) +#2154 := (uf_12 #13 #15) +#2231 := (uf_192 #2154) +#2232 := (uf_147 #26 #2231) +#10478 := (= uf_9 #2232) +#10467 := (= uf_9 #2225) +#22343 := (not #10467) +#1373 := (uf_15 #24) +#2228 := (up_197 #1373) +#22358 := (or #2228 #22343 #10478) +#22363 := (forall (vars (?x577 T5)) (:pat #2226) #22358) +#22369 := (not #22363) +#2140 := (uf_10 #14 #26) +#2141 := (uf_13 #15 #2140) +#10170 := (= uf_9 #2141) +#22177 := (not #10170) +#180 := (uf_24 #14 #15) +#3758 := (= uf_9 #180) +#10821 := (not #3758) +#22370 := (or #10821 #22177 #22369) +#22371 := (not #22370) +#10219 := (= uf_9 #2162) +#10502 := (not #10219) +#22376 := (or #10502 #22371) +#22379 := (forall (vars (?x574 T4) (?x575 T5) (?x576 T5)) (:pat #2223) #22376) +#2229 := (not #2228) +#10473 := (and #2229 #10467) +#10484 := (not #10473) +#10485 := (or #10484 #10478) +#10490 := (forall (vars (?x577 T5)) (:pat #2226) #10485) +#10511 := (and #3758 #10170 #10490) +#10514 := (or #10502 #10511) +#10517 := (forall (vars (?x574 T4) (?x575 T5) (?x576 T5)) (:pat #2223) #10514) +#22380 := (iff #10517 #22379) +#22377 := (iff #10514 #22376) +#22374 := (iff #10511 #22371) +#22366 := (and #3758 #10170 #22363) +#22372 := (iff #22366 #22371) +#22373 := [rewrite]: #22372 +#22367 := (iff #10511 #22366) +#22364 := (iff #10490 #22363) +#22361 := (iff #10485 #22358) +#22344 := (or #2228 #22343) +#22355 := (or #22344 #10478) +#22359 := (iff #22355 #22358) +#22360 := [rewrite]: #22359 +#22356 := (iff #10485 #22355) +#22353 := (iff #10484 #22344) +#22345 := (not #22344) +#22348 := (not #22345) +#22351 := (iff #22348 #22344) +#22352 := [rewrite]: #22351 +#22349 := (iff #10484 #22348) +#22346 := (iff #10473 #22345) +#22347 := [rewrite]: #22346 +#22350 := [monotonicity #22347]: #22349 +#22354 := [trans #22350 #22352]: #22353 +#22357 := [monotonicity #22354]: #22356 +#22362 := [trans #22357 #22360]: #22361 +#22365 := [quant-intro #22362]: #22364 +#22368 := [monotonicity #22365]: #22367 +#22375 := [trans #22368 #22373]: #22374 +#22378 := [monotonicity #22375]: #22377 +#22381 := [quant-intro #22378]: #22380 +#18361 := (~ #10517 #10517) +#18359 := (~ #10514 #10514) +#18357 := (~ #10511 #10511) +#18355 := (~ #10490 #10490) +#18353 := (~ #10485 #10485) +#18354 := [refl]: #18353 +#18356 := [nnf-pos #18354]: #18355 +#18351 := (~ #10170 #10170) +#18352 := [refl]: #18351 +#18349 := (~ #3758 #3758) +#18350 := [refl]: #18349 +#18358 := [monotonicity #18350 #18352 #18356]: #18357 +#18347 := (~ #10502 #10502) +#18348 := [refl]: #18347 +#18360 := [monotonicity #18348 #18358]: #18359 +#18362 := [nnf-pos #18360]: #18361 +#2145 := (= #2141 uf_9) +#184 := (= #180 uf_9) +#2236 := (and #184 #2145) +#2233 := (= #2232 uf_9) +#2227 := (= #2225 uf_9) +#2230 := (and #2227 #2229) +#2234 := (implies #2230 #2233) +#2235 := (forall (vars (?x577 T5)) (:pat #2226) #2234) +#2237 := (and #2235 #2236) +#2163 := (= #2162 uf_9) +#2238 := (implies #2163 #2237) +#2239 := (forall (vars (?x574 T4) (?x575 T5) (?x576 T5)) (:pat #2223) #2238) +#10520 := (iff #2239 #10517) +#10493 := (and #3758 #10170) +#10496 := (and #10490 #10493) +#10503 := (or #10502 #10496) +#10508 := (forall (vars (?x574 T4) (?x575 T5) (?x576 T5)) (:pat #2223) #10503) +#10518 := (iff #10508 #10517) +#10515 := (iff #10503 #10514) +#10512 := (iff #10496 #10511) +#10513 := [rewrite]: #10512 +#10516 := [monotonicity #10513]: #10515 +#10519 := [quant-intro #10516]: #10518 +#10509 := (iff #2239 #10508) +#10506 := (iff #2238 #10503) +#10499 := (implies #10219 #10496) +#10504 := (iff #10499 #10503) +#10505 := [rewrite]: #10504 +#10500 := (iff #2238 #10499) +#10497 := (iff #2237 #10496) +#10494 := (iff #2236 #10493) +#10171 := (iff #2145 #10170) +#10172 := [rewrite]: #10171 +#3759 := (iff #184 #3758) +#3760 := [rewrite]: #3759 +#10495 := [monotonicity #3760 #10172]: #10494 +#10491 := (iff #2235 #10490) +#10488 := (iff #2234 #10485) +#10481 := (implies #10473 #10478) +#10486 := (iff #10481 #10485) +#10487 := [rewrite]: #10486 +#10482 := (iff #2234 #10481) +#10479 := (iff #2233 #10478) +#10480 := [rewrite]: #10479 +#10476 := (iff #2230 #10473) +#10470 := (and #10467 #2229) +#10474 := (iff #10470 #10473) +#10475 := [rewrite]: #10474 +#10471 := (iff #2230 #10470) +#10468 := (iff #2227 #10467) +#10469 := [rewrite]: #10468 +#10472 := [monotonicity #10469]: #10471 +#10477 := [trans #10472 #10475]: #10476 +#10483 := [monotonicity #10477 #10480]: #10482 +#10489 := [trans #10483 #10487]: #10488 +#10492 := [quant-intro #10489]: #10491 +#10498 := [monotonicity #10492 #10495]: #10497 +#10220 := (iff #2163 #10219) +#10221 := [rewrite]: #10220 +#10501 := [monotonicity #10221 #10498]: #10500 +#10507 := [trans #10501 #10505]: #10506 +#10510 := [quant-intro #10507]: #10509 +#10521 := [trans #10510 #10519]: #10520 +#10466 := [asserted]: #2239 +#10522 := [mp #10466 #10521]: #10517 +#18363 := [mp~ #10522 #18362]: #10517 +#22382 := [mp #18363 #22381]: #22379 +#26123 := (not #22379) +#26129 := (or #26123 #26002 #26058) +#26035 := (or #26034 #26033 #26029) +#26037 := (forall (vars (?x577 T5)) (:pat #26036) #26035) +#26038 := (not #26037) +#26043 := (or #25880 #26042 #26038) +#26044 := (not #26043) +#26045 := (or #26002 #26044) +#26132 := (or #26123 #26045) +#26147 := (iff #26132 #26129) +#26144 := (or #26123 #26061) +#26145 := (iff #26144 #26129) +#26146 := [rewrite]: #26145 +#26142 := (iff #26132 #26144) +#26062 := (iff #26045 #26061) +#26059 := (iff #26044 #26058) +#26056 := (iff #26043 #26055) +#26053 := (iff #26038 #26052) +#26050 := (iff #26037 #26049) +#26047 := (iff #26035 #26046) +#26048 := [rewrite]: #26047 +#26051 := [quant-intro #26048]: #26050 +#26054 := [monotonicity #26051]: #26053 +#26057 := [monotonicity #26054]: #26056 +#26060 := [monotonicity #26057]: #26059 +#26063 := [monotonicity #26060]: #26062 +#26143 := [monotonicity #26063]: #26142 +#26148 := [trans #26143 #26146]: #26147 +#26133 := [quant-inst]: #26132 +#26149 := [mp #26133 #26148]: #26129 +#27928 := [unit-resolution #26149 #22382]: #26061 +#27929 := [unit-resolution #27928 #27927]: #26058 +#26216 := (or #26055 #26041) +#26217 := [def-axiom]: #26216 +#27930 := [unit-resolution #26217 #27929]: #26041 +#27956 := [trans #27930 #27953]: #26522 +#26523 := (not #26522) +#26711 := (or #12358 #26523) +#26511 := (uf_43 #24114 #2961) +#26512 := (uf_66 #26511 0::int #24114) +#26513 := (uf_27 uf_273 #26512) +#26514 := (= uf_9 #26513) +#26515 := (not #26514) +#26678 := (iff #18897 #26515) +#26674 := (iff #12358 #26514) +#26675 := (iff #26514 #12358) +#26687 := (= #26513 #3012) +#26683 := (= #26512 #3011) +#27689 := (= #26511 #2960) +#27687 := (= #2961 uf_274) +#24233 := (= uf_274 #2961) +#2693 := (uf_116 #2692) +#11669 := (= #161 #2693) +#23683 := (forall (vars (?x718 T3) (?x719 int)) (:pat #23682) #11669) +#11673 := (forall (vars (?x718 T3) (?x719 int)) #11669) +#23686 := (iff #11673 #23683) +#23684 := (iff #11669 #11669) +#23685 := [refl]: #23684 +#23687 := [quant-intro #23685]: #23686 +#18754 := (~ #11673 #11673) +#18752 := (~ #11669 #11669) +#18753 := [refl]: #18752 +#18755 := [nnf-pos #18753]: #18754 +#2694 := (= #2693 #161) +#2695 := (forall (vars (?x718 T3) (?x719 int)) #2694) +#11674 := (iff #2695 #11673) +#11671 := (iff #2694 #11669) +#11672 := [rewrite]: #11671 +#11675 := [quant-intro #11672]: #11674 +#11668 := [asserted]: #2695 +#11678 := [mp #11668 #11675]: #11673 +#18756 := [mp~ #11678 #18755]: #11673 +#23688 := [mp #18756 #23687]: #23683 +#24187 := (not #23683) +#24238 := (or #24187 #24233) +#24239 := [quant-inst]: #24238 +#27686 := [unit-resolution #24239 #23688]: #24233 +#27688 := [symm #27686]: #27687 +#27690 := [monotonicity #27683 #27688]: #27689 +#26686 := [monotonicity #27690 #27683]: #26683 +#26688 := [monotonicity #26686]: #26687 +#26676 := [monotonicity #26688]: #26675 +#26677 := [symm #26676]: #26674 +#26679 := [monotonicity #26677]: #26678 +#26638 := [hypothesis]: #18897 +#26680 := [mp #26638 #26679]: #26515 +#26516 := (uf_58 #3079 #26512) +#26517 := (uf_135 #26516) +#26518 := (= uf_9 #26517) +#26528 := (or #26515 #26518) +#26531 := (not #26528) +decl uf_23 :: (-> T3 T2) +#26524 := (uf_23 #24114) +#26525 := (= uf_9 #26524) +#2778 := (uf_23 uf_7) +#27721 := (= #2778 #26524) +#27718 := (= #26524 #2778) +#27719 := [monotonicity #27683]: #27718 +#27722 := [symm #27719]: #27721 +#11835 := (= uf_9 #2778) +#2779 := (= #2778 uf_9) +#11837 := (iff #2779 #11835) +#11838 := [rewrite]: #11837 +#11834 := [asserted]: #2779 +#11841 := [mp #11834 #11838]: #11835 +#27723 := [trans #11841 #27722]: #26525 +#26526 := (not #26525) +#26708 := (or #26526 #26531) +#27724 := [hypothesis]: #26522 +#26469 := (<= #24116 0::int) +#26682 := (not #26469) +#14790 := [not-or-elim #14776]: #13943 +#26452 := (* -1::int #24116) +#26584 := (+ uf_272 #26452) +#26585 := (<= #26584 0::int) +#27704 := (not #24117) +#27705 := (or #27704 #26585) +#27706 := [th-lemma]: #27705 +#27707 := [unit-resolution #27706 #27703]: #26585 +#27713 := (not #26585) +#26698 := (or #26682 #13942 #27713) +#26699 := [th-lemma]: #26698 +#26707 := [unit-resolution #26699 #27707 #14790]: #26682 +#237 := (uf_23 #233) +#758 := (:var 4 int) +#2062 := (uf_43 #233 #758) +#2063 := (uf_66 #2062 #247 #233) +#1364 := (:var 5 T4) +#2080 := (uf_25 #1364 #2063) +#1356 := (:var 3 T5) +#2060 := (uf_10 #1364 #1356) +#268 := (:var 2 int) +#2058 := (uf_124 #233 #268) +#2059 := (uf_43 #2058 #758) +#2061 := (uf_13 #2059 #2060) +#2081 := (pattern #2061 #2080 #237) +#1535 := (uf_59 #1364) +#2078 := (uf_58 #1535 #2063) +#2079 := (pattern #2061 #2078 #237) +#2085 := (uf_27 #1364 #2063) +#9989 := (= uf_9 #2085) +#22088 := (not #9989) +#2082 := (uf_135 #2078) +#9983 := (= uf_9 #2082) +#22089 := (or #9983 #22088) +#22090 := (not #22089) +#2067 := (uf_55 #1364) +#9932 := (= uf_9 #2067) +#22064 := (not #9932) +#9929 := (= uf_9 #2061) +#22063 := (not #9929) +#4079 := (* -1::int #268) +#6249 := (+ #247 #4079) +#6838 := (>= #6249 0::int) +#4346 := (>= #247 0::int) +#20033 := (not #4346) +#3963 := (= uf_9 #237) +#10698 := (not #3963) +#22096 := (or #10698 #20033 #6838 #22063 #22064 #22090) +#22101 := (forall (vars (?x526 T4) (?x527 int) (?x528 T5) (?x529 int) (?x530 int) (?x531 T3)) (:pat #2079 #2081) #22096) +#9986 := (not #9983) +#9992 := (and #9986 #9989) +#8189 := (not #6838) +#9965 := (and #3963 #4346 #8189 #9929 #9932) +#9970 := (not #9965) +#10006 := (or #9970 #9992) +#10009 := (forall (vars (?x526 T4) (?x527 int) (?x528 T5) (?x529 int) (?x530 int) (?x531 T3)) (:pat #2079 #2081) #10006) +#22102 := (iff #10009 #22101) +#22099 := (iff #10006 #22096) +#22065 := (or #10698 #20033 #6838 #22063 #22064) +#22093 := (or #22065 #22090) +#22097 := (iff #22093 #22096) +#22098 := [rewrite]: #22097 +#22094 := (iff #10006 #22093) +#22091 := (iff #9992 #22090) +#22092 := [rewrite]: #22091 +#22074 := (iff #9970 #22065) +#22066 := (not #22065) +#22069 := (not #22066) +#22072 := (iff #22069 #22065) +#22073 := [rewrite]: #22072 +#22070 := (iff #9970 #22069) +#22067 := (iff #9965 #22066) +#22068 := [rewrite]: #22067 +#22071 := [monotonicity #22068]: #22070 +#22075 := [trans #22071 #22073]: #22074 +#22095 := [monotonicity #22075 #22092]: #22094 +#22100 := [trans #22095 #22098]: #22099 +#22103 := [quant-intro #22100]: #22102 +#18227 := (~ #10009 #10009) +#18225 := (~ #10006 #10006) +#18226 := [refl]: #18225 +#18228 := [nnf-pos #18226]: #18227 +#2086 := (= #2085 uf_9) +#2083 := (= #2082 uf_9) +#2084 := (not #2083) +#2087 := (and #2084 #2086) +#2068 := (= #2067 uf_9) +#238 := (= #237 uf_9) +#2069 := (and #238 #2068) +#2066 := (= #2061 uf_9) +#2070 := (and #2066 #2069) +#400 := (<= 0::int #247) +#2071 := (and #400 #2070) +#1425 := (< #247 #268) +#2072 := (and #1425 #2071) +#2088 := (implies #2072 #2087) +#2089 := (forall (vars (?x526 T4) (?x527 int) (?x528 T5) (?x529 int) (?x530 int) (?x531 T3)) (:pat #2079 #2081) #2088) +#10012 := (iff #2089 #10009) +#9935 := (and #3963 #9932) +#9938 := (and #9929 #9935) +#9941 := (and #400 #9938) +#9944 := (and #1425 #9941) +#9950 := (not #9944) +#9998 := (or #9950 #9992) +#10003 := (forall (vars (?x526 T4) (?x527 int) (?x528 T5) (?x529 int) (?x530 int) (?x531 T3)) (:pat #2079 #2081) #9998) +#10010 := (iff #10003 #10009) +#10007 := (iff #9998 #10006) +#9971 := (iff #9950 #9970) +#9968 := (iff #9944 #9965) +#9959 := (and #4346 #9938) +#9962 := (and #8189 #9959) +#9966 := (iff #9962 #9965) +#9967 := [rewrite]: #9966 +#9963 := (iff #9944 #9962) +#9960 := (iff #9941 #9959) +#4345 := (iff #400 #4346) +#4347 := [rewrite]: #4345 +#9961 := [monotonicity #4347]: #9960 +#8190 := (iff #1425 #8189) +#8191 := [rewrite]: #8190 +#9964 := [monotonicity #8191 #9961]: #9963 +#9969 := [trans #9964 #9967]: #9968 +#9972 := [monotonicity #9969]: #9971 +#10008 := [monotonicity #9972]: #10007 +#10011 := [quant-intro #10008]: #10010 +#10004 := (iff #2089 #10003) +#10001 := (iff #2088 #9998) +#9995 := (implies #9944 #9992) +#9999 := (iff #9995 #9998) +#10000 := [rewrite]: #9999 +#9996 := (iff #2088 #9995) +#9993 := (iff #2087 #9992) +#9990 := (iff #2086 #9989) +#9991 := [rewrite]: #9990 +#9987 := (iff #2084 #9986) +#9984 := (iff #2083 #9983) +#9985 := [rewrite]: #9984 +#9988 := [monotonicity #9985]: #9987 +#9994 := [monotonicity #9988 #9991]: #9993 +#9945 := (iff #2072 #9944) +#9942 := (iff #2071 #9941) +#9939 := (iff #2070 #9938) +#9936 := (iff #2069 #9935) +#9933 := (iff #2068 #9932) +#9934 := [rewrite]: #9933 +#3964 := (iff #238 #3963) +#3965 := [rewrite]: #3964 +#9937 := [monotonicity #3965 #9934]: #9936 +#9930 := (iff #2066 #9929) +#9931 := [rewrite]: #9930 +#9940 := [monotonicity #9931 #9937]: #9939 +#9943 := [monotonicity #9940]: #9942 +#9946 := [monotonicity #9943]: #9945 +#9997 := [monotonicity #9946 #9994]: #9996 +#10002 := [trans #9997 #10000]: #10001 +#10005 := [quant-intro #10002]: #10004 +#10013 := [trans #10005 #10011]: #10012 +#9982 := [asserted]: #2089 +#10014 := [mp #9982 #10013]: #10009 +#18229 := [mp~ #10014 #18228]: #10009 +#22104 := [mp #18229 #22103]: #22101 +#26542 := (not #22101) +#26613 := (or #26542 #23209 #26469 #26523 #26526 #26531) +#26519 := (or #26518 #26515) +#26520 := (not #26519) +#26453 := (+ 0::int #26452) +#26454 := (>= #26453 0::int) +#26455 := (>= 0::int 0::int) +#26456 := (not #26455) +#26527 := (or #26526 #26456 #26454 #26523 #23209 #26520) +#26614 := (or #26542 #26527) +#26601 := (iff #26614 #26613) +#26537 := (or #23209 #26469 #26523 #26526 #26531) +#26616 := (or #26542 #26537) +#26619 := (iff #26616 #26613) +#26620 := [rewrite]: #26619 +#26617 := (iff #26614 #26616) +#26540 := (iff #26527 #26537) +#26534 := (or #26526 false #26469 #26523 #23209 #26531) +#26538 := (iff #26534 #26537) +#26539 := [rewrite]: #26538 +#26535 := (iff #26527 #26534) +#26532 := (iff #26520 #26531) +#26529 := (iff #26519 #26528) +#26530 := [rewrite]: #26529 +#26533 := [monotonicity #26530]: #26532 +#26472 := (iff #26454 #26469) +#26466 := (>= #26452 0::int) +#26470 := (iff #26466 #26469) +#26471 := [rewrite]: #26470 +#26467 := (iff #26454 #26466) +#26464 := (= #26453 #26452) +#26465 := [rewrite]: #26464 +#26468 := [monotonicity #26465]: #26467 +#26473 := [trans #26468 #26471]: #26472 +#26462 := (iff #26456 false) +#26460 := (iff #26456 #3294) +#26458 := (iff #26455 true) +#26459 := [rewrite]: #26458 +#26461 := [monotonicity #26459]: #26460 +#26463 := [trans #26461 #13445]: #26462 +#26536 := [monotonicity #26463 #26473 #26533]: #26535 +#26541 := [trans #26536 #26539]: #26540 +#26618 := [monotonicity #26541]: #26617 +#26602 := [trans #26618 #26620]: #26601 +#26615 := [quant-inst]: #26614 +#26603 := [mp #26615 #26602]: #26613 +#26706 := [unit-resolution #26603 #22104 #14784 #26707 #27724]: #26708 +#26709 := [unit-resolution #26706 #27723]: #26531 +#26604 := (or #26528 #26514) +#26605 := [def-axiom]: #26604 +#26710 := [unit-resolution #26605 #26709 #26680]: false +#26712 := [lemma #26710]: #26711 +#27952 := [unit-resolution #26712 #27956]: #12358 +#23238 := (or #23975 #18897 #18900 #23969) +#23239 := [def-axiom]: #23238 +#27957 := [unit-resolution #23239 #27952 #27925 #27936]: #23969 +#23252 := (or #23966 #23960) +#23263 := [def-axiom]: #23252 +#27958 := [unit-resolution #23263 #27957]: #23960 +#23245 := (or #23963 #18900 #18906 #23957) +#23258 := [def-axiom]: #23245 +#27959 := [unit-resolution #23258 #27958 #27925 #27858]: #18906 +#27024 := (or #27023 #12367 #26819) +#27025 := [def-axiom]: #27024 +#27961 := [unit-resolution #27025 #27959]: #27960 +#27962 := [unit-resolution #27961 #27857]: #26819 +#27902 := (or #26823 #26812) +#26997 := (or #26823 #18897 #26812) +#26998 := [def-axiom]: #26997 +#27904 := [unit-resolution #26998 #27952]: #27902 +#27905 := [unit-resolution #27904 #27962]: #26812 +#26956 := (or #26809 #26796) +#26991 := [def-axiom]: #26956 +#27903 := [unit-resolution #26991 #27905]: #26796 +#27585 := (not #26518) +#27980 := (iff #27585 #26552) +#27976 := (iff #26518 #26551) +#27987 := (= #26517 #26546) +#27910 := (= #26516 #26312) +#27911 := [monotonicity #26686]: #27910 +#27988 := [monotonicity #27911]: #27987 +#27979 := [monotonicity #27988]: #27976 +#27981 := [monotonicity #27979]: #27980 +#27907 := [unit-resolution #26603 #22104 #14784 #26707 #27956]: #26708 +#27908 := [unit-resolution #27907 #27723]: #26531 +#27597 := (or #26528 #27585) +#27598 := [def-axiom]: #27597 +#27909 := [unit-resolution #27598 #27908]: #27585 +#27982 := [mp #27909 #27981]: #26552 +#26910 := (or #26788 #26551) +#26911 := [def-axiom]: #26910 +#27983 := [unit-resolution #26911 #27982]: #26788 +#24653 := (uf_14 uf_7) +#27977 := (= #24653 #26634) +#27985 := (= #26634 #24653) +#27991 := (= #26287 uf_7) +#27989 := (= #26287 #24114) +#28002 := [mp #27925 #26277]: #26264 +#26014 := (or #26074 #26268 #26288) +#26016 := [def-axiom]: #26014 +#27986 := [unit-resolution #26016 #28002 #26214]: #26288 +#27990 := [symm #27986]: #27989 +#27992 := [trans #27990 #27683]: #27991 +#27993 := [monotonicity #27992]: #27985 +#27978 := [symm #27993]: #27977 +#24654 := (= uf_16 #24653) +#24661 := (iff #11835 #24654) +#2303 := (pattern #237) +#2831 := (uf_14 #233) +#12008 := (= uf_16 #2831) +#12012 := (iff #3963 #12008) +#12015 := (forall (vars (?x761 T3)) (:pat #2303) #12012) +#18854 := (~ #12015 #12015) +#18852 := (~ #12012 #12012) +#18853 := [refl]: #18852 +#18855 := [nnf-pos #18853]: #18854 +#2844 := (= #2831 uf_16) +#2845 := (iff #238 #2844) +#2846 := (forall (vars (?x761 T3)) (:pat #2303) #2845) +#12016 := (iff #2846 #12015) +#12013 := (iff #2845 #12012) +#12010 := (iff #2844 #12008) +#12011 := [rewrite]: #12010 +#12014 := [monotonicity #3965 #12011]: #12013 +#12017 := [quant-intro #12014]: #12016 +#12007 := [asserted]: #2846 +#12020 := [mp #12007 #12017]: #12015 +#18856 := [mp~ #12020 #18855]: #12015 +#24285 := (not #12015) +#24664 := (or #24285 #24661) +#24665 := [quant-inst]: #24664 +#27984 := [unit-resolution #24665 #18856]: #24661 +#24666 := (not #24661) +#28001 := (or #24666 #24654) +#24670 := (not #11835) +#24671 := (or #24666 #24670 #24654) +#24672 := [def-axiom]: #24671 +#28003 := [unit-resolution #24672 #11841]: #28001 +#28004 := [unit-resolution #28003 #27984]: #24654 +#28005 := [trans #28004 #27978]: #26726 +#26958 := (not #26629) +#28390 := (iff #12299 #26958) +#28388 := (iff #12296 #26629) +#28355 := (iff #26629 #12296) +#28362 := (= #26609 #2955) +#28360 := (= #26608 #2952) +#28357 := (= #26608 #24234) +#28329 := (= #26553 #2962) +#28302 := (= #26553 #26432) +#26435 := (uf_66 #26432 0::int #24114) +#26436 := (uf_58 #3079 #26435) +#26437 := (uf_136 #26436) +#28300 := (= #26437 #26432) +#26438 := (= #26432 #26437) +decl up_68 :: (-> T14 bool) +#26445 := (up_68 #26436) +#26446 := (not #26445) +#26442 := (uf_27 uf_273 #26435) +#26443 := (= uf_9 #26442) +#26444 := (not #26443) +#26440 := (uf_135 #26436) +#26441 := (= uf_9 #26440) +#26439 := (not #26438) +#26474 := (or #26439 #26441 #26444 #26446) +#26477 := (not #26474) +#26449 := (uf_27 uf_273 #26432) +#26450 := (= uf_9 #26449) +#28032 := (= #2963 #26449) +#28007 := (= #26449 #2963) +#28013 := [monotonicity #27935]: #28007 +#28033 := [symm #28013]: #28032 +#28031 := [trans #14797 #28033]: #26450 +#26451 := (not #26450) +#28034 := (or #26451 #26477) +#276 := (:var 3 int) +#310 := (:var 2 T3) +#1463 := (uf_124 #310 #247) +#1464 := (uf_43 #1463 #276) +#1460 := (uf_43 #310 #276) +#1461 := (uf_66 #1460 #161 #310) +#38 := (:var 4 T4) +#1466 := (uf_59 #38) +#1467 := (uf_58 #1466 #1461) +#1468 := (pattern #1467 #1464) +#1459 := (uf_41 #38) +#1462 := (uf_40 #1459 #1461) +#1465 := (pattern #1462 #1464) +#1471 := (uf_66 #1464 #161 #310) +#1474 := (uf_58 #1466 #1471) +#1479 := (uf_136 #1474) +#8354 := (= #1464 #1479) +#21428 := (not #8354) +#1476 := (uf_135 #1474) +#8348 := (= uf_9 #1476) +#1472 := (uf_27 #38 #1471) +#8345 := (= uf_9 #1472) +#21427 := (not #8345) +#1475 := (up_68 #1474) +#21426 := (not #1475) +#21429 := (or #21426 #21427 #8348 #21428) +#21430 := (not #21429) +#1469 := (uf_27 #38 #1464) +#8342 := (= uf_9 #1469) +#8377 := (not #8342) +#5373 := (* -1::int #247) +#6256 := (+ #161 #5373) +#6255 := (>= #6256 0::int) +#21436 := (or #5113 #6255 #8377 #21430) +#21441 := (forall (vars (?x346 T4) (?x347 int) (?x348 T3) (?x349 int) (?x350 int)) (:pat #1465 #1468) #21436) +#8351 := (not #8348) +#8386 := (and #1475 #8345 #8351 #8354) +#8026 := (not #6255) +#8029 := (and #4084 #8026) +#8032 := (not #8029) +#8395 := (or #8032 #8377 #8386) +#8400 := (forall (vars (?x346 T4) (?x347 int) (?x348 T3) (?x349 int) (?x350 int)) (:pat #1465 #1468) #8395) +#21442 := (iff #8400 #21441) +#21439 := (iff #8395 #21436) +#21311 := (or #5113 #6255) +#21433 := (or #21311 #8377 #21430) +#21437 := (iff #21433 #21436) +#21438 := [rewrite]: #21437 +#21434 := (iff #8395 #21433) +#21431 := (iff #8386 #21430) +#21432 := [rewrite]: #21431 +#21320 := (iff #8032 #21311) +#21312 := (not #21311) +#21315 := (not #21312) +#21318 := (iff #21315 #21311) +#21319 := [rewrite]: #21318 +#21316 := (iff #8032 #21315) +#21313 := (iff #8029 #21312) +#21314 := [rewrite]: #21313 +#21317 := [monotonicity #21314]: #21316 +#21321 := [trans #21317 #21319]: #21320 +#21435 := [monotonicity #21321 #21432]: #21434 +#21440 := [trans #21435 #21438]: #21439 +#21443 := [quant-intro #21440]: #21442 +#17588 := (~ #8400 #8400) +#17586 := (~ #8395 #8395) +#17587 := [refl]: #17586 +#17589 := [nnf-pos #17587]: #17588 +#1480 := (= #1479 #1464) +#1477 := (= #1476 uf_9) +#1478 := (not #1477) +#1481 := (and #1478 #1480) +#1482 := (and #1475 #1481) +#1473 := (= #1472 uf_9) +#1483 := (and #1473 #1482) +#1362 := (< #161 #247) +#1363 := (and #1362 #285) +#1484 := (implies #1363 #1483) +#1470 := (= #1469 uf_9) +#1485 := (implies #1470 #1484) +#1486 := (forall (vars (?x346 T4) (?x347 int) (?x348 T3) (?x349 int) (?x350 int)) (:pat #1465 #1468) #1485) +#8403 := (iff #1486 #8400) +#8357 := (and #8351 #8354) +#8360 := (and #1475 #8357) +#8363 := (and #8345 #8360) +#7987 := (and #285 #1362) +#7996 := (not #7987) +#8369 := (or #7996 #8363) +#8378 := (or #8377 #8369) +#8383 := (forall (vars (?x346 T4) (?x347 int) (?x348 T3) (?x349 int) (?x350 int)) (:pat #1465 #1468) #8378) +#8401 := (iff #8383 #8400) +#8398 := (iff #8378 #8395) +#8389 := (or #8032 #8386) +#8392 := (or #8377 #8389) +#8396 := (iff #8392 #8395) +#8397 := [rewrite]: #8396 +#8393 := (iff #8378 #8392) +#8390 := (iff #8369 #8389) +#8387 := (iff #8363 #8386) +#8388 := [rewrite]: #8387 +#8033 := (iff #7996 #8032) +#8030 := (iff #7987 #8029) +#8027 := (iff #1362 #8026) +#8028 := [rewrite]: #8027 +#8031 := [monotonicity #4085 #8028]: #8030 +#8034 := [monotonicity #8031]: #8033 +#8391 := [monotonicity #8034 #8388]: #8390 +#8394 := [monotonicity #8391]: #8393 +#8399 := [trans #8394 #8397]: #8398 +#8402 := [quant-intro #8399]: #8401 +#8384 := (iff #1486 #8383) +#8381 := (iff #1485 #8378) +#8374 := (implies #8342 #8369) +#8379 := (iff #8374 #8378) +#8380 := [rewrite]: #8379 +#8375 := (iff #1485 #8374) +#8372 := (iff #1484 #8369) +#8366 := (implies #7987 #8363) +#8370 := (iff #8366 #8369) +#8371 := [rewrite]: #8370 +#8367 := (iff #1484 #8366) +#8364 := (iff #1483 #8363) +#8361 := (iff #1482 #8360) +#8358 := (iff #1481 #8357) +#8355 := (iff #1480 #8354) +#8356 := [rewrite]: #8355 +#8352 := (iff #1478 #8351) +#8349 := (iff #1477 #8348) +#8350 := [rewrite]: #8349 +#8353 := [monotonicity #8350]: #8352 +#8359 := [monotonicity #8353 #8356]: #8358 +#8362 := [monotonicity #8359]: #8361 +#8346 := (iff #1473 #8345) +#8347 := [rewrite]: #8346 +#8365 := [monotonicity #8347 #8362]: #8364 +#7988 := (iff #1363 #7987) +#7989 := [rewrite]: #7988 +#8368 := [monotonicity #7989 #8365]: #8367 +#8373 := [trans #8368 #8371]: #8372 +#8343 := (iff #1470 #8342) +#8344 := [rewrite]: #8343 +#8376 := [monotonicity #8344 #8373]: #8375 +#8382 := [trans #8376 #8380]: #8381 +#8385 := [quant-intro #8382]: #8384 +#8404 := [trans #8385 #8402]: #8403 +#8341 := [asserted]: #1486 +#8405 := [mp #8341 #8404]: #8400 +#17590 := [mp~ #8405 #17589]: #8400 +#21444 := [mp #17590 #21443]: #21441 +#27098 := (not #21441) +#27099 := (or #27098 #26451 #26469 #26477) +#26447 := (or #26446 #26444 #26441 #26439) +#26448 := (not #26447) +#26457 := (or #26456 #26454 #26451 #26448) +#27111 := (or #27098 #26457) +#27522 := (iff #27111 #27099) +#26483 := (or #26451 #26469 #26477) +#27450 := (or #27098 #26483) +#27435 := (iff #27450 #27099) +#27448 := [rewrite]: #27435 +#27451 := (iff #27111 #27450) +#26486 := (iff #26457 #26483) +#26480 := (or false #26469 #26451 #26477) +#26484 := (iff #26480 #26483) +#26485 := [rewrite]: #26484 +#26481 := (iff #26457 #26480) +#26478 := (iff #26448 #26477) +#26475 := (iff #26447 #26474) +#26476 := [rewrite]: #26475 +#26479 := [monotonicity #26476]: #26478 +#26482 := [monotonicity #26463 #26473 #26479]: #26481 +#26487 := [trans #26482 #26485]: #26486 +#27428 := [monotonicity #26487]: #27451 +#27523 := [trans #27428 #27448]: #27522 +#27449 := [quant-inst]: #27111 +#27524 := [mp #27449 #27523]: #27099 +#28015 := [unit-resolution #27524 #21444 #26707]: #28034 +#28035 := [unit-resolution #28015 #28031]: #26477 +#27525 := (or #26474 #26438) +#27526 := [def-axiom]: #27525 +#28036 := [unit-resolution #27526 #28035]: #26438 +#28301 := [symm #28036]: #28300 +#28299 := (= #26553 #26437) +#28298 := (= #26312 #26436) +#28308 := (= #26436 #26312) +#28325 := (= #26435 #3011) +#26269 := (uf_116 #3011) +#26270 := (uf_43 #24114 #26269) +#28320 := (= #26270 #3011) +#26271 := (= #3011 #26270) +#26500 := (or #25416 #26268 #26271) +#26272 := (or #26271 #26268) +#26501 := (or #25416 #26272) +#26508 := (iff #26501 #26500) +#26273 := (or #26268 #26271) +#26503 := (or #25416 #26273) +#26506 := (iff #26503 #26500) +#26507 := [rewrite]: #26506 +#26504 := (iff #26501 #26503) +#26274 := (iff #26272 #26273) +#26275 := [rewrite]: #26274 +#26505 := [monotonicity #26275]: #26504 +#26509 := [trans #26505 #26507]: #26508 +#26502 := [quant-inst]: #26501 +#26396 := [mp #26502 #26509]: #26500 +#28037 := [unit-resolution #26396 #18736 #28002]: #26271 +#28321 := [symm #28037]: #28320 +#28324 := (= #26435 #26270) +#26643 := (uf_116 #25404) +#26651 := (+ #26356 #26643) +#26654 := (uf_43 #24114 #26651) +#28304 := (= #26654 #26270) +#28237 := (= #26651 #26269) +#26563 := (uf_116 #26368) +#28283 := (= #26563 #26269) +#28058 := (= #26368 #3011) +#28056 := (= #26346 #3011) +#28038 := (= #23223 #2960) +#28039 := [symm #26284]: #28038 +#28057 := [monotonicity #28039 #27683]: #28056 +#28040 := (= #26368 #26346) +#28050 := [symm #26283]: #28040 +#28059 := [trans #28050 #28057]: #28058 +#28284 := [monotonicity #28059]: #28283 +#28282 := (= #26651 #26563) +#28272 := (= #26563 #26651) +#27071 := (* -1::int #26357) +#27072 := (+ #24016 #27071) +#27073 := (<= #27072 0::int) +#27070 := (= #24016 #26357) +#28067 := (= #2961 #26357) +#28087 := (= #26357 #2961) +#28088 := [monotonicity #28039]: #28087 +#28068 := [symm #28088]: #28067 +#28085 := (= #24016 #2961) +#24240 := (= #2961 #24016) +#24245 := (or #24187 #24240) +#24246 := [quant-inst]: #24245 +#28060 := [unit-resolution #24246 #23688]: #24240 +#28086 := [symm #28060]: #28085 +#28069 := [trans #28086 #28068]: #27070 +#28070 := (not #27070) +#28049 := (or #28070 #27073) +#28066 := [th-lemma]: #28049 +#28051 := [unit-resolution #28066 #28069]: #27073 +#27068 := (>= #27072 0::int) +#28052 := (or #28070 #27068) +#28053 := [th-lemma]: #28052 +#28054 := [unit-resolution #28053 #28069]: #27068 +#26567 := (* -1::int #26563) +#26568 := (+ #26357 #26567) +#26569 := (+ #26356 #26568) +#27092 := (<= #26569 0::int) +#26570 := (= #26569 0::int) +#27074 := (or #24187 #26570) +#26564 := (= #26365 #26563) +#27075 := (or #24187 #26564) +#27077 := (iff #27075 #27074) +#27083 := (iff #27074 #27074) +#27084 := [rewrite]: #27083 +#26571 := (iff #26564 #26570) +#26572 := [rewrite]: #26571 +#27078 := [monotonicity #26572]: #27077 +#27093 := [trans #27078 #27084]: #27077 +#27076 := [quant-inst]: #27075 +#27094 := [mp #27076 #27093]: #27074 +#28055 := [unit-resolution #27094 #23688]: #26570 +#28076 := (not #26570) +#28079 := (or #28076 #27092) +#28080 := [th-lemma]: #28079 +#28081 := [unit-resolution #28080 #28055]: #27092 +#27095 := (>= #26569 0::int) +#28082 := (or #28076 #27095) +#28078 := [th-lemma]: #28082 +#28083 := [unit-resolution #28078 #28055]: #27095 +#27032 := (<= #26356 1::int) +#27031 := (= #26356 1::int) +#2927 := (uf_138 uf_7) +#2928 := (= #2927 1::int) +#12262 := [asserted]: #2928 +#28084 := (= #26356 #2927) +#28103 := [monotonicity #27683]: #28084 +#28105 := [trans #28103 #12262]: #27031 +#28106 := (not #27031) +#28261 := (or #28106 #27032) +#28262 := [th-lemma]: #28261 +#28263 := [unit-resolution #28262 #28105]: #27032 +#27069 := (>= #26356 1::int) +#28264 := (or #28106 #27069) +#28265 := [th-lemma]: #28264 +#28266 := [unit-resolution #28265 #28105]: #27069 +#27890 := (* -1::int #26643) +#27891 := (+ #24016 #27890) +#27892 := (<= #27891 0::int) +#27887 := (= #24016 #26643) +#28253 := (= #26643 #24016) +#28254 := [monotonicity #27941]: #28253 +#28252 := [symm #28254]: #27887 +#28255 := (not #27887) +#28256 := (or #28255 #27892) +#28257 := [th-lemma]: #28256 +#28258 := [unit-resolution #28257 #28252]: #27892 +#27893 := (>= #27891 0::int) +#28259 := (or #28255 #27893) +#28260 := [th-lemma]: #28259 +#28271 := [unit-resolution #28260 #28252]: #27893 +#28281 := [th-lemma #28266 #28263 #28271 #28258 #28266 #28263 #28083 #28081 #28054 #28051]: #28272 +#28280 := [symm #28281]: #28282 +#28239 := [trans #28280 #28284]: #28237 +#28305 := [monotonicity #28239]: #28304 +#28322 := (= #26435 #26654) +#26639 := (uf_66 #25404 0::int #24114) +#26657 := (= #26639 #26654) +#26660 := (not #26657) +#26640 := (uf_139 #26639 #25404) +#26641 := (= uf_9 #26640) +#26642 := (not #26641) +#26666 := (or #26642 #26660) +#26671 := (not #26666) +#27881 := (or #26114 #26671) +#26644 := (+ #26643 #26356) +#26645 := (+ 0::int #26644) +#26646 := (uf_43 #24114 #26645) +#26647 := (= #26639 #26646) +#26648 := (not #26647) +#26649 := (or #26648 #26642) +#26650 := (not #26649) +#27869 := (or #26114 #26650) +#27883 := (iff #27869 #27881) +#27885 := (iff #27881 #27881) +#27886 := [rewrite]: #27885 +#26672 := (iff #26650 #26671) +#26669 := (iff #26649 #26666) +#26663 := (or #26660 #26642) +#26667 := (iff #26663 #26666) +#26668 := [rewrite]: #26667 +#26664 := (iff #26649 #26663) +#26661 := (iff #26648 #26660) +#26658 := (iff #26647 #26657) +#26655 := (= #26646 #26654) +#26652 := (= #26645 #26651) +#26653 := [rewrite]: #26652 +#26656 := [monotonicity #26653]: #26655 +#26659 := [monotonicity #26656]: #26658 +#26662 := [monotonicity #26659]: #26661 +#26665 := [monotonicity #26662]: #26664 +#26670 := [trans #26665 #26668]: #26669 +#26673 := [monotonicity #26670]: #26672 +#27884 := [monotonicity #26673]: #27883 +#27896 := [trans #27884 #27886]: #27883 +#27882 := [quant-inst]: #27869 +#27897 := [mp #27882 #27896]: #27881 +#28240 := [unit-resolution #27897 #21660]: #26671 +#27900 := (or #26666 #26657) +#27901 := [def-axiom]: #27900 +#28238 := [unit-resolution #27901 #28240]: #26657 +#28310 := (= #26435 #26639) +#28311 := [monotonicity #27948]: #28310 +#28323 := [trans #28311 #28238]: #28322 +#28319 := [trans #28323 #28305]: #28324 +#28326 := [trans #28319 #28321]: #28325 +#28296 := [monotonicity #28326]: #28308 +#28309 := [symm #28296]: #28298 +#28297 := [monotonicity #28309]: #28299 +#28303 := [trans #28297 #28301]: #28302 +#28335 := [trans #28303 #27935]: #28329 +#28334 := [monotonicity #28335]: #28357 +#28361 := [trans #28334 #28359]: #28360 +#28363 := [monotonicity #28361]: #28362 +#28356 := [monotonicity #28363]: #28355 +#28389 := [symm #28356]: #28388 +#28391 := [monotonicity #28389]: #28390 +#28392 := [mp #14796 #28391]: #26958 +#28395 := (= #2967 #26630) +#28387 := (= #26630 #2967) +#28393 := [monotonicity #28335]: #28387 +#28328 := [symm #28393]: #28395 +#28349 := [trans #14799 #28328]: #26610 +#26877 := (not #26610) +#26878 := (or #26753 #26877) +#26905 := [def-axiom]: #26878 +#28327 := [unit-resolution #26905 #28349]: #26753 +#26952 := (or #26807 #26629 #26750 #26766 #26791) +#26953 := [def-axiom]: #26952 +#28350 := [unit-resolution #26953 #28327 #28392 #28005 #27983 #27903]: false +#28351 := [lemma #28350]: #28348 +#28242 := [unit-resolution #28351 #28241]: #23957 +#23303 := (or #23954 #3022) +#23302 := [def-axiom]: #23303 +#28243 := [unit-resolution #23302 #28242]: #3022 +#28633 := (+ #3021 #18936) +#26421 := (>= #28633 0::int) +#28632 := (= #3021 #18935) +#27102 := (= #18935 #3021) +#26769 := (= #18934 #3011) +#26767 := (= ?x773!13 0::int) +#23266 := (not #18939) +#26720 := [hypothesis]: #22757 +#23257 := (or #22752 #23266) +#23268 := [def-axiom]: #23257 +#26762 := [unit-resolution #23268 #26720]: #23266 +#23178 := (or #22752 #18931) +#23264 := [def-axiom]: #23178 +#26763 := [unit-resolution #23264 #26720]: #18931 +#26768 := [th-lemma #26763 #26762]: #26767 +#27101 := [monotonicity #26768]: #26769 +#27157 := [monotonicity #27101]: #27102 +#28041 := [symm #27157]: #28632 +#28023 := (not #28632) +#28024 := (or #28023 #26421) +#28022 := [th-lemma]: #28024 +#28025 := [unit-resolution #28022 #28041]: #26421 +#23179 := (not #18938) +#23265 := (or #22752 #23179) +#23180 := [def-axiom]: #23265 +#28026 := [unit-resolution #23180 #26720]: #23179 +#26970 := (* -1::int #3021) +#26971 := (+ uf_285 #26970) +#26972 := (>= #26971 0::int) +#28244 := (or #13672 #26972) +#28245 := [th-lemma]: #28244 +#28246 := [unit-resolution #28245 #28243]: #26972 +#28641 := [th-lemma #28246 #28026 #28025]: false +#28642 := [lemma #28641]: #22752 +#23280 := (or #23954 #23948) +#23281 := [def-axiom]: #23280 +#29203 := [unit-resolution #23281 #28242]: #23948 +#28560 := [hypothesis]: #13906 +#28561 := [th-lemma #14790 #28560]: false +#28562 := [lemma #28561]: #13903 +#23300 := (or #23951 #13906 #23945) +#23301 := [def-axiom]: #23300 +#29204 := [unit-resolution #23301 #28562 #29203]: #23945 +#23309 := (or #23942 #23936) +#23310 := [def-axiom]: #23309 +#29207 := [unit-resolution #23310 #29204]: #23936 +#23328 := (or #23939 #22757 #23933) +#23305 := [def-axiom]: #23328 +#29208 := [unit-resolution #23305 #29207 #28642]: #23933 +#23321 := (or #23930 #23924) +#23322 := [def-axiom]: #23321 +#29209 := [unit-resolution #23322 #29208]: #23924 +#29210 := (or #23927 #13672 #23921) +#23317 := (or #23927 #13672 #13942 #23921) +#23318 := [def-axiom]: #23317 +#29211 := [unit-resolution #23318 #14790]: #29210 +#29212 := [unit-resolution #29211 #29209 #28243]: #23921 +#23351 := (or #23918 #13947) +#23355 := [def-axiom]: #23351 +#29213 := [unit-resolution #23355 #29212]: #13947 +#27053 := (* -1::int #26964) +#27103 := (+ uf_293 #27053) +#27104 := (<= #27103 0::int) +#26965 := (= uf_293 #26964) +#1382 := (uf_66 #15 #161 #1381) +#1383 := (pattern #1382) +#1384 := (uf_125 #1382 #15) +#8071 := (= #161 #1384) +#8075 := (forall (vars (?x319 T5) (?x320 int)) (:pat #1383) #8071) +#17553 := (~ #8075 #8075) +#17551 := (~ #8071 #8071) +#17552 := [refl]: #17551 +#17554 := [nnf-pos #17552]: #17553 +#1385 := (= #1384 #161) +#1386 := (forall (vars (?x319 T5) (?x320 int)) (:pat #1383) #1385) +#8076 := (iff #1386 #8075) +#8073 := (iff #1385 #8071) +#8074 := [rewrite]: #8073 +#8077 := [quant-intro #8074]: #8076 +#8070 := [asserted]: #1386 +#8080 := [mp #8070 #8077]: #8075 +#17555 := [mp~ #8080 #17554]: #8075 +#26411 := (not #8075) +#26968 := (or #26411 #26965) +#26969 := [quant-inst]: #26968 +#27438 := [unit-resolution #26969 #17555]: #26965 +#27439 := (not #26965) +#29214 := (or #27439 #27104) +#29215 := [th-lemma]: #29214 +#29216 := [unit-resolution #29215 #27438]: #27104 +#29217 := (not #27104) +#29218 := (or #27037 #22873 #29217) +#29219 := [th-lemma]: #29218 +#29220 := [unit-resolution #29219 #29216 #29213]: #27037 +#23345 := (or #23918 #23754) +#23338 := [def-axiom]: #23345 +#29221 := [unit-resolution #23338 #29212]: #23754 +#23365 := (or #23918 #12426) +#23366 := [def-axiom]: #23365 +#29222 := [unit-resolution #23366 #29212]: #12426 +#27373 := (+ uf_272 #27053) +#27374 := (<= #27373 0::int) +#27445 := (not #27374) +#23356 := (or #23918 #14405) +#23359 := [def-axiom]: #23356 +#29223 := [unit-resolution #23359 #29212]: #14405 +#27446 := (or #27445 #14404) +#27437 := [hypothesis]: #14405 +#27105 := (>= #27103 0::int) +#27440 := (or #27439 #27105) +#27441 := [th-lemma]: #27440 +#27442 := [unit-resolution #27441 #27438]: #27105 +#27443 := [hypothesis]: #27374 +#27444 := [th-lemma #27443 #27442 #27437]: false +#27447 := [lemma #27444]: #27446 +#29224 := [unit-resolution #27447 #29223]: #27445 +#23346 := (or #23918 #23912) +#23339 := [def-axiom]: #23346 +#29225 := [unit-resolution #23339 #29212]: #23912 +#27311 := (<= #26964 131073::int) +#23336 := (or #23918 #16332) +#23337 := [def-axiom]: #23336 +#29226 := [unit-resolution #23337 #29212]: #16332 +#29227 := (not #27105) +#29228 := (or #27311 #23042 #29227) +#29229 := [th-lemma]: #29228 +#29230 := [unit-resolution #29229 #27442 #29226]: #27311 +#27312 := (not #27311) +#27038 := (not #27037) +#27757 := (or #14049 #27038 #27312 #27374 #23037 #23759 #23915) +#27327 := (uf_66 #2960 #26964 uf_7) +#27328 := (uf_110 uf_273 #27327) +#27331 := (= uf_299 #27328) +#27162 := (= #3068 #27328) +#27733 := (= #27328 #3068) +#27638 := (= #27327 #3067) +#27584 := (= #26964 uf_293) +#27589 := [symm #27438]: #27584 +#27639 := [monotonicity #27589]: #27638 +#27734 := [monotonicity #27639]: #27733 +#27665 := [symm #27734]: #27162 +#27735 := (= uf_299 #3068) +#27640 := [hypothesis]: #12426 +#27641 := [hypothesis]: #23912 +#27352 := [hypothesis]: #14046 +#23335 := (or #23872 #14049) +#23446 := [def-axiom]: #23335 +#27720 := [unit-resolution #23446 #27352]: #23872 +#23378 := (or #23915 #23875 #23909) +#23380 := [def-axiom]: #23378 +#27731 := [unit-resolution #23380 #27720 #27641]: #23909 +#23397 := (or #23906 #12576) +#23398 := [def-axiom]: #23397 +#27666 := [unit-resolution #23398 #27731]: #12576 +#27732 := [symm #27666]: #3139 +#27736 := [trans #27732 #27640]: #27735 +#27737 := [trans #27736 #27665]: #27331 +#27738 := [hypothesis]: #27445 +#27675 := [hypothesis]: #27311 +#27739 := [hypothesis]: #27037 +#23405 := (or #23906 #23900) +#23406 := [def-axiom]: #23405 +#27740 := [unit-resolution #23406 #27731]: #23900 +#27363 := [hypothesis]: #23754 +#27108 := (+ uf_292 #13970) +#27109 := (<= #27108 0::int) +#27741 := (or #12644 #27109) +#27742 := [th-lemma]: #27741 +#27743 := [unit-resolution #27742 #27666]: #27109 +#27349 := (not #27109) +#27367 := (or #23008 #23759 #27349 #14049) +#27179 := (+ uf_294 #19528) +#27180 := (<= #27179 0::int) +#27355 := (not #27180) +#23419 := (not #19530) +#27353 := [hypothesis]: #23013 +#23443 := (or #23008 #23419) +#23444 := [def-axiom]: #23443 +#27354 := [unit-resolution #23444 #27353]: #23419 +#27356 := (or #27355 #14049 #19530) +#27357 := [th-lemma]: #27356 +#27358 := [unit-resolution #27357 #27354 #27352]: #27355 +#27191 := (+ uf_292 #19541) +#27192 := (>= #27191 0::int) +#27348 := (not #27192) +#27342 := [hypothesis]: #27109 +#23439 := (not #19543) +#23445 := (or #23008 #23439) +#23413 := [def-axiom]: #23445 +#27359 := [unit-resolution #23413 #27353]: #23439 +#27350 := (or #27348 #19543 #27349) +#27343 := [hypothesis]: #23439 +#27346 := [hypothesis]: #27192 +#27347 := [th-lemma #27346 #27343 #27342]: false +#27351 := [lemma #27347]: #27350 +#27360 := [unit-resolution #27351 #27359 #27342]: #27348 +#27364 := (or #27180 #27192) +#23383 := (or #23008 #19192) +#23438 := [def-axiom]: #23383 +#27361 := [unit-resolution #23438 #27353]: #19192 +#23457 := (or #23008 #19191) +#23437 := [def-axiom]: #23457 +#27362 := [unit-resolution #23437 #27353]: #19191 +#27205 := (or #23759 #22992 #22993 #27180 #27192) +#27168 := (+ #19196 #14431) +#27169 := (<= #27168 0::int) +#27170 := (+ ?x781!15 #14044) +#27171 := (>= #27170 0::int) +#27172 := (or #22993 #27171 #27169 #22992) +#27206 := (or #23759 #27172) +#27213 := (iff #27206 #27205) +#27200 := (or #22992 #22993 #27180 #27192) +#27208 := (or #23759 #27200) +#27211 := (iff #27208 #27205) +#27212 := [rewrite]: #27211 +#27209 := (iff #27206 #27208) +#27203 := (iff #27172 #27200) +#27197 := (or #22993 #27180 #27192 #22992) +#27201 := (iff #27197 #27200) +#27202 := [rewrite]: #27201 +#27198 := (iff #27172 #27197) +#27195 := (iff #27169 #27192) +#27185 := (+ #14431 #19196) +#27188 := (<= #27185 0::int) +#27193 := (iff #27188 #27192) +#27194 := [rewrite]: #27193 +#27189 := (iff #27169 #27188) +#27186 := (= #27168 #27185) +#27187 := [rewrite]: #27186 +#27190 := [monotonicity #27187]: #27189 +#27196 := [trans #27190 #27194]: #27195 +#27183 := (iff #27171 #27180) +#27173 := (+ #14044 ?x781!15) +#27176 := (>= #27173 0::int) +#27181 := (iff #27176 #27180) +#27182 := [rewrite]: #27181 +#27177 := (iff #27171 #27176) +#27174 := (= #27170 #27173) +#27175 := [rewrite]: #27174 +#27178 := [monotonicity #27175]: #27177 +#27184 := [trans #27178 #27182]: #27183 +#27199 := [monotonicity #27184 #27196]: #27198 +#27204 := [trans #27199 #27202]: #27203 +#27210 := [monotonicity #27204]: #27209 +#27214 := [trans #27210 #27212]: #27213 +#27207 := [quant-inst]: #27206 +#27215 := [mp #27207 #27214]: #27205 +#27365 := [unit-resolution #27215 #27363 #27362 #27361]: #27364 +#27366 := [unit-resolution #27365 #27360 #27358]: false +#27368 := [lemma #27366]: #27367 +#27753 := [unit-resolution #27368 #27743 #27352 #27363]: #23008 +#23423 := (or #23903 #23897 #23013) +#23424 := [def-axiom]: #23423 +#27754 := [unit-resolution #23424 #27753 #27740]: #23897 +#23454 := (or #23894 #23886) +#23455 := [def-axiom]: #23454 +#27755 := [unit-resolution #23455 #27754]: #23886 +#27334 := (not #27331) +#27520 := (or #23891 #27038 #27312 #27334 #27374) +#27317 := (+ #26964 #13873) +#27318 := (>= #27317 0::int) +#27326 := (= #27328 uf_299) +#27329 := (not #27326) +#27330 := (or #27329 #27038 #27318 #27312) +#27518 := (or #23891 #27330) +#27642 := (iff #27518 #27520) +#27382 := (or #27038 #27312 #27334 #27374) +#27590 := (or #23891 #27382) +#27593 := (iff #27590 #27520) +#27594 := [rewrite]: #27593 +#27591 := (iff #27518 #27590) +#27385 := (iff #27330 #27382) +#27379 := (or #27334 #27038 #27374 #27312) +#27383 := (iff #27379 #27382) +#27384 := [rewrite]: #27383 +#27380 := (iff #27330 #27379) +#27377 := (iff #27318 #27374) +#27335 := (+ #13873 #26964) +#27370 := (>= #27335 0::int) +#27375 := (iff #27370 #27374) +#27376 := [rewrite]: #27375 +#27371 := (iff #27318 #27370) +#27336 := (= #27317 #27335) +#27369 := [rewrite]: #27336 +#27372 := [monotonicity #27369]: #27371 +#27378 := [trans #27372 #27376]: #27377 +#27344 := (iff #27329 #27334) +#27332 := (iff #27326 #27331) +#27333 := [rewrite]: #27332 +#27345 := [monotonicity #27333]: #27344 +#27381 := [monotonicity #27345 #27378]: #27380 +#27386 := [trans #27381 #27384]: #27385 +#27592 := [monotonicity #27386]: #27591 +#27647 := [trans #27592 #27594]: #27642 +#27521 := [quant-inst]: #27518 +#27648 := [mp #27521 #27647]: #27520 +#27756 := [unit-resolution #27648 #27755 #27739 #27675 #27738 #27737]: false +#27758 := [lemma #27756]: #27757 +#29231 := [unit-resolution #27758 #29230 #29225 #29224 #29222 #29221 #29220]: #14049 +#23541 := (+ uf_294 #14142) +#23536 := (>= #23541 0::int) +#27163 := (uf_58 #3079 #3175) +#27762 := (uf_136 #27163) +#27763 := (uf_24 uf_273 #27762) +#27764 := (= uf_9 #27763) +#27765 := (not #27764) +#27759 := (uf_135 #27163) +#27760 := (= uf_9 #27759) +#27761 := (not #27760) +#27819 := (or #27761 #27765) +#27822 := (not #27819) +#27773 := (uf_210 uf_273 #27762) +#27774 := (= uf_9 #27773) +#27771 := (uf_25 uf_273 #27762) +#27772 := (= uf_26 #27771) +#27813 := (or #27772 #27774) +#27816 := (not #27813) +#27527 := (uf_15 #3175) +#27777 := (uf_14 #27527) +#27795 := (= uf_16 #27777) +#27810 := (not #27795) +#27768 := (uf_15 #27762) +#27769 := (uf_14 #27768) +#27770 := (= uf_16 #27769) +#27828 := (or #27770 #27810 #27816 #27822) +#27833 := (not #27828) +#27784 := (uf_25 uf_273 #3175) +#27785 := (= uf_26 #27784) +#27782 := (uf_210 uf_273 #3175) +#27783 := (= uf_9 #27782) +#27798 := (or #27783 #27785) +#27801 := (not #27798) +#27804 := (or #27795 #27801) +#27807 := (not #27804) +#27836 := (or #27807 #27833) +#27839 := (not #27836) +#27842 := (or #19008 #27839) +#27845 := (not #27842) +#27848 := (iff #12812 #27845) +#29445 := (or #26854 #27848) +#27766 := (or #27765 #27761) +#27767 := (not #27766) +#27775 := (or #27774 #27772) +#27776 := (not #27775) +#27778 := (= #27777 uf_16) +#27779 := (not #27778) +#27780 := (or #27779 #27776 #27770 #27767) +#27781 := (not #27780) +#27786 := (or #27785 #27783) +#27787 := (not #27786) +#27788 := (or #27778 #27787) +#27789 := (not #27788) +#27790 := (or #27789 #27781) +#27791 := (not #27790) +#27792 := (or #19008 #27791) +#27793 := (not #27792) +#27794 := (iff #12812 #27793) +#29439 := (or #26854 #27794) +#29438 := (iff #29439 #29445) +#29456 := (iff #29445 #29445) +#29454 := [rewrite]: #29456 +#27849 := (iff #27794 #27848) +#27846 := (iff #27793 #27845) +#27843 := (iff #27792 #27842) +#27840 := (iff #27791 #27839) +#27837 := (iff #27790 #27836) +#27834 := (iff #27781 #27833) +#27831 := (iff #27780 #27828) +#27825 := (or #27810 #27816 #27770 #27822) +#27829 := (iff #27825 #27828) +#27830 := [rewrite]: #27829 +#27826 := (iff #27780 #27825) +#27823 := (iff #27767 #27822) +#27820 := (iff #27766 #27819) +#27821 := [rewrite]: #27820 +#27824 := [monotonicity #27821]: #27823 +#27817 := (iff #27776 #27816) +#27814 := (iff #27775 #27813) +#27815 := [rewrite]: #27814 +#27818 := [monotonicity #27815]: #27817 +#27811 := (iff #27779 #27810) +#27796 := (iff #27778 #27795) +#27797 := [rewrite]: #27796 +#27812 := [monotonicity #27797]: #27811 +#27827 := [monotonicity #27812 #27818 #27824]: #27826 +#27832 := [trans #27827 #27830]: #27831 +#27835 := [monotonicity #27832]: #27834 +#27808 := (iff #27789 #27807) +#27805 := (iff #27788 #27804) +#27802 := (iff #27787 #27801) +#27799 := (iff #27786 #27798) +#27800 := [rewrite]: #27799 +#27803 := [monotonicity #27800]: #27802 +#27806 := [monotonicity #27797 #27803]: #27805 +#27809 := [monotonicity #27806]: #27808 +#27838 := [monotonicity #27809 #27835]: #27837 +#27841 := [monotonicity #27838]: #27840 +#27844 := [monotonicity #27841]: #27843 +#27847 := [monotonicity #27844]: #27846 +#27850 := [monotonicity #27847]: #27849 +#29455 := [monotonicity #27850]: #29438 +#29457 := [trans #29455 #29454]: #29438 +#29446 := [quant-inst]: #29439 +#29459 := [mp #29446 #29457]: #29445 +#29640 := [unit-resolution #29459 #22514]: #27848 +#29381 := (not #27848) +#29642 := (or #29381 #27842) +#29641 := [hypothesis]: #19017 +#29377 := (or #29381 #12812 #27842) +#29382 := [def-axiom]: #29377 +#28866 := [unit-resolution #29382 #29641]: #29642 +#29632 := [unit-resolution #28866 #29640]: #27842 +#29635 := (or #27845 #27839) +#23357 := (or #23918 #13950) +#23358 := [def-axiom]: #23357 +#28992 := [unit-resolution #23358 #29212]: #13950 +#28994 := [trans #26494 #27946]: #25983 +#28995 := [unit-resolution #26021 #28994 #27937]: #25981 +#28996 := [unit-resolution #27928 #28995]: #26058 +#28997 := [unit-resolution #26217 #28996]: #26041 +#29000 := [trans #28997 #27953]: #26522 +#27729 := (or #12803 #14243 #26523 #14046) +#27672 := [hypothesis]: #13950 +#27528 := (uf_66 #23223 uf_294 #26404) +#27529 := (uf_125 #27528 #23223) +#27558 := (* -1::int #27529) +#27667 := (+ uf_294 #27558) +#27668 := (<= #27667 0::int) +#27530 := (= uf_294 #27529) +#27533 := (or #26411 #27530) +#27534 := [quant-inst]: #27533 +#27673 := [unit-resolution #27534 #17555]: #27530 +#27676 := (not #27530) +#27677 := (or #27676 #27668) +#27678 := [th-lemma]: #27677 +#27679 := [unit-resolution #27678 #27673]: #27668 +#27549 := (>= #27529 0::int) +#27550 := (not #27549) +#27601 := (uf_66 #26511 #27529 #24114) +#27605 := (uf_58 #3079 #27601) +#27606 := (uf_135 #27605) +#27607 := (= uf_9 #27606) +#27602 := (uf_27 uf_273 #27601) +#27603 := (= uf_9 #27602) +#27604 := (not #27603) +#27611 := (or #27604 #27607) +#27699 := (iff #19008 #27604) +#27697 := (iff #12803 #27603) +#27695 := (iff #27603 #12803) +#27693 := (= #27602 #3176) +#27691 := (= #27601 #3175) +#27684 := (= #27529 uf_294) +#27685 := [symm #27673]: #27684 +#27692 := [monotonicity #27690 #27685 #27683]: #27691 +#27694 := [monotonicity #27692]: #27693 +#27696 := [monotonicity #27694]: #27695 +#27698 := [symm #27696]: #27697 +#27700 := [monotonicity #27698]: #27699 +#27680 := [hypothesis]: #19008 +#27701 := [mp #27680 #27700]: #27604 +#27636 := (or #27611 #27603) +#27637 := [def-axiom]: #27636 +#27702 := [unit-resolution #27637 #27701]: #27611 +#27559 := (+ #24116 #27558) +#27560 := (<= #27559 0::int) +#27712 := (not #27560) +#27708 := [hypothesis]: #14049 +#27669 := (>= #27667 0::int) +#27709 := (or #27676 #27669) +#27710 := [th-lemma]: #27709 +#27711 := [unit-resolution #27710 #27673]: #27669 +#27714 := (not #27669) +#27715 := (or #27712 #27713 #27714 #14046) +#27716 := [th-lemma]: #27715 +#27717 := [unit-resolution #27716 #27711 #27708 #27707]: #27712 +#27614 := (not #27611) +#27725 := (or #27550 #27560 #27614) +#27625 := (or #26542 #23209 #26523 #26526 #27550 #27560 #27614) +#27608 := (or #27607 #27604) +#27609 := (not #27608) +#27547 := (+ #27529 #26452) +#27548 := (>= #27547 0::int) +#27610 := (or #26526 #27550 #27548 #26523 #23209 #27609) +#27626 := (or #26542 #27610) +#27633 := (iff #27626 #27625) +#27620 := (or #23209 #26523 #26526 #27550 #27560 #27614) +#27628 := (or #26542 #27620) +#27631 := (iff #27628 #27625) +#27632 := [rewrite]: #27631 +#27629 := (iff #27626 #27628) +#27623 := (iff #27610 #27620) +#27617 := (or #26526 #27550 #27560 #26523 #23209 #27614) +#27621 := (iff #27617 #27620) +#27622 := [rewrite]: #27621 +#27618 := (iff #27610 #27617) +#27615 := (iff #27609 #27614) +#27612 := (iff #27608 #27611) +#27613 := [rewrite]: #27612 +#27616 := [monotonicity #27613]: #27615 +#27563 := (iff #27548 #27560) +#27552 := (+ #26452 #27529) +#27555 := (>= #27552 0::int) +#27561 := (iff #27555 #27560) +#27562 := [rewrite]: #27561 +#27556 := (iff #27548 #27555) +#27553 := (= #27547 #27552) +#27554 := [rewrite]: #27553 +#27557 := [monotonicity #27554]: #27556 +#27564 := [trans #27557 #27562]: #27563 +#27619 := [monotonicity #27564 #27616]: #27618 +#27624 := [trans #27619 #27622]: #27623 +#27630 := [monotonicity #27624]: #27629 +#27634 := [trans #27630 #27632]: #27633 +#27627 := [quant-inst]: #27626 +#27635 := [mp #27627 #27634]: #27625 +#27726 := [unit-resolution #27635 #22104 #14784 #27724 #27723]: #27725 +#27727 := [unit-resolution #27726 #27717 #27702]: #27550 +#27728 := [th-lemma #27727 #27679 #27672]: false +#27730 := [lemma #27728]: #27729 +#29001 := [unit-resolution #27730 #29231 #29000 #28992]: #12803 +#29437 := (or #27845 #19008 #27839) +#29380 := [def-axiom]: #29437 +#29636 := [unit-resolution #29380 #29001]: #29635 +#29634 := [unit-resolution #29636 #29632]: #27839 +#29590 := (or #27836 #27828) +#29500 := [def-axiom]: #29590 +#29637 := [unit-resolution #29500 #29634]: #27828 +#29123 := (= #24653 #27777) +#29378 := (= #27777 #24653) +#29114 := (= #27527 uf_7) +#28862 := (= #27527 #24114) +#27514 := (= #24114 #27527) +#27302 := (uf_48 #3175 #24114) +#27308 := (= uf_9 #27302) +#27513 := (iff #27308 #27514) +#28999 := (or #25432 #27513) +#27515 := (iff #27514 #27308) +#28993 := (or #25432 #27515) +#29008 := (iff #28993 #28999) +#28998 := (iff #28999 #28999) +#29010 := [rewrite]: #28998 +#27516 := (iff #27515 #27513) +#27517 := [rewrite]: #27516 +#29009 := [monotonicity #27517]: #29008 +#29011 := [trans #29009 #29010]: #29008 +#29007 := [quant-inst]: #28993 +#29012 := [mp #29007 #29011]: #28999 +#29076 := [unit-resolution #29012 #23681]: #27513 +#29751 := (= #3178 #27302) +#29068 := (= #27302 #3178) +#29078 := [monotonicity #27683]: #29068 +#29752 := [symm #29078]: #29751 +#27490 := (+ uf_294 #26365) +#27493 := (uf_43 #24114 #27490) +#27643 := (uf_15 #27493) +#29135 := (= #27643 #27527) +#29116 := (= #27527 #27643) +#29048 := (= #3175 #27493) +#27480 := (uf_66 #23223 uf_294 #24114) +#27496 := (= #27480 #27493) +#27499 := (not #27496) +#27481 := (uf_139 #27480 #23223) +#27482 := (= uf_9 #27481) +#27483 := (not #27482) +#27505 := (or #27483 #27499) +#27510 := (not #27505) +#29033 := (or #26114 #27510) +#27484 := (+ uf_294 #26358) +#27485 := (uf_43 #24114 #27484) +#27486 := (= #27480 #27485) +#27487 := (not #27486) +#27488 := (or #27487 #27483) +#27489 := (not #27488) +#29034 := (or #26114 #27489) +#29030 := (iff #29034 #29033) +#29036 := (iff #29033 #29033) +#29037 := [rewrite]: #29036 +#27511 := (iff #27489 #27510) +#27508 := (iff #27488 #27505) +#27502 := (or #27499 #27483) +#27506 := (iff #27502 #27505) +#27507 := [rewrite]: #27506 +#27503 := (iff #27488 #27502) +#27500 := (iff #27487 #27499) +#27497 := (iff #27486 #27496) +#27494 := (= #27485 #27493) +#27491 := (= #27484 #27490) +#27492 := [rewrite]: #27491 +#27495 := [monotonicity #27492]: #27494 +#27498 := [monotonicity #27495]: #27497 +#27501 := [monotonicity #27498]: #27500 +#27504 := [monotonicity #27501]: #27503 +#27509 := [trans #27504 #27507]: #27508 +#27512 := [monotonicity #27509]: #27511 +#29035 := [monotonicity #27512]: #29030 +#29038 := [trans #29035 #29037]: #29030 +#29029 := [quant-inst]: #29034 +#29039 := [mp #29029 #29038]: #29033 +#29088 := [unit-resolution #29039 #21660]: #27510 +#28968 := (or #27505 #27496) +#29050 := [def-axiom]: #28968 +#29049 := [unit-resolution #29050 #29088]: #27496 +#29056 := (= #3175 #27480) +#29054 := (= #27480 #3175) +#29055 := [monotonicity #28039 #27683]: #29054 +#29057 := [symm #29055]: #29056 +#29082 := [trans #29057 #29049]: #29048 +#29117 := [monotonicity #29082]: #29116 +#29115 := [symm #29117]: #29135 +#27644 := (= #24114 #27643) +#29031 := (or #24181 #27644) +#29032 := [quant-inst]: #29031 +#29087 := [unit-resolution #29032 #23694]: #27644 +#29136 := [trans #29087 #29115]: #27514 +#28947 := (not #27514) +#27309 := (not #27308) +#29080 := (iff #19011 #27309) +#29071 := (iff #12806 #27308) +#29081 := (iff #27308 #12806) +#29066 := [monotonicity #29078]: #29081 +#29072 := [symm #29066]: #29071 +#29083 := [monotonicity #29072]: #29080 +#29077 := [hypothesis]: #19011 +#29079 := [mp #29077 #29083]: #27309 +#28946 := (not #27513) +#29042 := (or #28946 #27308 #28947) +#29043 := [def-axiom]: #29042 +#29084 := [unit-resolution #29043 #29079 #29076]: #28947 +#29137 := [unit-resolution #29084 #29136]: false +#29138 := [lemma #29137]: #12806 +#29753 := [trans #29138 #29752]: #27308 +#28964 := (or #28946 #27309 #27514) +#28951 := [def-axiom]: #28964 +#28867 := [unit-resolution #28951 #29753 #29076]: #27514 +#29113 := [symm #28867]: #28862 +#28864 := [trans #29113 #27683]: #29114 +#29141 := [monotonicity #28864]: #29378 +#29124 := [symm #29141]: #29123 +#29188 := [trans #28004 #29124]: #27795 +#29563 := (not #27770) +#29681 := (iff #12299 #29563) +#29679 := (iff #12296 #27770) +#29461 := (iff #27770 #12296) +#29267 := (= #27769 #2955) +#29265 := (= #27768 #2952) +#29264 := (= #27768 #24234) +#29783 := (= #27762 #2962) +#29781 := (= #27762 #26432) +#27531 := (uf_66 #26432 #27529 #24114) +#27532 := (uf_58 #3079 #27531) +#27535 := (uf_136 #27532) +#29779 := (= #27535 #26432) +#27536 := (= #26432 #27535) +#27543 := (up_68 #27532) +#27544 := (not #27543) +#27540 := (uf_27 uf_273 #27531) +#27541 := (= uf_9 #27540) +#27542 := (not #27541) +#27538 := (uf_135 #27532) +#27539 := (= uf_9 #27538) +#27537 := (not #27536) +#27565 := (or #27537 #27539 #27542 #27544) +#27568 := (not #27565) +#28420 := (or #27549 #14243) +#28416 := [hypothesis]: #27550 +#28417 := [th-lemma #28416 #27679 #27672]: false +#28421 := [lemma #28417]: #28420 +#29742 := [unit-resolution #28421 #28992]: #27549 +#29745 := (or #27712 #27714) +#29743 := (or #27712 #27714 #14046) +#29744 := [unit-resolution #27716 #27707]: #29743 +#29746 := [unit-resolution #29744 #29231]: #29745 +#29747 := [unit-resolution #29746 #27711]: #27712 +#29144 := (or #27098 #26451 #27550 #27560 #27568) +#27545 := (or #27544 #27542 #27539 #27537) +#27546 := (not #27545) +#27551 := (or #27550 #27548 #26451 #27546) +#29145 := (or #27098 #27551) +#29157 := (iff #29145 #29144) +#27574 := (or #26451 #27550 #27560 #27568) +#29168 := (or #27098 #27574) +#29156 := (iff #29168 #29144) +#29154 := [rewrite]: #29156 +#29155 := (iff #29145 #29168) +#27577 := (iff #27551 #27574) +#27571 := (or #27550 #27560 #26451 #27568) +#27575 := (iff #27571 #27574) +#27576 := [rewrite]: #27575 +#27572 := (iff #27551 #27571) +#27569 := (iff #27546 #27568) +#27566 := (iff #27545 #27565) +#27567 := [rewrite]: #27566 +#27570 := [monotonicity #27567]: #27569 +#27573 := [monotonicity #27564 #27570]: #27572 +#27578 := [trans #27573 #27576]: #27577 +#29164 := [monotonicity #27578]: #29155 +#29158 := [trans #29164 #29154]: #29157 +#29167 := [quant-inst]: #29145 +#29159 := [mp #29167 #29158]: #29144 +#29748 := [unit-resolution #29159 #21444 #29747 #29742 #28031]: #27568 +#29175 := (or #27565 #27536) +#29176 := [def-axiom]: #29175 +#29749 := [unit-resolution #29176 #29748]: #27536 +#29780 := [symm #29749]: #29779 +#29777 := (= #27762 #27535) +#29775 := (= #27163 #27532) +#29773 := (= #27532 #27163) +#29771 := (= #27531 #3175) +#27310 := (uf_116 #3175) +#27388 := (uf_43 #24114 #27310) +#29765 := (= #27388 #3175) +#27429 := (= #3175 #27388) +#27431 := (or #27309 #27429) +#29044 := (or #25416 #27309 #27429) +#27430 := (or #27429 #27309) +#29045 := (or #25416 #27430) +#28949 := (iff #29045 #29044) +#28962 := (or #25416 #27431) +#28965 := (iff #28962 #29044) +#28948 := [rewrite]: #28965 +#28960 := (iff #29045 #28962) +#27432 := (iff #27430 #27431) +#27433 := [rewrite]: #27432 +#28963 := [monotonicity #27433]: #28960 +#28945 := [trans #28963 #28948]: #28949 +#28961 := [quant-inst]: #29045 +#28950 := [mp #28961 #28945]: #29044 +#29754 := [unit-resolution #28950 #18736]: #27431 +#29755 := [unit-resolution #29754 #29753]: #27429 +#29766 := [symm #29755]: #29765 +#29769 := (= #27531 #27388) +#28122 := (+ #26643 #27529) +#28146 := (+ #26356 #28122) +#28149 := (uf_43 #24114 #28146) +#29763 := (= #28149 #27388) +#29757 := (= #28146 #27310) +#29735 := (= #27310 #28146) +#29736 := (* -1::int #28146) +#29737 := (+ #27310 #29736) +#29738 := (<= #29737 0::int) +#27645 := (uf_116 #27493) +#27649 := (* -1::int #27645) +#29118 := (+ #27310 #27649) +#29119 := (<= #29118 0::int) +#29091 := (= #27310 #27645) +#29592 := (= #27645 #27310) +#29585 := (= #27493 #3175) +#29612 := (= #27493 #27480) +#29613 := [symm #29049]: #29612 +#29591 := [trans #29613 #29055]: #29585 +#29588 := [monotonicity #29591]: #29592 +#29593 := [symm #29588]: #29091 +#29604 := (not #29091) +#29605 := (or #29604 #29119) +#29606 := [th-lemma]: #29605 +#29607 := [unit-resolution #29606 #29593]: #29119 +#27650 := (+ #26357 #27649) +#27651 := (+ #26356 #27650) +#27652 := (+ uf_294 #27651) +#29185 := (>= #27652 0::int) +#27653 := (= #27652 0::int) +#29089 := (or #24187 #27653) +#27646 := (= #27490 #27645) +#29085 := (or #24187 #27646) +#29108 := (iff #29085 #29089) +#29110 := (iff #29089 #29089) +#29111 := [rewrite]: #29110 +#27654 := (iff #27646 #27653) +#27655 := [rewrite]: #27654 +#29109 := [monotonicity #27655]: #29108 +#29112 := [trans #29109 #29111]: #29108 +#29086 := [quant-inst]: #29085 +#29107 := [mp #29086 #29112]: #29089 +#29608 := [unit-resolution #29107 #23688]: #27653 +#29595 := (not #27653) +#29596 := (or #29595 #29185) +#29597 := [th-lemma]: #29596 +#29598 := [unit-resolution #29597 #29608]: #29185 +#29601 := (not #27668) +#29600 := (not #27892) +#29599 := (not #27068) +#29594 := (not #29185) +#29586 := (not #29119) +#29602 := (or #29738 #29586 #29594 #29599 #29600 #29601) +#29603 := [th-lemma]: #29602 +#29625 := [unit-resolution #29603 #28258 #27679 #29598 #29607 #28054]: #29738 +#29739 := (>= #29737 0::int) +#29120 := (>= #29118 0::int) +#29626 := (or #29604 #29120) +#29616 := [th-lemma]: #29626 +#29614 := [unit-resolution #29616 #29593]: #29120 +#29196 := (<= #27652 0::int) +#29617 := (or #29595 #29196) +#29618 := [th-lemma]: #29617 +#29619 := [unit-resolution #29618 #29608]: #29196 +#29627 := (not #27893) +#29624 := (not #27073) +#29621 := (not #29196) +#29620 := (not #29120) +#29623 := (or #29739 #29620 #29621 #29624 #29627 #27714) +#29628 := [th-lemma]: #29623 +#29629 := [unit-resolution #29628 #28051 #28271 #29619 #29614 #27711]: #29739 +#29503 := (not #29739) +#29630 := (not #29738) +#29518 := (or #29735 #29630 #29503) +#29532 := [th-lemma]: #29518 +#29517 := [unit-resolution #29532 #29629 #29625]: #29735 +#29263 := [symm #29517]: #29757 +#29472 := [monotonicity #29263]: #29763 +#29767 := (= #27531 #28149) +#28104 := (uf_66 #25404 #27529 #24114) +#28136 := (= #28104 #28149) +#28137 := (not #28136) +#28107 := (uf_139 #28104 #25404) +#28108 := (= uf_9 #28107) +#28109 := (not #28108) +#28145 := (or #28109 #28137) +#28249 := (not #28145) +#29261 := (or #26114 #28249) +#28110 := (+ #27529 #26644) +#28111 := (uf_43 #24114 #28110) +#28112 := (= #28104 #28111) +#28117 := (not #28112) +#28118 := (or #28117 #28109) +#28121 := (not #28118) +#29262 := (or #26114 #28121) +#29295 := (iff #29262 #29261) +#29335 := (iff #29261 #29261) +#29336 := [rewrite]: #29335 +#28250 := (iff #28121 #28249) +#28247 := (iff #28118 #28145) +#28142 := (or #28137 #28109) +#28156 := (iff #28142 #28145) +#28157 := [rewrite]: #28156 +#28143 := (iff #28118 #28142) +#28140 := (iff #28117 #28137) +#28138 := (iff #28112 #28136) +#28150 := (= #28111 #28149) +#28147 := (= #28110 #28146) +#28148 := [rewrite]: #28147 +#28151 := [monotonicity #28148]: #28150 +#28139 := [monotonicity #28151]: #28138 +#28141 := [monotonicity #28139]: #28140 +#28144 := [monotonicity #28141]: #28143 +#28248 := [trans #28144 #28157]: #28247 +#28251 := [monotonicity #28248]: #28250 +#29334 := [monotonicity #28251]: #29295 +#29337 := [trans #29334 #29336]: #29295 +#29294 := [quant-inst]: #29262 +#29338 := [mp #29294 #29337]: #29261 +#29759 := [unit-resolution #29338 #21660]: #28249 +#29340 := (or #28145 #28136) +#29279 := [def-axiom]: #29340 +#29760 := [unit-resolution #29279 #29759]: #28136 +#29761 := (= #27531 #28104) +#29762 := [monotonicity #27948]: #29761 +#29768 := [trans #29762 #29760]: #29767 +#29473 := [trans #29768 #29472]: #29769 +#29499 := [trans #29473 #29766]: #29771 +#29656 := [monotonicity #29499]: #29773 +#29657 := [symm #29656]: #29775 +#29180 := [monotonicity #29657]: #29777 +#29677 := [trans #29180 #29780]: #29781 +#29199 := [trans #29677 #27935]: #29783 +#29181 := [monotonicity #29199]: #29264 +#29266 := [trans #29181 #28359]: #29265 +#29460 := [monotonicity #29266]: #29267 +#29462 := [monotonicity #29460]: #29461 +#29680 := [symm #29462]: #29679 +#29682 := [monotonicity #29680]: #29681 +#29683 := [mp #14796 #29682]: #29563 +#29170 := (not #27607) +#29696 := (iff #29170 #27761) +#29689 := (iff #27607 #27760) +#29693 := (iff #27760 #27607) +#29691 := (= #27759 #27606) +#29688 := (= #27163 #27605) +#29686 := (= #27605 #27163) +#29687 := [monotonicity #27692]: #29686 +#29690 := [symm #29687]: #29688 +#29692 := [monotonicity #29690]: #29691 +#29694 := [monotonicity #29692]: #29693 +#29695 := [symm #29694]: #29689 +#29697 := [monotonicity #29695]: #29696 +#29684 := [unit-resolution #27635 #22104 #14784 #29000 #29747 #29742 #27723]: #27614 +#29173 := (or #27611 #29170) +#29169 := [def-axiom]: #29173 +#29685 := [unit-resolution #29169 #29684]: #29170 +#29698 := [mp #29685 #29697]: #27761 +#29577 := (or #27819 #27760) +#29578 := [def-axiom]: #29577 +#29699 := [unit-resolution #29578 #29698]: #27819 +#29709 := (or #27833 #27770 #27810 #27822) +#29792 := (not #29735) +#29793 := (or #29792 #27772) +#29788 := (= #2967 #27771) +#29785 := (= #27771 #2967) +#29756 := [hypothesis]: #29735 +#29758 := [symm #29756]: #29757 +#29764 := [monotonicity #29758]: #29763 +#29770 := [trans #29768 #29764]: #29769 +#29772 := [trans #29770 #29766]: #29771 +#29774 := [monotonicity #29772]: #29773 +#29776 := [symm #29774]: #29775 +#29778 := [monotonicity #29776]: #29777 +#29782 := [trans #29778 #29780]: #29781 +#29784 := [trans #29782 #27935]: #29783 +#29786 := [monotonicity #29784]: #29785 +#29789 := [symm #29786]: #29788 +#29790 := [trans #14799 #29789]: #27772 +#29458 := (not #27772) +#29740 := [hypothesis]: #29458 +#29791 := [unit-resolution #29740 #29790]: false +#29794 := [lemma #29791]: #29793 +#29702 := [unit-resolution #29794 #29517]: #27772 +#29528 := (or #27813 #29458) +#29529 := [def-axiom]: #29528 +#29703 := [unit-resolution #29529 #29702]: #27813 +#29571 := (or #27833 #27770 #27810 #27816 #27822) +#29572 := [def-axiom]: #29571 +#29710 := [unit-resolution #29572 #29703]: #29709 +#29711 := [unit-resolution #29710 #29699 #29683 #29188 #29637]: false +#29712 := [lemma #29711]: #12812 +#23247 := (not #19374) +#29440 := [hypothesis]: #23803 +#23427 := (or #23812 #23800) +#23522 := [def-axiom]: #23427 +#29504 := [unit-resolution #23522 #29440]: #23812 +#23396 := (or #23806 #23800) +#23538 := [def-axiom]: #23396 +#29505 := [unit-resolution #23538 #29440]: #23806 +#29542 := (or #23818 #23809) +#23403 := (or #23906 #14046) +#23404 := [def-axiom]: #23403 +#29536 := [unit-resolution #23404 #29231]: #23906 +#29537 := [unit-resolution #23380 #29536 #29225]: #23875 +#23447 := (or #23872 #23866) +#23448 := [def-axiom]: #23447 +#29538 := [unit-resolution #23448 #29537]: #23866 +#27307 := (or #23818 #23809 #19008 #23869) +#27389 := [hypothesis]: #23821 +#23428 := (or #23818 #12812) +#23429 := [def-axiom]: #23428 +#27390 := [unit-resolution #23429 #27389]: #12812 +#23411 := (or #23818 #12806) +#23426 := [def-axiom]: #23411 +#27391 := [unit-resolution #23426 #27389]: #12806 +#27392 := [hypothesis]: #12803 +#27387 := [hypothesis]: #23866 +#23466 := (or #23869 #19008 #19011 #23863) +#23461 := [def-axiom]: #23466 +#27393 := [unit-resolution #23461 #27391 #27387 #27392]: #23863 +#23475 := (or #23860 #23854) +#23470 := [def-axiom]: #23475 +#27394 := [unit-resolution #23470 #27393]: #23854 +#23468 := (or #23857 #19011 #19017 #23851) +#23469 := [def-axiom]: #23468 +#27395 := [unit-resolution #23469 #27394 #27391 #27390]: #23851 +#27396 := [hypothesis]: #23806 +#23528 := (or #23824 #23818) +#23515 := [def-axiom]: #23528 +#27397 := [unit-resolution #23515 #27389]: #23824 +#23521 := (or #23833 #19008 #19011 #23827) +#23510 := [def-axiom]: #23521 +#27304 := [unit-resolution #23510 #27397 #27392 #27391]: #23833 +#23499 := (or #23836 #23830) +#23501 := [def-axiom]: #23499 +#27305 := [unit-resolution #23501 #27304]: #23836 +#23492 := (or #23845 #23809 #23839) +#23494 := [def-axiom]: #23492 +#27306 := [unit-resolution #23494 #27305 #27396]: #23845 +#23482 := (or #23848 #23842) +#23483 := [def-axiom]: #23482 +#27269 := [unit-resolution #23483 #27306 #27395]: false +#27303 := [lemma #27269]: #27307 +#29521 := [unit-resolution #27303 #29001 #29538]: #29542 +#29525 := [unit-resolution #29521 #29505]: #23818 +#29520 := (or #23821 #19017 #23815) +#23431 := (or #23821 #19011 #19017 #23815) +#23432 := [def-axiom]: #23431 +#29526 := [unit-resolution #23432 #29138]: #29520 +#29527 := [unit-resolution #29526 #29525 #29504 #29712]: false +#29530 := [lemma #29527]: #23800 +#29896 := (or #23803 #23797) +#16270 := (<= uf_272 131073::int) +#16273 := (iff #13872 #16270) +#16264 := (+ 131073::int #13873) +#16267 := (>= #16264 0::int) +#16271 := (iff #16267 #16270) +#16272 := [rewrite]: #16271 +#16268 := (iff #13872 #16267) +#16265 := (= #13874 #16264) +#16266 := [monotonicity #7888]: #16265 +#16269 := [monotonicity #16266]: #16268 +#16274 := [trans #16269 #16272]: #16273 +#14787 := [not-or-elim #14776]: #13880 +#14788 := [and-elim #14787]: #13872 +#16275 := [mp #14788 #16274]: #16270 +#29232 := [hypothesis]: #19037 +#29233 := [th-lemma #29232 #29231 #16275]: false +#29234 := [lemma #29233]: #16368 +#29894 := (or #23803 #19037 #23797) +#29891 := (or #14243 #14088) +#29892 := [th-lemma]: #29891 +#29893 := [unit-resolution #29892 #28992]: #14088 +#23558 := (or #23803 #19034 #19037 #23797) +#23555 := [def-axiom]: #23558 +#29895 := [unit-resolution #23555 #29893]: #29894 +#29897 := [unit-resolution #29895 #29234]: #29896 +#29898 := [unit-resolution #29897 #29530]: #23797 +#23561 := (or #23794 #23788) +#23565 := [def-axiom]: #23561 +#29899 := [unit-resolution #23565 #29898]: #23788 +#23271 := (>= #14169 -1::int) +#23285 := (or #23794 #14168) +#23286 := [def-axiom]: #23285 +#29900 := [unit-resolution #23286 #29898]: #14168 +#29901 := (or #14172 #23271) +#29902 := [th-lemma]: #29901 +#29903 := [unit-resolution #29902 #29900]: #23271 +#29238 := (not #23271) +#29239 := (or #14100 #29238) +#29201 := [hypothesis]: #23271 +#29202 := [hypothesis]: #14105 +#29237 := [th-lemma #29202 #29231 #29201]: false +#29240 := [lemma #29237]: #29239 +#29904 := [unit-resolution #29240 #29903]: #14100 +#23580 := (or #23791 #14105 #23785) +#23566 := [def-axiom]: #23580 +#29905 := [unit-resolution #23566 #29904 #29899]: #23785 +#23575 := (or #23782 #23776) +#23213 := [def-axiom]: #23575 +#29906 := [unit-resolution #23213 #29905]: #23776 +#29920 := (= #3068 #3209) +#29918 := (= #3209 #3068) +#29914 := (= #3208 #3067) +#29912 := (= #3208 #27327) +#29910 := (= uf_301 #26964) +#29907 := [hypothesis]: #23809 +#23549 := (or #23806 #12826) +#23550 := [def-axiom]: #23549 +#29908 := [unit-resolution #23550 #29907]: #12826 +#29909 := [symm #29908]: #3189 +#29911 := [trans #29909 #27438]: #29910 +#29913 := [monotonicity #29911]: #29912 +#29915 := [trans #29913 #27639]: #29914 +#29919 := [monotonicity #29915]: #29918 +#29921 := [symm #29919]: #29920 +#29922 := (= uf_300 #3068) +#23559 := (or #23806 #12823) +#23548 := [def-axiom]: #23559 +#29916 := [unit-resolution #23548 #29907]: #12823 +#29917 := [symm #29916]: #3187 +#29923 := [trans #29917 #29222]: #29922 +#29924 := [trans #29923 #29921]: #12862 +#28863 := (+ uf_293 #14142) +#28865 := (>= #28863 0::int) +#29925 := (or #12993 #28865) +#29926 := [th-lemma]: #29925 +#29927 := [unit-resolution #29926 #29908]: #28865 +#29483 := (not #28865) +#29930 := (or #14145 #29483) +#29928 := (or #14145 #14404 #29483) +#29929 := [th-lemma]: #29928 +#29931 := [unit-resolution #29929 #29223]: #29930 +#29932 := [unit-resolution #29931 #29927]: #14145 +#23374 := (or #22784 #22782 #14144) +#23581 := [def-axiom]: #23374 +#29933 := [unit-resolution #23581 #29932 #29924]: #22784 +#23255 := (or #23770 #22783) +#23256 := [def-axiom]: #23255 +#29934 := [unit-resolution #23256 #29933]: #23770 +#23572 := (or #23779 #23773 #22836) +#23573 := [def-axiom]: #23572 +#29935 := [unit-resolution #23573 #29934 #29906]: #22836 +#23583 := (or #22831 #23247) +#23243 := [def-axiom]: #23583 +#29936 := [unit-resolution #23243 #29935]: #23247 +#29307 := (+ uf_294 #19372) +#29860 := (>= #29307 0::int) +#29955 := (not #29860) +#29855 := (= uf_294 ?x785!14) +#29888 := (not #29855) +#29858 := (= #3184 #19060) +#29864 := (not #29858) +#29859 := (+ #3184 #19385) +#29861 := (>= #29859 0::int) +#29871 := (not #29861) +#23394 := (or #23806 #14207) +#23395 := [def-axiom]: #23394 +#29937 := [unit-resolution #23395 #29907]: #14207 +#29541 := (+ uf_292 #14120) +#29539 := (<= #29541 0::int) +#29938 := (or #13002 #29539) +#29939 := [th-lemma]: #29938 +#29940 := [unit-resolution #29939 #29916]: #29539 +#23227 := (or #22831 #23584) +#23568 := [def-axiom]: #23227 +#29941 := [unit-resolution #23568 #29935]: #23584 +#29872 := (not #29539) +#29873 := (or #29871 #19387 #29872 #14211) +#29866 := [hypothesis]: #14207 +#29867 := [hypothesis]: #29539 +#29868 := [hypothesis]: #23584 +#29869 := [hypothesis]: #29861 +#29870 := [th-lemma #29869 #29868 #29867 #29866]: false +#29874 := [lemma #29870]: #29873 +#29942 := [unit-resolution #29874 #29941 #29940 #29937]: #29871 +#29865 := (or #29864 #29861) +#29875 := [th-lemma]: #29865 +#29943 := [unit-resolution #29875 #29942]: #29864 +#29889 := (or #29888 #29858) +#29884 := (= #19060 #3184) +#29882 := (= #19059 #3175) +#29880 := (= ?x785!14 uf_294) +#29879 := [hypothesis]: #29855 +#29881 := [symm #29879]: #29880 +#29883 := [monotonicity #29881]: #29882 +#29885 := [monotonicity #29883]: #29884 +#29886 := [symm #29885]: #29858 +#29878 := [hypothesis]: #29864 +#29887 := [unit-resolution #29878 #29886]: false +#29890 := [lemma #29887]: #29889 +#29944 := [unit-resolution #29890 #29943]: #29888 +#29958 := (or #29855 #29955) +#29308 := (<= #29307 0::int) +#29319 := (+ uf_292 #19385) +#29320 := (>= #29319 0::int) +#29945 := (not #29320) +#29946 := (or #29945 #19387 #29872) +#29947 := [th-lemma]: #29946 +#29948 := [unit-resolution #29947 #29940 #29941]: #29945 +#29951 := (or #29308 #29320) +#23582 := (or #22831 #19056) +#23242 := [def-axiom]: #23582 +#29949 := [unit-resolution #23242 #29935]: #19056 +#23586 := (or #22831 #19055) +#23592 := [def-axiom]: #23586 +#29950 := [unit-resolution #23592 #29935]: #19055 +#29802 := (or #23759 #22815 #22816 #29308 #29320) +#29296 := (+ #19060 #14431) +#29297 := (<= #29296 0::int) +#29298 := (+ ?x785!14 #14044) +#29299 := (>= #29298 0::int) +#29300 := (or #22816 #29299 #29297 #22815) +#29803 := (or #23759 #29300) +#29810 := (iff #29803 #29802) +#29328 := (or #22815 #22816 #29308 #29320) +#29805 := (or #23759 #29328) +#29808 := (iff #29805 #29802) +#29809 := [rewrite]: #29808 +#29806 := (iff #29803 #29805) +#29331 := (iff #29300 #29328) +#29325 := (or #22816 #29308 #29320 #22815) +#29329 := (iff #29325 #29328) +#29330 := [rewrite]: #29329 +#29326 := (iff #29300 #29325) +#29323 := (iff #29297 #29320) +#29313 := (+ #14431 #19060) +#29316 := (<= #29313 0::int) +#29321 := (iff #29316 #29320) +#29322 := [rewrite]: #29321 +#29317 := (iff #29297 #29316) +#29314 := (= #29296 #29313) +#29315 := [rewrite]: #29314 +#29318 := [monotonicity #29315]: #29317 +#29324 := [trans #29318 #29322]: #29323 +#29311 := (iff #29299 #29308) +#29301 := (+ #14044 ?x785!14) +#29304 := (>= #29301 0::int) +#29309 := (iff #29304 #29308) +#29310 := [rewrite]: #29309 +#29305 := (iff #29299 #29304) +#29302 := (= #29298 #29301) +#29303 := [rewrite]: #29302 +#29306 := [monotonicity #29303]: #29305 +#29312 := [trans #29306 #29310]: #29311 +#29327 := [monotonicity #29312 #29324]: #29326 +#29332 := [trans #29327 #29330]: #29331 +#29807 := [monotonicity #29332]: #29806 +#29811 := [trans #29807 #29809]: #29810 +#29804 := [quant-inst]: #29803 +#29812 := [mp #29804 #29811]: #29802 +#29952 := [unit-resolution #29812 #29221 #29950 #29949]: #29951 +#29953 := [unit-resolution #29952 #29948]: #29308 +#29954 := (not #29308) +#29956 := (or #29855 #29954 #29955) +#29957 := [th-lemma]: #29956 +#29959 := [unit-resolution #29957 #29953]: #29958 +#29960 := [unit-resolution #29959 #29944]: #29955 +#29961 := [th-lemma #29903 #29960 #29936]: false +#29962 := [lemma #29961]: #23806 +#29633 := [unit-resolution #29521 #29962]: #23818 +#29615 := [unit-resolution #29526 #29633 #29712]: #23815 +#23534 := (or #23812 #13075) +#23416 := [def-axiom]: #23534 +#29713 := [unit-resolution #23416 #29615]: #13075 +#29142 := (or #13081 #23536) +#29708 := [th-lemma]: #29142 +#29714 := [unit-resolution #29708 #29713]: #23536 +#29715 := [hypothesis]: #14144 +#29716 := [th-lemma #29715 #29714 #29231]: false +#29717 := [lemma #29716]: #14145 +#29990 := (or #22784 #14144) +#29985 := (= #3184 #3209) +#29982 := (= #3209 #3184) +#29979 := (= #3208 #3175) +#29978 := [symm #29713]: #3247 +#29980 := [monotonicity #29978]: #29979 +#29983 := [monotonicity #29980]: #29982 +#29986 := [symm #29983]: #29985 +#29987 := (= uf_300 #3184) +#23533 := (or #23812 #13070) +#23531 := [def-axiom]: #23533 +#29977 := [unit-resolution #23531 #29615]: #13070 +#29984 := [symm #29977]: #3240 +#23373 := (or #23812 #3246) +#23375 := [def-axiom]: #23373 +#29981 := [unit-resolution #23375 #29615]: #3246 +#29988 := [trans #29981 #29984]: #29987 +#29989 := [trans #29988 #29986]: #12862 +#29991 := [unit-resolution #23581 #29989]: #29990 +#29992 := [unit-resolution #29991 #29717]: #22784 +#29993 := [unit-resolution #23256 #29992]: #23770 +#29994 := [unit-resolution #23573 #29906]: #23776 +#29995 := [unit-resolution #29994 #29993]: #22836 +#30004 := [unit-resolution #23568 #29995]: #23584 +#30026 := (or #29945 #19387) +#29570 := (+ #3184 #14120) +#29587 := (<= #29570 0::int) +#29569 := (= #3184 uf_300) +#30005 := (= uf_304 uf_300) +#30006 := [symm #29981]: #30005 +#30007 := [trans #29977 #30006]: #29569 +#30008 := (not #29569) +#30009 := (or #30008 #29587) +#30010 := [th-lemma]: #30009 +#30011 := [unit-resolution #30010 #30007]: #29587 +#30016 := (or #19017 #23851) +#30012 := (or #19011 #23863) +#30013 := [unit-resolution #23461 #29001 #29538]: #30012 +#30014 := [unit-resolution #30013 #29138]: #23863 +#30015 := [unit-resolution #23470 #30014]: #23854 +#30017 := [unit-resolution #23469 #29138 #30015]: #30016 +#30018 := [unit-resolution #30017 #29712]: #23851 +#30019 := [unit-resolution #23483 #30018]: #23842 +#30020 := [unit-resolution #23494 #29962 #30019]: #23839 +#23514 := (or #23836 #14211) +#23498 := [def-axiom]: #23514 +#30021 := [unit-resolution #23498 #30020]: #14211 +#30022 := (not #29587) +#30023 := (or #29539 #14207 #30022) +#30024 := [th-lemma]: #30023 +#30025 := [unit-resolution #30024 #30021 #30011]: #29539 +#30027 := [unit-resolution #29947 #30025]: #30026 +#30028 := [unit-resolution #30027 #30004]: #29945 +#30029 := [unit-resolution #23242 #29995]: #19056 +#30030 := [unit-resolution #23592 #29995]: #19055 +#30031 := [unit-resolution #29812 #29221 #30030 #30029 #30028]: #29308 +#29996 := [unit-resolution #23243 #29995]: #23247 +#29997 := [hypothesis]: #29955 +#29998 := [th-lemma #29903 #29997 #29996]: false +#29999 := [lemma #29998]: #29860 +#30032 := [unit-resolution #29957 #29999 #30031]: #29855 +#30033 := [unit-resolution #29890 #30032]: #29858 +#30034 := [unit-resolution #29875 #30033]: #29861 +[th-lemma #30011 #30034 #30004]: false +unsat diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/Boogie/ROOT.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Boogie/ROOT.ML Wed Nov 04 11:40:59 2009 +0100 @@ -0,0 +1,1 @@ +use_thy "Boogie"; diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/Boogie/Tools/boogie_commands.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Boogie/Tools/boogie_commands.ML Wed Nov 04 11:40:59 2009 +0100 @@ -0,0 +1,90 @@ +(* 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 + +fun boogie_load (verbose, base_name) thy = + let + val path = Path.explode (base_name ^ ".b2i") + val _ = File.exists path orelse + error ("Unable to find file " ^ quote (Path.implode path)) + val _ = Boogie_VCs.is_closed thy orelse + error ("Found the beginning of a new Boogie environment, " ^ + "but another Boogie environment is still open.") + in Boogie_Loader.load_b2i verbose path thy end + +fun boogie_status thy = + let + fun pretty_vc_name (name, _, proved) = + let val s = if proved then "proved" else "not proved" + in Pretty.str (name ^ " (" ^ s ^ ")") end + in + Pretty.writeln (Pretty.big_list "Boogie verification conditions:" + (map pretty_vc_name (Boogie_VCs.as_list thy))) + end + +fun boogie_vc name lthy = + (case Boogie_VCs.lookup (ProofContext.theory_of lthy) name of + NONE => error ("There is no verification condition " ^ quote name ^ ".") + | SOME vc => + let + fun store thm = Boogie_VCs.discharge (name, thm) + fun after_qed [[thm]] = LocalTheory.theory (store thm) + | after_qed _ = I + in lthy |> Proof.theorem_i NONE after_qed [[(vc, [])]] end) + +fun boogie_end thy = + let + fun not_proved (name, _, proved) = if not proved then SOME name else NONE + val unproved = map_filter not_proved (Boogie_VCs.as_list 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 + +val _ = + OuterSyntax.command "boogie_open" + "Open a new Boogie environment and load a Boogie-generated .b2i file." + OuterKeyword.thy_decl + (Scan.optional (Args.parens (Args.$$$ "quiet") >> K false) true -- + OuterParse.name >> (Toplevel.theory o boogie_load)) + +val _ = + OuterSyntax.improper_command "boogie_status" + "Show the name and state of all loaded verification conditions." + OuterKeyword.diag + (Scan.succeed (Toplevel.no_timing o Toplevel.keep (fn state => + boogie_status (Toplevel.theory_of state)))) + +val _ = + OuterSyntax.command "boogie_vc" + "Enter the proof mode for a specific verification condition." + OuterKeyword.thy_goal + (OuterParse.name >> (fn name => Toplevel.print o + Toplevel.local_theory_to_proof NONE (boogie_vc name))) + +val _ = + OuterSyntax.command "boogie_end" + "Close the current Boogie environment." + OuterKeyword.thy_decl + (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 diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/Boogie/Tools/boogie_loader.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Boogie/Tools/boogie_loader.ML Wed Nov 04 11:40:59 2009 +0100 @@ -0,0 +1,475 @@ +(* 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 -> Path.T -> theory -> theory +end + +structure Boogie_Loader: BOOGIE_LOADER = +struct + +fun log verbose text args thy = + if verbose + then (Pretty.writeln (Pretty.big_list text (map Pretty.str args)); thy) + else thy + +val isabelle_name = + let + fun purge s = if Symbol.is_letdig 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 + +datatype attribute_value = StringValue of string | TermValue of Term.term + + + +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 declare (name, arity) thy = + let val isa_name = isabelle_name name + in + (case lookup_type_name thy isa_name arity of + SOME type_name => ((type_name, false), thy) + | NONE => + let + val bname = Binding.name isa_name + val args = Name.variant_list [] (replicate arity "'") + val (T, thy') = + ObjectLogic.typedecl (Binding.name isa_name, args, NoSyn) thy + val type_name = fst (Term.dest_Type T) + in ((type_name, true), thy') end) + end + + fun type_names ((name, _), (new_name, new)) = + if new then SOME (new_name ^ " (was " ^ name ^ ")") else NONE +in +fun declare_types verbose tys = + fold_map declare tys #-> (fn tds => + log verbose "Declared types:" (map_filter type_names (tys ~~ tds)) #> + rpair (Symtab.make (map fst tys ~~ map fst tds))) +end + + + +local + val quote_mixfix = + Symbol.explode #> map + (fn "_" => "'_" + | "(" => "'(" + | ")" => "')" + | "/" => "'/" + | s => s) #> + implode + + fun mk_syntax name i = + let + val syn = quote_mixfix name + val args = concat (separate ",/ " (replicate i "_")) + in + if i = 0 then Mixfix (syn, [], 1000) + else Mixfix (syn ^ "()'(/" ^ args ^ "')", replicate i 0, 1000) + end + + fun process_attributes T = + let + fun const name = SOME (Const (name, T)) + + fun choose builtin = + (case builtin of + "bvnot" => const @{const_name z3_bvnot} + | "bvand" => const @{const_name z3_bvand} + | "bvor" => const @{const_name z3_bvor} + | "bvxor" => const @{const_name z3_bvxor} + | "bvadd" => const @{const_name z3_bvadd} + | "bvneg" => const @{const_name z3_bvneg} + | "bvsub" => const @{const_name z3_bvsub} + | "bvmul" => const @{const_name z3_bvmul} + | "bvudiv" => const @{const_name z3_bvudiv} + | "bvurem" => const @{const_name z3_bvurem} + | "bvsdiv" => const @{const_name z3_bvsdiv} + | "bvsrem" => const @{const_name z3_bvsrem} + | "bvshl" => const @{const_name z3_bvshl} + | "bvlshr" => const @{const_name z3_bvlshr} + | "bvashr" => const @{const_name z3_bvashr} + | "bvult" => const @{const_name z3_bvult} + | "bvule" => const @{const_name z3_bvule} + | "bvugt" => const @{const_name z3_bvugt} + | "bvuge" => const @{const_name z3_bvuge} + | "bvslt" => const @{const_name z3_bvslt} + | "bvsle" => const @{const_name z3_bvsle} + | "bvsgt" => const @{const_name z3_bvsgt} + | "bvsge" => const @{const_name z3_bvsge} + | "sign_extend" => const @{const_name z3_sign_extend} + | _ => 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 + val is_type_instance = Type.typ_instance o Sign.tsig_of + in + if Sign.declared_const thy intern + then + if is_type_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 declare (name, ((Ts, T), atts)) thy = + let val isa_name = isabelle_name name and U = Ts ---> T + in + (case lookup_const thy isa_name U of + SOME t => (((name, t), false), thy) + | NONE => + (case process_attributes U atts of + SOME t => (((name, t), false), thy) + | NONE => + thy + |> Sign.declare_const ((Binding.name isa_name, U), + mk_syntax name (length Ts)) + |> apfst (rpair true o pair name))) + end + + fun const_names ((name, _), ((_, t), new)) = + if new then SOME (fst (Term.dest_Const t) ^ " (as " ^ name ^ ")") else NONE +in +fun declare_functions verbose fns = + fold_map declare fns #-> (fn fds => + log verbose "Declared constants:" (map_filter const_names (fns ~~ fds)) #> + rpair (Symtab.make (map fst 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 + + fun only_new_boogie_axioms thy = + let val baxs = map Thm.prop_of (Boogie_Axioms.get (ProofContext.init thy)) + in filter_out (member (op aconv) baxs o snd) end +in +fun add_axioms verbose axs thy = + let val axs' = only_new_boogie_axioms thy (name_axioms axs) + in + thy + |> PureThy.add_axioms (map (rpair [] o apfst Binding.name) axs') + |-> Context.theory_map o fold Boogie_Axioms.add_thm + |> log verbose "The following axioms were added:" (map fst axs') + |> Context.theory_map (fn context => fold Split_VC_SMT_Rules.add_thm + (Boogie_Axioms.get (Context.proof_of context)) context) + end +end + + + +fun add_vcs verbose vcs thy = + let + val reused = duplicates (op =) (map (fst o fst) vcs) + fun rename (n, i) = isabelle_name n ^ + (if member (op =) reused n then "_" ^ string_of_int i else "") + + fun decorate (name, t) = (rename name, HOLogic.mk_Trueprop t) + val vcs' = map decorate vcs + in + thy + |> Boogie_VCs.set vcs' + |> log verbose "The following verification conditions were loaded:" + (map fst vcs') + 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 path = + 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) (File.fold_lines split path (1, [])) end + +fun stopper i = Scan.stopper (K (i, EOF)) (fn (_, EOF) => true | _ => false) + +fun scan_err msg [] = "Boogie (at end of input): " ^ msg + | scan_err msg ((i, _) :: _) = "Boogie (at line " ^ string_of_int i ^ "): " ^ + msg + +fun scan_fail msg = Scan.fail_with (scan_err msg) + +fun finite scan path = + let val (i, ts) = tokenize path + in + (case Scan.error (Scan.finite (stopper i) scan) ts of + (x, []) => x + | (_, ts) => error (scan_err "unparsed input" ts)) + end + +fun read_int' s = (case read_int (explode s) of (i, []) => SOME i | _ => NONE) + +fun $$$ s = Scan.one (fn (_, Token s') => s = s' | _ => false) +val str = Scan.some (fn (_, Token s) => SOME s | _ => NONE) +val num = Scan.some (fn (_, Token s) => read_int' s | _ => NONE) + +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 ("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 _ (t :: ts) = fold f ts t + + fun mk_distinct T ts = + Const (@{const_name distinct}, HOLogic.listT T --> @{typ bool}) $ + HOLogic.mk_list T ts + + 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 (n, T) t = q T $ Term.absfree (isabelle_name n, T, t) + + val patternT = @{typ pattern} + fun mk_pattern _ [] = raise TERM ("mk_pattern", []) + | mk_pattern n [t] = Const (n, Term.fastype_of t --> patternT) $ t + | mk_pattern n (t :: ts) = + let val T = patternT --> Term.fastype_of t --> patternT + in Const (@{const_name andpat}, T) $ mk_pattern n ts $ t end + fun patt n c scan = + scan_line n num :|-- scan_count scan >> (mk_pattern c) + fun pattern scan = + patt "pat" @{const_name pat} scan || + patt "nopat" @{const_name nopat} scan || + scan_fail "illegal pattern kind" + fun mk_trigger [] t = t + | mk_trigger ps t = @{term trigger} $ HOLogic.mk_list patternT ps $ t + + val label_kind = + $$$ "pos" >> K @{term block_at} || + $$$ "neg" >> K @{term assert_at} || + scan_fail "illegal label kind" + + fun mk_select (m, k) = + let + val mT = Term.fastype_of m and kT = Term.fastype_of k + val vT = Term.range_type mT + in Const (@{const_name boogie_select}, mT --> kT --> vT) $ m $ k 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 boogie_store}, 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} + val m = HOLogic.mk_number @{typ nat} + in + Const (@{const_name boogie_bv_extract}, [nT, nT, dT] ---> rT) $ + mk_nat_num msb $ 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 boogie_bv_concat}, [T1, T2] ---> U) $ t1 $ t2 end +in +fun expr 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} || + binexp "implies" (binop @{term "op -->"}) || + scan_line "distinct" num :|-- scan_count exp >> + (fn [] => @{term True} + | ts as (t :: _) => mk_distinct (Term.fastype_of t) ts) || + binexp "=" HOLogic.mk_eq || + scan_line "var" str -- typ tds >> (Free o apfst isabelle_name) || + 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" str -- typ tds) n -- + scan_count (pattern exp) k -- + scan_count (attribute tds fds) i -- + exp >> (fn (((vs, ps), _), t) => + fold_rev (mk_quant q) vs (mk_trigger ps t))) || + scan_line "label" (label_kind -- num -- num) -- exp >> + (fn (((l, line), col), t) => + if line = 0 orelse col = 0 then t + else l $ Free ("L_" ^ string_of_int line ^ "_" ^ string_of_int col, + @{typ bool}) $ 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) >> + mk_select) || + 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 end + +and attribute tds fds = + let + val attr_val = + scan_line' "expr-attr" |-- expr 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 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 tds Symtab.empty) i >> pair name)) >> + (fn fns => declare_functions verbose fns thy)) + +fun axioms verbose tds fds = Scan.depend (fn thy => + Scan.repeat (scan_line "axiom" num :|-- (fn i => + expr tds fds --| scan_count (attribute tds fds) i)) >> + (fn axs => (add_axioms verbose 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 tds fds) i >> K ())) >> K (thy, ())) + +fun vcs verbose tds fds = Scan.depend (fn thy => + Scan.repeat (scan_line "vc" (str -- num) -- + (expr tds fds >> Sign.cert_term thy)) >> + (fn vcs => ((), add_vcs verbose vcs thy))) + +fun parse verbose thy = Scan.pass thy + (type_decls verbose :|-- (fn tds => + fun_decls verbose tds :|-- (fn fds => + axioms verbose tds fds |-- + var_decls tds fds |-- + vcs verbose tds fds))) + +fun load_b2i verbose path thy = finite (parse verbose thy) path + +end diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/Boogie/Tools/boogie_split.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Boogie/Tools/boogie_split.ML Wed Nov 04 11:40:59 2009 +0100 @@ -0,0 +1,167 @@ +(* Title: HOL/Boogie/Tools/boogie_split.ML + Author: Sascha Boehme, TU Muenchen + +Method to split Boogie-generate verification conditions and pass the splinters +to user-supplied automated sub-tactics. +*) + +signature BOOGIE_SPLIT = +sig + val add_sub_tactic: string * (Proof.context -> int -> tactic) -> theory -> + theory + val sub_tactic_names: theory -> string list + val setup: theory -> theory +end + +structure Boogie_Split: BOOGIE_SPLIT = +struct + +(* sub-tactics store *) + +structure Sub_Tactics = TheoryDataFun +( + type T = (Proof.context -> int -> tactic) Symtab.table + val empty = Symtab.empty + val copy = I + val extend = I + fun merge _ = Symtab.merge (K true) +) + +fun add_sub_tactic tac = Sub_Tactics.map (Symtab.update_new tac) + +val sub_tactic_names = Symtab.keys o Sub_Tactics.get + +fun lookup_sub_tactic ctxt name = + (case Symtab.lookup (Sub_Tactics.get (ProofContext.theory_of ctxt)) name of + SOME tac => SOME (name, tac) + | NONE => (warning ("Unknown sub-tactic: " ^ quote name); NONE)) + +fun as_meta_eq eq = eq RS @{thm eq_reflection} + +fun full_implies_conv cv ct = + (case try Thm.dest_implies ct of + NONE => cv ct + | SOME (cp, cc) => Drule.imp_cong_rule (cv cp) (full_implies_conv cv cc)) + +fun sub_tactics_tac ctxt (verbose, sub_tac_names) case_name = + let + fun trace msg = if verbose then tracing msg else () + fun trying name = trace ("Trying tactic " ^ quote name ^ " on case " ^ + quote case_name ^ " ...") + fun solved () = trace ("Case solved: " ^ quote case_name) + fun failed () = trace ("Case remains unsolved: " ^ quote case_name) + + infix IF_UNSOLVED + fun (tac1 IF_UNSOLVED tac2) i st = st |> (tac1 i THEN (fn st' => + let val j = i + Thm.nprems_of st' - Thm.nprems_of st + in + if i > j then (solved (); Tactical.all_tac st') + else Seq.INTERVAL tac2 i j st' + end)) + + fun TRY_ALL [] _ st = (failed (); Tactical.no_tac st) + | TRY_ALL ((name, tac) :: tacs) i st = (trying name; + (TRY o tac ctxt IF_UNSOLVED TRY_ALL tacs) i st) + + val unfold_labels = Conv.try_conv (Conv.arg_conv + (More_Conv.rewrs_conv (map as_meta_eq @{thms labels}))) + in + CONVERSION (full_implies_conv unfold_labels) + THEN' TRY_ALL (map_filter (lookup_sub_tactic ctxt) sub_tac_names) + end + + +(* case names *) + +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])) + +local + fun make_case_name (i, t) = + the_default ("B_" ^ string_of_int (i + 1)) (try case_name_of t) + + val assert_intro = Thm.symmetric (as_meta_eq @{thm assert_at_def}) + val l = Thm.dest_arg1 (Thm.rhs_of assert_intro) + fun intro new = + Conv.rewr_conv (Thm.instantiate ([], [(l, new)]) assert_intro) + + val assert_eq = @{lemma "assert_at x t == assert_at y t" + by (simp add: assert_at_def)} + val (x, y) = pairself Thm.dest_arg1 (Thm.dest_binop (Thm.cprop_of assert_eq)) + fun rename old new = + Conv.rewr_conv (Thm.instantiate ([], [(x, old), (y, new)]) assert_eq) + + fun at_concl cv = Conv.concl_conv ~1 (Conv.arg_conv cv) + fun make_label thy name = Thm.cterm_of thy (Free (name, @{typ bool})) + + fun rename_case thy new ct = + (case try case_name_of (Thm.term_of ct) of + NONE => at_concl (intro (make_label thy new)) + | SOME old => if new = old then Conv.all_conv + else at_concl (rename (make_label thy old) (make_label thy new))) ct +in +fun unique_case_names thy st = + let + val names = map_index make_case_name (Thm.prems_of st) + |> ` (duplicates (op =)) |-> Name.variant_list + in ALLGOALS (fn i => CONVERSION (rename_case thy (nth names (i-1))) i) st end +end + + +(* splitting method *) + +local + val split_rules = [@{thm impI}, @{thm conjI}, @{thm TrueI}] + + fun prep_tac ctxt args facts = + HEADGOAL (Method.insert_tac facts) + THEN HEADGOAL (REPEAT_ALL_NEW (Tactic.resolve_tac split_rules)) + THEN unique_case_names (ProofContext.theory_of ctxt) + THEN ALLGOALS (SUBGOAL (fn (t, i) => + TRY (sub_tactics_tac ctxt args (case_name_of t) i))) + + val drop_assert_at = Conv.concl_conv ~1 (Conv.try_conv (Conv.arg_conv + (Conv.rewr_conv (as_meta_eq @{thm assert_at_def})))) + + fun split_vc args ctxt = METHOD_CASES (fn facts => + prep_tac ctxt args facts #> + Seq.maps (fn st => + st + |> ALLGOALS (CONVERSION drop_assert_at) + |> Seq.map (pair (map (rpair [] o case_name_of) (Thm.prems_of st)))) #> + Seq.maps (fn (names, st) => + CASES + (Rule_Cases.make_common false + (ProofContext.theory_of ctxt, Thm.prop_of st) names) + Tactical.all_tac st)) + + val verbose_opt = Args.parens (Args.$$$ "verbose") >> K true + val sub_tactics = Args.$$$ "try" -- Args.colon |-- Scan.repeat1 Args.name +in +val setup_split_vc = Method.setup @{binding split_vc} + (Scan.lift (Scan.optional verbose_opt false -- Scan.optional sub_tactics []) + >> split_vc) + ("Splits a Boogie-generated verification conditions into smaller problems" ^ + " and tries to solve the splinters with the supplied sub-tactics.") +end + + +(* predefined sub-tactics *) + +fun fast_sub_tac ctxt = Classical.fast_tac (Classical.claset_of ctxt) +fun simp_sub_tac ctxt = + Simplifier.asm_full_simp_tac (Simplifier.simpset_of ctxt) +fun smt_sub_tac ctxt = SMT_Solver.smt_tac ctxt (Split_VC_SMT_Rules.get ctxt) + + +(* setup *) + +val setup = + setup_split_vc #> + add_sub_tactic ("fast", fast_sub_tac) #> + add_sub_tactic ("simp", simp_sub_tac) #> + add_sub_tactic ("smt", smt_sub_tac) + +end diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/Boogie/Tools/boogie_vcs.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Boogie/Tools/boogie_vcs.ML Wed Nov 04 11:40:59 2009 +0100 @@ -0,0 +1,60 @@ +(* Title: HOL/Boogie/Tools/boogie_vcs.ML + Author: Sascha Boehme, TU Muenchen + +Store for Boogie's verification conditions. +*) + +signature BOOGIE_VCS = +sig + val set: (string * term) list -> theory -> theory + val lookup: theory -> string -> term option + val discharge: string * thm -> theory -> theory + val close: theory -> theory + val is_closed: theory -> bool + val as_list: theory -> (string * term * bool) list +end + +structure Boogie_VCs: BOOGIE_VCS = +struct + +fun err_vcs () = error "undischarged Boogie verification conditions found" + +structure VCs = TheoryDataFun +( + type T = (Term.term * bool) Symtab.table option + val empty = NONE + val copy = I + val extend = I + fun merge _ (NONE, NONE) = NONE + | merge _ (_, _) = err_vcs () +) + +fun set vcs = VCs.map (fn + NONE => SOME (Symtab.make (map (apsnd (rpair false)) vcs)) + | SOME _ => err_vcs ()) + +fun lookup thy name = + (case VCs.get thy of + SOME vcs => Option.map fst (Symtab.lookup vcs name) + | NONE => NONE) + +fun discharge (name, thm) thy = VCs.map (fn + SOME vcs => SOME (Symtab.map_entry name (fn (t, proved) => + if proved then (t, proved) + else (t, Pattern.matches thy (Thm.prop_of thm, t))) vcs) + | NONE => NONE) thy + +val close = VCs.map (fn + SOME vcs => + if Symtab.exists (fn (_, (_, proved)) => not proved) vcs then err_vcs () + else NONE + | NONE => NONE) + +val is_closed = is_none o VCs.get + +fun as_list thy = + (case VCs.get thy of + SOME vcs => map (fn (n, (t, p)) => (n, t, p)) (Symtab.dest vcs) + | NONE => []) + +end diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/Code_Numeral.thy --- a/src/HOL/Code_Numeral.thy Thu Oct 29 16:23:57 2009 +0100 +++ b/src/HOL/Code_Numeral.thy Wed Nov 04 11:40:59 2009 +0100 @@ -172,7 +172,7 @@ "n < m \ nat_of n < nat_of m" instance proof -qed (auto simp add: code_numeral left_distrib div_mult_self1) +qed (auto simp add: code_numeral left_distrib intro: mult_commute) end @@ -268,7 +268,15 @@ lemma int_of_code [code]: "int_of k = (if k = 0 then 0 else (if k mod 2 = 0 then 2 * int_of (k div 2) else 2 * int_of (k div 2) + 1))" - by (auto simp add: int_of_def mod_div_equality') +proof - + have "(nat_of k div 2) * 2 + nat_of k mod 2 = nat_of k" + by (rule mod_div_equality) + then have "int ((nat_of k div 2) * 2 + nat_of k mod 2) = int (nat_of k)" + by simp + then have "int (nat_of k) = int (nat_of k div 2) * 2 + int (nat_of k mod 2)" + unfolding int_mult zadd_int [symmetric] by simp + then show ?thesis by (auto simp add: int_of_def mult_ac) +qed hide (open) const of_nat nat_of int_of diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/Decision_Procs/Commutative_Ring.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Decision_Procs/Commutative_Ring.thy Wed Nov 04 11:40:59 2009 +0100 @@ -0,0 +1,319 @@ +(* Author: Bernhard Haeupler + +Proving equalities in commutative rings done "right" in Isabelle/HOL. +*) + +header {* Proving equalities in commutative rings *} + +theory Commutative_Ring +imports Main Parity +uses ("commutative_ring_tac.ML") +begin + +text {* Syntax of multivariate polynomials (pol) and polynomial expressions. *} + +datatype 'a pol = + Pc 'a + | Pinj nat "'a pol" + | PX "'a pol" nat "'a pol" + +datatype 'a polex = + Pol "'a pol" + | Add "'a polex" "'a polex" + | Sub "'a polex" "'a polex" + | Mul "'a polex" "'a polex" + | Pow "'a polex" nat + | Neg "'a polex" + +text {* Interpretation functions for the shadow syntax. *} + +primrec + Ipol :: "'a::{comm_ring_1} list \ 'a pol \ 'a" +where + "Ipol l (Pc c) = c" + | "Ipol l (Pinj i P) = Ipol (drop i l) P" + | "Ipol l (PX P x Q) = Ipol l P * (hd l)^x + Ipol (drop 1 l) Q" + +primrec + Ipolex :: "'a::{comm_ring_1} list \ 'a polex \ 'a" +where + "Ipolex l (Pol P) = Ipol l P" + | "Ipolex l (Add P Q) = Ipolex l P + Ipolex l Q" + | "Ipolex l (Sub P Q) = Ipolex l P - Ipolex l Q" + | "Ipolex l (Mul P Q) = Ipolex l P * Ipolex l Q" + | "Ipolex l (Pow p n) = Ipolex l p ^ n" + | "Ipolex l (Neg P) = - Ipolex l P" + +text {* Create polynomial normalized polynomials given normalized inputs. *} + +definition + mkPinj :: "nat \ 'a pol \ 'a pol" where + "mkPinj x P = (case P of + Pc c \ Pc c | + Pinj y P \ Pinj (x + y) P | + PX p1 y p2 \ Pinj x P)" + +definition + mkPX :: "'a::{comm_ring} pol \ nat \ 'a pol \ 'a pol" where + "mkPX P i Q = (case P of + Pc c \ (if (c = 0) then (mkPinj 1 Q) else (PX P i Q)) | + Pinj j R \ PX P i Q | + PX P2 i2 Q2 \ (if (Q2 = (Pc 0)) then (PX P2 (i+i2) Q) else (PX P i Q)) )" + +text {* Defining the basic ring operations on normalized polynomials *} + +function + add :: "'a::{comm_ring} pol \ 'a pol \ 'a pol" (infixl "\" 65) +where + "Pc a \ Pc b = Pc (a + b)" + | "Pc c \ Pinj i P = Pinj i (P \ Pc c)" + | "Pinj i P \ Pc c = Pinj i (P \ Pc c)" + | "Pc c \ PX P i Q = PX P i (Q \ Pc c)" + | "PX P i Q \ Pc c = PX P i (Q \ Pc c)" + | "Pinj x P \ Pinj y Q = + (if x = y then mkPinj x (P \ Q) + else (if x > y then mkPinj y (Pinj (x - y) P \ Q) + else mkPinj x (Pinj (y - x) Q \ P)))" + | "Pinj x P \ PX Q y R = + (if x = 0 then P \ PX Q y R + else (if x = 1 then PX Q y (R \ P) + else PX Q y (R \ Pinj (x - 1) P)))" + | "PX P x R \ Pinj y Q = + (if y = 0 then PX P x R \ Q + else (if y = 1 then PX P x (R \ Q) + else PX P x (R \ Pinj (y - 1) Q)))" + | "PX P1 x P2 \ PX Q1 y Q2 = + (if x = y then mkPX (P1 \ Q1) x (P2 \ Q2) + else (if x > y then mkPX (PX P1 (x - y) (Pc 0) \ Q1) y (P2 \ Q2) + else mkPX (PX Q1 (y-x) (Pc 0) \ P1) x (P2 \ Q2)))" +by pat_completeness auto +termination by (relation "measure (\(x, y). size x + size y)") auto + +function + mul :: "'a::{comm_ring} pol \ 'a pol \ 'a pol" (infixl "\" 70) +where + "Pc a \ Pc b = Pc (a * b)" + | "Pc c \ Pinj i P = + (if c = 0 then Pc 0 else mkPinj i (P \ Pc c))" + | "Pinj i P \ Pc c = + (if c = 0 then Pc 0 else mkPinj i (P \ Pc c))" + | "Pc c \ PX P i Q = + (if c = 0 then Pc 0 else mkPX (P \ Pc c) i (Q \ Pc c))" + | "PX P i Q \ Pc c = + (if c = 0 then Pc 0 else mkPX (P \ Pc c) i (Q \ Pc c))" + | "Pinj x P \ Pinj y Q = + (if x = y then mkPinj x (P \ Q) else + (if x > y then mkPinj y (Pinj (x-y) P \ Q) + else mkPinj x (Pinj (y - x) Q \ P)))" + | "Pinj x P \ PX Q y R = + (if x = 0 then P \ PX Q y R else + (if x = 1 then mkPX (Pinj x P \ Q) y (R \ P) + else mkPX (Pinj x P \ Q) y (R \ Pinj (x - 1) P)))" + | "PX P x R \ Pinj y Q = + (if y = 0 then PX P x R \ Q else + (if y = 1 then mkPX (Pinj y Q \ P) x (R \ Q) + else mkPX (Pinj y Q \ P) x (R \ Pinj (y - 1) Q)))" + | "PX P1 x P2 \ PX Q1 y Q2 = + mkPX (P1 \ Q1) (x + y) (P2 \ Q2) \ + (mkPX (P1 \ mkPinj 1 Q2) x (Pc 0) \ + (mkPX (Q1 \ mkPinj 1 P2) y (Pc 0)))" +by pat_completeness auto +termination by (relation "measure (\(x, y). size x + size y)") + (auto simp add: mkPinj_def split: pol.split) + +text {* Negation*} +primrec + neg :: "'a::{comm_ring} pol \ 'a pol" +where + "neg (Pc c) = Pc (-c)" + | "neg (Pinj i P) = Pinj i (neg P)" + | "neg (PX P x Q) = PX (neg P) x (neg Q)" + +text {* Substraction *} +definition + sub :: "'a::{comm_ring} pol \ 'a pol \ 'a pol" (infixl "\" 65) +where + "sub P Q = P \ neg Q" + +text {* Square for Fast Exponentation *} +primrec + sqr :: "'a::{comm_ring_1} pol \ 'a pol" +where + "sqr (Pc c) = Pc (c * c)" + | "sqr (Pinj i P) = mkPinj i (sqr P)" + | "sqr (PX A x B) = mkPX (sqr A) (x + x) (sqr B) \ + mkPX (Pc (1 + 1) \ A \ mkPinj 1 B) x (Pc 0)" + +text {* Fast Exponentation *} +fun + pow :: "nat \ 'a::{comm_ring_1} pol \ 'a pol" +where + "pow 0 P = Pc 1" + | "pow n P = (if even n then pow (n div 2) (sqr P) + else P \ pow (n div 2) (sqr P))" + +lemma pow_if: + "pow n P = + (if n = 0 then Pc 1 else if even n then pow (n div 2) (sqr P) + else P \ pow (n div 2) (sqr P))" + by (cases n) simp_all + + +text {* Normalization of polynomial expressions *} + +primrec + norm :: "'a::{comm_ring_1} polex \ 'a pol" +where + "norm (Pol P) = P" + | "norm (Add P Q) = norm P \ norm Q" + | "norm (Sub P Q) = norm P \ norm Q" + | "norm (Mul P Q) = norm P \ norm Q" + | "norm (Pow P n) = pow n (norm P)" + | "norm (Neg P) = neg (norm P)" + +text {* mkPinj preserve semantics *} +lemma mkPinj_ci: "Ipol l (mkPinj a B) = Ipol l (Pinj a B)" + by (induct B) (auto simp add: mkPinj_def algebra_simps) + +text {* mkPX preserves semantics *} +lemma mkPX_ci: "Ipol l (mkPX A b C) = Ipol l (PX A b C)" + by (cases A) (auto simp add: mkPX_def mkPinj_ci power_add algebra_simps) + +text {* Correctness theorems for the implemented operations *} + +text {* Negation *} +lemma neg_ci: "Ipol l (neg P) = -(Ipol l P)" + by (induct P arbitrary: l) auto + +text {* Addition *} +lemma add_ci: "Ipol l (P \ Q) = Ipol l P + Ipol l Q" +proof (induct P Q arbitrary: l rule: add.induct) + case (6 x P y Q) + show ?case + proof (rule linorder_cases) + assume "x < y" + with 6 show ?case by (simp add: mkPinj_ci algebra_simps) + next + assume "x = y" + with 6 show ?case by (simp add: mkPinj_ci) + next + assume "x > y" + with 6 show ?case by (simp add: mkPinj_ci algebra_simps) + qed +next + case (7 x P Q y R) + have "x = 0 \ x = 1 \ x > 1" by arith + moreover + { assume "x = 0" with 7 have ?case by simp } + moreover + { assume "x = 1" with 7 have ?case by (simp add: algebra_simps) } + moreover + { assume "x > 1" from 7 have ?case by (cases x) simp_all } + ultimately show ?case by blast +next + case (8 P x R y Q) + have "y = 0 \ y = 1 \ y > 1" by arith + moreover + { assume "y = 0" with 8 have ?case by simp } + moreover + { assume "y = 1" with 8 have ?case by simp } + moreover + { assume "y > 1" with 8 have ?case by simp } + ultimately show ?case by blast +next + case (9 P1 x P2 Q1 y Q2) + show ?case + proof (rule linorder_cases) + assume a: "x < y" hence "EX d. d + x = y" by arith + with 9 a show ?case by (auto simp add: mkPX_ci power_add algebra_simps) + next + assume a: "y < x" hence "EX d. d + y = x" by arith + with 9 a show ?case by (auto simp add: power_add mkPX_ci algebra_simps) + next + assume "x = y" + with 9 show ?case by (simp add: mkPX_ci algebra_simps) + qed +qed (auto simp add: algebra_simps) + +text {* Multiplication *} +lemma mul_ci: "Ipol l (P \ Q) = Ipol l P * Ipol l Q" + by (induct P Q arbitrary: l rule: mul.induct) + (simp_all add: mkPX_ci mkPinj_ci algebra_simps add_ci power_add) + +text {* Substraction *} +lemma sub_ci: "Ipol l (P \ Q) = Ipol l P - Ipol l Q" + by (simp add: add_ci neg_ci sub_def) + +text {* Square *} +lemma sqr_ci: "Ipol ls (sqr P) = Ipol ls P * Ipol ls P" + by (induct P arbitrary: ls) + (simp_all add: add_ci mkPinj_ci mkPX_ci mul_ci algebra_simps power_add) + +text {* Power *} +lemma even_pow:"even n \ pow n P = pow (n div 2) (sqr P)" + by (induct n) simp_all + +lemma pow_ci: "Ipol ls (pow n P) = Ipol ls P ^ n" +proof (induct n arbitrary: P rule: nat_less_induct) + case (1 k) + show ?case + proof (cases k) + case 0 + then show ?thesis by simp + next + case (Suc l) + show ?thesis + proof cases + assume "even l" + then have "Suc l div 2 = l div 2" + by (simp add: nat_number even_nat_plus_one_div_two) + moreover + from Suc have "l < k" by simp + with 1 have "\P. Ipol ls (pow l P) = Ipol ls P ^ l" by simp + moreover + note Suc `even l` even_nat_plus_one_div_two + ultimately show ?thesis by (auto simp add: mul_ci power_Suc even_pow) + next + assume "odd l" + { + fix p + have "Ipol ls (sqr P) ^ (Suc l div 2) = Ipol ls P ^ Suc l" + proof (cases l) + case 0 + with `odd l` show ?thesis by simp + next + case (Suc w) + with `odd l` have "even w" by simp + have two_times: "2 * (w div 2) = w" + by (simp only: numerals even_nat_div_two_times_two [OF `even w`]) + have "Ipol ls P * Ipol ls P = Ipol ls P ^ Suc (Suc 0)" + by (simp add: power_Suc) + then have "Ipol ls P * Ipol ls P = Ipol ls P ^ 2" + by (simp add: numerals) + with Suc show ?thesis + by (auto simp add: power_mult [symmetric, of _ 2 _] two_times mul_ci sqr_ci + simp del: power_Suc) + qed + } with 1 Suc `odd l` show ?thesis by simp + qed + qed +qed + +text {* Normalization preserves semantics *} +lemma norm_ci: "Ipolex l Pe = Ipol l (norm Pe)" + by (induct Pe) (simp_all add: add_ci sub_ci mul_ci neg_ci pow_ci) + +text {* Reflection lemma: Key to the (incomplete) decision procedure *} +lemma norm_eq: + assumes "norm P1 = norm P2" + shows "Ipolex l P1 = Ipolex l P2" +proof - + from prems have "Ipol l (norm P1) = Ipol l (norm P2)" by simp + then show ?thesis by (simp only: norm_ci) +qed + + +use "commutative_ring_tac.ML" +setup Commutative_Ring_Tac.setup + +end diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/Decision_Procs/Commutative_Ring_Complete.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Decision_Procs/Commutative_Ring_Complete.thy Wed Nov 04 11:40:59 2009 +0100 @@ -0,0 +1,391 @@ +(* Author: Bernhard Haeupler + +This theory is about of the relative completeness of method comm-ring +method. As long as the reified atomic polynomials of type 'a pol are +in normal form, the cring method is complete. +*) + +header {* Proof of the relative completeness of method comm-ring *} + +theory Commutative_Ring_Complete +imports Commutative_Ring +begin + +text {* Formalization of normal form *} +fun + isnorm :: "('a::{comm_ring}) pol \ bool" +where + "isnorm (Pc c) \ True" + | "isnorm (Pinj i (Pc c)) \ False" + | "isnorm (Pinj i (Pinj j Q)) \ False" + | "isnorm (Pinj 0 P) \ False" + | "isnorm (Pinj i (PX Q1 j Q2)) \ isnorm (PX Q1 j Q2)" + | "isnorm (PX P 0 Q) \ False" + | "isnorm (PX (Pc c) i Q) \ c \ 0 \ isnorm Q" + | "isnorm (PX (PX P1 j (Pc c)) i Q) \ c \ 0 \ isnorm (PX P1 j (Pc c)) \ isnorm Q" + | "isnorm (PX P i Q) \ isnorm P \ isnorm Q" + +(* Some helpful lemmas *) +lemma norm_Pinj_0_False:"isnorm (Pinj 0 P) = False" +by(cases P, auto) + +lemma norm_PX_0_False:"isnorm (PX (Pc 0) i Q) = False" +by(cases i, auto) + +lemma norm_Pinj:"isnorm (Pinj i Q) \ isnorm Q" +by(cases i,simp add: norm_Pinj_0_False norm_PX_0_False,cases Q) auto + +lemma norm_PX2:"isnorm (PX P i Q) \ isnorm Q" +by(cases i, auto, cases P, auto, case_tac pol2, auto) + +lemma norm_PX1:"isnorm (PX P i Q) \ isnorm P" +by(cases i, auto, cases P, auto, case_tac pol2, auto) + +lemma mkPinj_cn:"\y~=0; isnorm Q\ \ isnorm (mkPinj y Q)" +apply(auto simp add: mkPinj_def norm_Pinj_0_False split: pol.split) +apply(case_tac nat, auto simp add: norm_Pinj_0_False) +by(case_tac pol, auto) (case_tac y, auto) + +lemma norm_PXtrans: + assumes A:"isnorm (PX P x Q)" and "isnorm Q2" + shows "isnorm (PX P x Q2)" +proof(cases P) + case (PX p1 y p2) from prems show ?thesis by(cases x, auto, cases p2, auto) +next + case Pc from prems show ?thesis by(cases x, auto) +next + case Pinj from prems show ?thesis by(cases x, auto) +qed + +lemma norm_PXtrans2: assumes A:"isnorm (PX P x Q)" and "isnorm Q2" shows "isnorm (PX P (Suc (n+x)) Q2)" +proof(cases P) + case (PX p1 y p2) + from prems show ?thesis by(cases x, auto, cases p2, auto) +next + case Pc + from prems show ?thesis by(cases x, auto) +next + case Pinj + from prems show ?thesis by(cases x, auto) +qed + +text {* mkPX conserves normalizedness (@{text "_cn"}) *} +lemma mkPX_cn: + assumes "x \ 0" and "isnorm P" and "isnorm Q" + shows "isnorm (mkPX P x Q)" +proof(cases P) + case (Pc c) + from prems show ?thesis by (cases x) (auto simp add: mkPinj_cn mkPX_def) +next + case (Pinj i Q) + from prems show ?thesis by (cases x) (auto simp add: mkPinj_cn mkPX_def) +next + case (PX P1 y P2) + from prems have Y0:"y>0" by(cases y, auto) + from prems have "isnorm P1" "isnorm P2" by (auto simp add: norm_PX1[of P1 y P2] norm_PX2[of P1 y P2]) + with prems Y0 show ?thesis by (cases x, auto simp add: mkPX_def norm_PXtrans2[of P1 y _ Q _], cases P2, auto) +qed + +text {* add conserves normalizedness *} +lemma add_cn:"isnorm P \ isnorm Q \ isnorm (P \ Q)" +proof(induct P Q rule: add.induct) + case (2 c i P2) thus ?case by (cases P2, simp_all, cases i, simp_all) +next + case (3 i P2 c) thus ?case by (cases P2, simp_all, cases i, simp_all) +next + case (4 c P2 i Q2) + from prems have "isnorm P2" "isnorm Q2" by (auto simp only: norm_PX1[of P2 i Q2] norm_PX2[of P2 i Q2]) + with prems show ?case by(cases i, simp, cases P2, auto, case_tac pol2, auto) +next + case (5 P2 i Q2 c) + from prems have "isnorm P2" "isnorm Q2" by (auto simp only: norm_PX1[of P2 i Q2] norm_PX2[of P2 i Q2]) + with prems show ?case by(cases i, simp, cases P2, auto, case_tac pol2, auto) +next + case (6 x P2 y Q2) + from prems have Y0:"y>0" by (cases y, auto simp add: norm_Pinj_0_False) + from prems have X0:"x>0" by (cases x, auto simp add: norm_Pinj_0_False) + have "x < y \ x = y \ x > y" by arith + moreover + { assume "xy" hence "EX d. x=d+y" by arith + then obtain d where "x=d+y".. + moreover + note prems Y0 + moreover + from prems have "isnorm P2" "isnorm Q2" by (auto simp add: norm_Pinj[of _ P2] norm_Pinj[of _ Q2]) + moreover + with prems have "isnorm (Pinj d P2)" by (cases d, simp, cases P2, auto) + ultimately have ?case by (simp add: mkPinj_cn)} + ultimately show ?case by blast +next + case (7 x P2 Q2 y R) + have "x=0 \ (x = 1) \ (x > 1)" by arith + moreover + { assume "x=0" with prems have ?case by (auto simp add: norm_Pinj_0_False)} + moreover + { assume "x=1" + from prems have "isnorm R" "isnorm P2" by (auto simp add: norm_Pinj[of _ P2] norm_PX2[of Q2 y R]) + with prems have "isnorm (R \ P2)" by simp + with prems have ?case by (simp add: norm_PXtrans[of Q2 y _]) } + moreover + { assume "x > 1" hence "EX d. x=Suc (Suc d)" by arith + then obtain d where X:"x=Suc (Suc d)" .. + from prems have NR:"isnorm R" "isnorm P2" by (auto simp add: norm_Pinj[of _ P2] norm_PX2[of Q2 y R]) + with prems have "isnorm (Pinj (x - 1) P2)" by(cases P2, auto) + with prems NR have "isnorm (R \ Pinj (x - 1) P2)" "isnorm (PX Q2 y R)" by simp fact + with X have ?case by (simp add: norm_PXtrans[of Q2 y _]) } + ultimately show ?case by blast +next + case (8 Q2 y R x P2) + have "x = 0 \ x = 1 \ x > 1" by arith + moreover + { assume "x=0" with prems have ?case by (auto simp add: norm_Pinj_0_False)} + moreover + { assume "x=1" + from prems have "isnorm R" "isnorm P2" by (auto simp add: norm_Pinj[of _ P2] norm_PX2[of Q2 y R]) + with prems have "isnorm (R \ P2)" by simp + with prems have ?case by (simp add: norm_PXtrans[of Q2 y _]) } + moreover + { assume "x > 1" hence "EX d. x=Suc (Suc d)" by arith + then obtain d where X:"x=Suc (Suc d)" .. + from prems have NR:"isnorm R" "isnorm P2" by (auto simp add: norm_Pinj[of _ P2] norm_PX2[of Q2 y R]) + with prems have "isnorm (Pinj (x - 1) P2)" by(cases P2, auto) + with prems NR have "isnorm (R \ Pinj (x - 1) P2)" "isnorm (PX Q2 y R)" by simp fact + with X have ?case by (simp add: norm_PXtrans[of Q2 y _]) } + ultimately show ?case by blast +next + case (9 P1 x P2 Q1 y Q2) + from prems have Y0:"y>0" by(cases y, auto) + from prems have X0:"x>0" by(cases x, auto) + from prems have NP1:"isnorm P1" and NP2:"isnorm P2" by (auto simp add: norm_PX1[of P1 _ P2] norm_PX2[of P1 _ P2]) + from prems have NQ1:"isnorm Q1" and NQ2:"isnorm Q2" by (auto simp add: norm_PX1[of Q1 _ Q2] norm_PX2[of Q1 _ Q2]) + have "y < x \ x = y \ x < y" by arith + moreover + {assume sm1:"y < x" hence "EX d. x=d+y" by arith + then obtain d where sm2:"x=d+y".. + note prems NQ1 NP1 NP2 NQ2 sm1 sm2 + moreover + have "isnorm (PX P1 d (Pc 0))" + proof(cases P1) + case (PX p1 y p2) + with prems show ?thesis by(cases d, simp,cases p2, auto) + next case Pc from prems show ?thesis by(cases d, auto) + next case Pinj from prems show ?thesis by(cases d, auto) + qed + ultimately have "isnorm (P2 \ Q2)" "isnorm (PX P1 (x - y) (Pc 0) \ Q1)" by auto + with Y0 sm1 sm2 have ?case by (simp add: mkPX_cn)} + moreover + {assume "x=y" + from prems NP1 NP2 NQ1 NQ2 have "isnorm (P2 \ Q2)" "isnorm (P1 \ Q1)" by auto + with Y0 prems have ?case by (simp add: mkPX_cn) } + moreover + {assume sm1:"x Q2)" "isnorm (PX Q1 (y - x) (Pc 0) \ P1)" by auto + with X0 sm1 sm2 have ?case by (simp add: mkPX_cn)} + ultimately show ?case by blast +qed simp + +text {* mul concerves normalizedness *} +lemma mul_cn :"isnorm P \ isnorm Q \ isnorm (P \ Q)" +proof(induct P Q rule: mul.induct) + case (2 c i P2) thus ?case + by (cases P2, simp_all) (cases "i",simp_all add: mkPinj_cn) +next + case (3 i P2 c) thus ?case + by (cases P2, simp_all) (cases "i",simp_all add: mkPinj_cn) +next + case (4 c P2 i Q2) + from prems have "isnorm P2" "isnorm Q2" by (auto simp only: norm_PX1[of P2 i Q2] norm_PX2[of P2 i Q2]) + with prems show ?case + by - (case_tac "c=0",simp_all,case_tac "i=0",simp_all add: mkPX_cn) +next + case (5 P2 i Q2 c) + from prems have "isnorm P2" "isnorm Q2" by (auto simp only: norm_PX1[of P2 i Q2] norm_PX2[of P2 i Q2]) + with prems show ?case + by - (case_tac "c=0",simp_all,case_tac "i=0",simp_all add: mkPX_cn) +next + case (6 x P2 y Q2) + have "x < y \ x = y \ x > y" by arith + moreover + { assume "x0" by (cases x, auto simp add: norm_Pinj_0_False) + moreover + from prems have "isnorm P2" "isnorm Q2" by (auto simp add: norm_Pinj[of _ P2] norm_Pinj[of _ Q2]) + moreover + with prems have "isnorm (Pinj d Q2)" by (cases d, simp, cases Q2, auto) + ultimately have ?case by (simp add: mkPinj_cn)} + moreover + { assume "x=y" + moreover + from prems have "isnorm P2" "isnorm Q2" by(auto simp add: norm_Pinj[of _ P2] norm_Pinj[of _ Q2]) + moreover + with prems have "y>0" by (cases y, auto simp add: norm_Pinj_0_False) + moreover + note prems + moreover + ultimately have ?case by (simp add: mkPinj_cn) } + moreover + { assume "x>y" hence "EX d. x=d+y" by arith + then obtain d where "x=d+y".. + moreover + note prems + moreover + from prems have "y>0" by (cases y, auto simp add: norm_Pinj_0_False) + moreover + from prems have "isnorm P2" "isnorm Q2" by (auto simp add: norm_Pinj[of _ P2] norm_Pinj[of _ Q2]) + moreover + with prems have "isnorm (Pinj d P2)" by (cases d, simp, cases P2, auto) + ultimately have ?case by (simp add: mkPinj_cn) } + ultimately show ?case by blast +next + case (7 x P2 Q2 y R) + from prems have Y0:"y>0" by(cases y, auto) + have "x=0 \ (x = 1) \ (x > 1)" by arith + moreover + { assume "x=0" with prems have ?case by (auto simp add: norm_Pinj_0_False)} + moreover + { assume "x=1" + from prems have "isnorm R" "isnorm P2" by (auto simp add: norm_Pinj[of _ P2] norm_PX2[of Q2 y R]) + with prems have "isnorm (R \ P2)" "isnorm Q2" by (auto simp add: norm_PX1[of Q2 y R]) + with Y0 prems have ?case by (simp add: mkPX_cn)} + moreover + { assume "x > 1" hence "EX d. x=Suc (Suc d)" by arith + then obtain d where X:"x=Suc (Suc d)" .. + from prems have NR:"isnorm R" "isnorm Q2" by (auto simp add: norm_PX2[of Q2 y R] norm_PX1[of Q2 y R]) + moreover + from prems have "isnorm (Pinj (x - 1) P2)" by(cases P2, auto) + moreover + from prems have "isnorm (Pinj x P2)" by(cases P2, auto) + moreover + note prems + ultimately have "isnorm (R \ Pinj (x - 1) P2)" "isnorm (Pinj x P2 \ Q2)" by auto + with Y0 X have ?case by (simp add: mkPX_cn)} + ultimately show ?case by blast +next + case (8 Q2 y R x P2) + from prems have Y0:"y>0" by(cases y, auto) + have "x=0 \ (x = 1) \ (x > 1)" by arith + moreover + { assume "x=0" with prems have ?case by (auto simp add: norm_Pinj_0_False)} + moreover + { assume "x=1" + from prems have "isnorm R" "isnorm P2" by (auto simp add: norm_Pinj[of _ P2] norm_PX2[of Q2 y R]) + with prems have "isnorm (R \ P2)" "isnorm Q2" by (auto simp add: norm_PX1[of Q2 y R]) + with Y0 prems have ?case by (simp add: mkPX_cn) } + moreover + { assume "x > 1" hence "EX d. x=Suc (Suc d)" by arith + then obtain d where X:"x=Suc (Suc d)" .. + from prems have NR:"isnorm R" "isnorm Q2" by (auto simp add: norm_PX2[of Q2 y R] norm_PX1[of Q2 y R]) + moreover + from prems have "isnorm (Pinj (x - 1) P2)" by(cases P2, auto) + moreover + from prems have "isnorm (Pinj x P2)" by(cases P2, auto) + moreover + note prems + ultimately have "isnorm (R \ Pinj (x - 1) P2)" "isnorm (Pinj x P2 \ Q2)" by auto + with Y0 X have ?case by (simp add: mkPX_cn) } + ultimately show ?case by blast +next + case (9 P1 x P2 Q1 y Q2) + from prems have X0:"x>0" by(cases x, auto) + from prems have Y0:"y>0" by(cases y, auto) + note prems + moreover + from prems have "isnorm P1" "isnorm P2" by (auto simp add: norm_PX1[of P1 x P2] norm_PX2[of P1 x P2]) + moreover + from prems have "isnorm Q1" "isnorm Q2" by (auto simp add: norm_PX1[of Q1 y Q2] norm_PX2[of Q1 y Q2]) + ultimately have "isnorm (P1 \ Q1)" "isnorm (P2 \ Q2)" + "isnorm (P1 \ mkPinj 1 Q2)" "isnorm (Q1 \ mkPinj 1 P2)" + by (auto simp add: mkPinj_cn) + with prems X0 Y0 have + "isnorm (mkPX (P1 \ Q1) (x + y) (P2 \ Q2))" + "isnorm (mkPX (P1 \ mkPinj (Suc 0) Q2) x (Pc 0))" + "isnorm (mkPX (Q1 \ mkPinj (Suc 0) P2) y (Pc 0))" + by (auto simp add: mkPX_cn) + thus ?case by (simp add: add_cn) +qed(simp) + +text {* neg conserves normalizedness *} +lemma neg_cn: "isnorm P \ isnorm (neg P)" +proof (induct P) + case (Pinj i P2) + from prems have "isnorm P2" by (simp add: norm_Pinj[of i P2]) + with prems show ?case by(cases P2, auto, cases i, auto) +next + case (PX P1 x P2) + from prems have "isnorm P2" "isnorm P1" by (auto simp add: norm_PX1[of P1 x P2] norm_PX2[of P1 x P2]) + with prems show ?case + proof(cases P1) + case (PX p1 y p2) + with prems show ?thesis by(cases x, auto, cases p2, auto) + next + case Pinj + with prems show ?thesis by(cases x, auto) + qed(cases x, auto) +qed(simp) + +text {* sub conserves normalizedness *} +lemma sub_cn:"isnorm p \ isnorm q \ isnorm (p \ q)" +by (simp add: sub_def add_cn neg_cn) + +text {* sqr conserves normalizizedness *} +lemma sqr_cn:"isnorm P \ isnorm (sqr P)" +proof(induct P) + case (Pinj i Q) + from prems show ?case by(cases Q, auto simp add: mkPX_cn mkPinj_cn, cases i, auto simp add: mkPX_cn mkPinj_cn) +next + case (PX P1 x P2) + from prems have "x+x~=0" "isnorm P2" "isnorm P1" by (cases x, auto simp add: norm_PX1[of P1 x P2] norm_PX2[of P1 x P2]) + with prems have + "isnorm (mkPX (Pc (1 + 1) \ P1 \ mkPinj (Suc 0) P2) x (Pc 0))" + and "isnorm (mkPX (sqr P1) (x + x) (sqr P2))" + by (auto simp add: add_cn mkPX_cn mkPinj_cn mul_cn) + thus ?case by (auto simp add: add_cn mkPX_cn mkPinj_cn mul_cn) +qed simp + +text {* pow conserves normalizedness *} +lemma pow_cn:"isnorm P \ isnorm (pow n P)" +proof (induct n arbitrary: P rule: nat_less_induct) + case (1 k) + show ?case + proof (cases "k=0") + case False + then have K2:"k div 2 < k" by (cases k, auto) + from prems have "isnorm (sqr P)" by (simp add: sqr_cn) + with prems K2 show ?thesis + by (simp add: allE[of _ "(k div 2)" _] allE[of _ "(sqr P)" _], cases k, auto simp add: mul_cn) + qed simp +qed + +end diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/Decision_Procs/Decision_Procs.thy --- a/src/HOL/Decision_Procs/Decision_Procs.thy Thu Oct 29 16:23:57 2009 +0100 +++ b/src/HOL/Decision_Procs/Decision_Procs.thy Wed Nov 04 11:40:59 2009 +0100 @@ -1,7 +1,10 @@ -header {* Various decision procedures. typically involving reflection *} +header {* Various decision procedures, typically involving reflection *} theory Decision_Procs -imports Cooper Ferrack MIR Approximation Dense_Linear_Order "ex/Approximation_Ex" "ex/Dense_Linear_Order_Ex" Parametric_Ferrante_Rackoff +imports + Commutative_Ring Cooper Ferrack MIR Approximation Dense_Linear_Order Parametric_Ferrante_Rackoff + Commutative_Ring_Complete + "ex/Commutative_Ring_Ex" "ex/Approximation_Ex" "ex/Dense_Linear_Order_Ex" begin end \ No newline at end of file diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/Decision_Procs/commutative_ring_tac.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Decision_Procs/commutative_ring_tac.ML Wed Nov 04 11:40:59 2009 +0100 @@ -0,0 +1,104 @@ +(* Author: Amine Chaieb + +Tactic for solving equalities over commutative rings. +*) + +signature COMMUTATIVE_RING_TAC = +sig + val tac: Proof.context -> int -> tactic + val setup: theory -> theory +end + +structure Commutative_Ring_Tac: COMMUTATIVE_RING_TAC = +struct + +(* Zero and One of the commutative ring *) +fun cring_zero T = Const (@{const_name HOL.zero}, T); +fun cring_one T = Const (@{const_name HOL.one}, T); + +(* reification functions *) +(* add two polynom expressions *) +fun polT t = Type (@{type_name Commutative_Ring.pol}, [t]); +fun polexT t = Type (@{type_name Commutative_Ring.polex}, [t]); + +(* pol *) +fun pol_Pc t = Const (@{const_name Commutative_Ring.pol.Pc}, t --> polT t); +fun pol_Pinj t = Const (@{const_name Commutative_Ring.pol.Pinj}, HOLogic.natT --> polT t --> polT t); +fun pol_PX t = Const (@{const_name Commutative_Ring.pol.PX}, polT t --> HOLogic.natT --> polT t --> polT t); + +(* polex *) +fun polex_add t = Const (@{const_name Commutative_Ring.polex.Add}, polexT t --> polexT t --> polexT t); +fun polex_sub t = Const (@{const_name Commutative_Ring.polex.Sub}, polexT t --> polexT t --> polexT t); +fun polex_mul t = Const (@{const_name Commutative_Ring.polex.Mul}, polexT t --> polexT t --> polexT t); +fun polex_neg t = Const (@{const_name Commutative_Ring.polex.Neg}, polexT t --> polexT t); +fun polex_pol t = Const (@{const_name Commutative_Ring.polex.Pol}, polT t --> polexT t); +fun polex_pow t = Const (@{const_name Commutative_Ring.polex.Pow}, polexT t --> HOLogic.natT --> polexT t); + +(* reification of polynoms : primitive cring expressions *) +fun reif_pol T vs (t as Free _) = + let + val one = @{term "1::nat"}; + val i = find_index (fn t' => t' = t) vs + in if i = 0 + then pol_PX T $ (pol_Pc T $ cring_one T) + $ one $ (pol_Pc T $ cring_zero T) + else pol_Pinj T $ HOLogic.mk_nat i + $ (pol_PX T $ (pol_Pc T $ cring_one T) + $ one $ (pol_Pc T $ cring_zero T)) + end + | reif_pol T vs t = pol_Pc T $ t; + +(* reification of polynom expressions *) +fun reif_polex T vs (Const (@{const_name HOL.plus}, _) $ a $ b) = + polex_add T $ reif_polex T vs a $ reif_polex T vs b + | reif_polex T vs (Const (@{const_name HOL.minus}, _) $ a $ b) = + polex_sub T $ reif_polex T vs a $ reif_polex T vs b + | reif_polex T vs (Const (@{const_name HOL.times}, _) $ a $ b) = + polex_mul T $ reif_polex T vs a $ reif_polex T vs b + | reif_polex T vs (Const (@{const_name HOL.uminus}, _) $ a) = + polex_neg T $ reif_polex T vs a + | reif_polex T vs (Const (@{const_name Power.power}, _) $ a $ n) = + polex_pow T $ reif_polex T vs a $ n + | reif_polex T vs t = polex_pol T $ reif_pol T vs t; + +(* reification of the equation *) +val cr_sort = @{sort "comm_ring_1"}; + +fun reif_eq thy (eq as Const(@{const_name "op ="}, Type("fun", [T, _])) $ lhs $ rhs) = + if Sign.of_sort thy (T, cr_sort) then + let + val fs = OldTerm.term_frees eq; + val cvs = cterm_of thy (HOLogic.mk_list T fs); + val clhs = cterm_of thy (reif_polex T fs lhs); + val crhs = cterm_of thy (reif_polex T fs rhs); + val ca = ctyp_of thy T; + in (ca, cvs, clhs, crhs) end + else error ("reif_eq: not an equation over " ^ Syntax.string_of_sort_global thy cr_sort) + | reif_eq _ _ = error "reif_eq: not an equation"; + +(* The cring tactic *) +(* Attention: You have to make sure that no t^0 is in the goal!! *) +(* Use simply rewriting t^0 = 1 *) +val cring_simps = + [@{thm mkPX_def}, @{thm mkPinj_def}, @{thm sub_def}, @{thm power_add}, + @{thm even_def}, @{thm pow_if}, sym OF [@{thm power_add}]]; + +fun tac ctxt = SUBGOAL (fn (g, i) => + let + val thy = ProofContext.theory_of ctxt; + val cring_ss = Simplifier.simpset_of ctxt (*FIXME really the full simpset!?*) + addsimps cring_simps; + val (ca, cvs, clhs, crhs) = reif_eq thy (HOLogic.dest_Trueprop g) + val norm_eq_th = + simplify cring_ss (instantiate' [SOME ca] [SOME clhs, SOME crhs, SOME cvs] @{thm norm_eq}) + in + cut_rules_tac [norm_eq_th] i + THEN (simp_tac cring_ss i) + THEN (simp_tac cring_ss i) + end); + +val setup = + Method.setup @{binding comm_ring} (Scan.succeed (SIMPLE_METHOD' o tac)) + "reflective decision procedure for equalities over commutative rings"; + +end; diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/Decision_Procs/ex/Commutative_Ring_Ex.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Decision_Procs/ex/Commutative_Ring_Ex.thy Wed Nov 04 11:40:59 2009 +0100 @@ -0,0 +1,48 @@ +(* Author: Bernhard Haeupler *) + +header {* Some examples demonstrating the comm-ring method *} + +theory Commutative_Ring_Ex +imports Commutative_Ring +begin + +lemma "4*(x::int)^5*y^3*x^2*3 + x*z + 3^5 = 12*x^7*y^3 + z*x + 243" +by comm_ring + +lemma "((x::int) + y)^2 = x^2 + y^2 + 2*x*y" +by comm_ring + +lemma "((x::int) + y)^3 = x^3 + y^3 + 3*x^2*y + 3*y^2*x" +by comm_ring + +lemma "((x::int) - y)^3 = x^3 + 3*x*y^2 + (-3)*y*x^2 - y^3" +by comm_ring + +lemma "((x::int) - y)^2 = x^2 + y^2 - 2*x*y" +by comm_ring + +lemma " ((a::int) + b + c)^2 = a^2 + b^2 + c^2 + 2*a*b + 2*b*c + 2*a*c" +by comm_ring + +lemma "((a::int) - b - c)^2 = a^2 + b^2 + c^2 - 2*a*b + 2*b*c - 2*a*c" +by comm_ring + +lemma "(a::int)*b + a*c = a*(b+c)" +by comm_ring + +lemma "(a::int)^2 - b^2 = (a - b) * (a + b)" +by comm_ring + +lemma "(a::int)^3 - b^3 = (a - b) * (a^2 + a*b + b^2)" +by comm_ring + +lemma "(a::int)^3 + b^3 = (a + b) * (a^2 - a*b + b^2)" +by comm_ring + +lemma "(a::int)^4 - b^4 = (a - b) * (a + b)*(a^2 + b^2)" +by comm_ring + +lemma "(a::int)^10 - b^10 = (a - b) * (a^9 + a^8*b + a^7*b^2 + a^6*b^3 + a^5*b^4 + a^4*b^5 + a^3*b^6 + a^2*b^7 + a*b^8 + b^9 )" +by comm_ring + +end diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/Divides.thy --- a/src/HOL/Divides.thy Thu Oct 29 16:23:57 2009 +0100 +++ b/src/HOL/Divides.thy Wed Nov 04 11:40:59 2009 +0100 @@ -7,15 +7,7 @@ theory Divides imports Nat_Numeral Nat_Transfer -uses - "~~/src/Provers/Arith/assoc_fold.ML" - "~~/src/Provers/Arith/cancel_numerals.ML" - "~~/src/Provers/Arith/combine_numerals.ML" - "~~/src/Provers/Arith/cancel_numeral_factor.ML" - "~~/src/Provers/Arith/extract_common_term.ML" - ("Tools/numeral_simprocs.ML") - ("Tools/nat_numeral_simprocs.ML") - "~~/src/Provers/Arith/cancel_div_mod.ML" +uses "~~/src/Provers/Arith/cancel_div_mod.ML" begin subsection {* Syntactic division operations *} @@ -139,7 +131,7 @@ note that ultimately show thesis by blast qed -lemma dvd_eq_mod_eq_0 [code_unfold]: "a dvd b \ b mod a = 0" +lemma dvd_eq_mod_eq_0 [code, code_unfold, code_inline del]: "a dvd b \ b mod a = 0" proof assume "b mod a = 0" with mod_div_equality [of b a] have "b div a * a = b" by simp @@ -435,18 +427,18 @@ @{term "q\nat"}(uotient) and @{term "r\nat"}(emainder). *} -definition divmod_rel :: "nat \ nat \ nat \ nat \ bool" where - "divmod_rel m n qr \ +definition divmod_nat_rel :: "nat \ nat \ nat \ nat \ bool" where + "divmod_nat_rel m n qr \ m = fst qr * n + snd qr \ (if n = 0 then fst qr = 0 else if n > 0 then 0 \ snd qr \ snd qr < n else n < snd qr \ snd qr \ 0)" -text {* @{const divmod_rel} is total: *} - -lemma divmod_rel_ex: - obtains q r where "divmod_rel m n (q, r)" +text {* @{const divmod_nat_rel} is total: *} + +lemma divmod_nat_rel_ex: + obtains q r where "divmod_nat_rel m n (q, r)" proof (cases "n = 0") case True with that show thesis - by (auto simp add: divmod_rel_def) + by (auto simp add: divmod_nat_rel_def) next case False have "\q r. m = q * n + r \ r < n" @@ -470,19 +462,19 @@ qed qed with that show thesis - using `n \ 0` by (auto simp add: divmod_rel_def) + using `n \ 0` by (auto simp add: divmod_nat_rel_def) qed -text {* @{const divmod_rel} is injective: *} - -lemma divmod_rel_unique: - assumes "divmod_rel m n qr" - and "divmod_rel m n qr'" +text {* @{const divmod_nat_rel} is injective: *} + +lemma divmod_nat_rel_unique: + assumes "divmod_nat_rel m n qr" + and "divmod_nat_rel m n qr'" shows "qr = qr'" proof (cases "n = 0") case True with assms show ?thesis by (cases qr, cases qr') - (simp add: divmod_rel_def) + (simp add: divmod_nat_rel_def) next case False have aux: "\q r q' r'. q' * n + r' = q * n + r \ r < n \ q' \ (q\nat)" @@ -491,91 +483,91 @@ apply (auto simp add: add_mult_distrib) done from `n \ 0` assms have "fst qr = fst qr'" - by (auto simp add: divmod_rel_def intro: order_antisym dest: aux sym) + by (auto simp add: divmod_nat_rel_def intro: order_antisym dest: aux sym) moreover from this assms have "snd qr = snd qr'" - by (simp add: divmod_rel_def) + by (simp add: divmod_nat_rel_def) ultimately show ?thesis by (cases qr, cases qr') simp qed text {* We instantiate divisibility on the natural numbers by - means of @{const divmod_rel}: + means of @{const divmod_nat_rel}: *} instantiation nat :: semiring_div begin -definition divmod :: "nat \ nat \ nat \ nat" where - [code del]: "divmod m n = (THE qr. divmod_rel m n qr)" - -lemma divmod_rel_divmod: - "divmod_rel m n (divmod m n)" +definition divmod_nat :: "nat \ nat \ nat \ nat" where + [code del]: "divmod_nat m n = (THE qr. divmod_nat_rel m n qr)" + +lemma divmod_nat_rel_divmod_nat: + "divmod_nat_rel m n (divmod_nat m n)" proof - - from divmod_rel_ex - obtain qr where rel: "divmod_rel m n qr" . + from divmod_nat_rel_ex + obtain qr where rel: "divmod_nat_rel m n qr" . then show ?thesis - by (auto simp add: divmod_def intro: theI elim: divmod_rel_unique) + by (auto simp add: divmod_nat_def intro: theI elim: divmod_nat_rel_unique) qed -lemma divmod_eq: - assumes "divmod_rel m n qr" - shows "divmod m n = qr" - using assms by (auto intro: divmod_rel_unique divmod_rel_divmod) +lemma divmod_nat_eq: + assumes "divmod_nat_rel m n qr" + shows "divmod_nat m n = qr" + using assms by (auto intro: divmod_nat_rel_unique divmod_nat_rel_divmod_nat) definition div_nat where - "m div n = fst (divmod m n)" + "m div n = fst (divmod_nat m n)" definition mod_nat where - "m mod n = snd (divmod m n)" - -lemma divmod_div_mod: - "divmod m n = (m div n, m mod n)" + "m mod n = snd (divmod_nat m n)" + +lemma divmod_nat_div_mod: + "divmod_nat m n = (m div n, m mod n)" unfolding div_nat_def mod_nat_def by simp lemma div_eq: - assumes "divmod_rel m n (q, r)" + assumes "divmod_nat_rel m n (q, r)" shows "m div n = q" - using assms by (auto dest: divmod_eq simp add: divmod_div_mod) + using assms by (auto dest: divmod_nat_eq simp add: divmod_nat_div_mod) lemma mod_eq: - assumes "divmod_rel m n (q, r)" + assumes "divmod_nat_rel m n (q, r)" shows "m mod n = r" - using assms by (auto dest: divmod_eq simp add: divmod_div_mod) - -lemma divmod_rel: "divmod_rel m n (m div n, m mod n)" - by (simp add: div_nat_def mod_nat_def divmod_rel_divmod) - -lemma divmod_zero: - "divmod m 0 = (0, m)" + using assms by (auto dest: divmod_nat_eq simp add: divmod_nat_div_mod) + +lemma divmod_nat_rel: "divmod_nat_rel m n (m div n, m mod n)" + by (simp add: div_nat_def mod_nat_def divmod_nat_rel_divmod_nat) + +lemma divmod_nat_zero: + "divmod_nat m 0 = (0, m)" proof - - from divmod_rel [of m 0] show ?thesis - unfolding divmod_div_mod divmod_rel_def by simp + from divmod_nat_rel [of m 0] show ?thesis + unfolding divmod_nat_div_mod divmod_nat_rel_def by simp qed -lemma divmod_base: +lemma divmod_nat_base: assumes "m < n" - shows "divmod m n = (0, m)" + shows "divmod_nat m n = (0, m)" proof - - from divmod_rel [of m n] show ?thesis - unfolding divmod_div_mod divmod_rel_def + from divmod_nat_rel [of m n] show ?thesis + unfolding divmod_nat_div_mod divmod_nat_rel_def using assms by (cases "m div n = 0") (auto simp add: gr0_conv_Suc [of "m div n"]) qed -lemma divmod_step: +lemma divmod_nat_step: assumes "0 < n" and "n \ m" - shows "divmod m n = (Suc ((m - n) div n), (m - n) mod n)" + shows "divmod_nat m n = (Suc ((m - n) div n), (m - n) mod n)" proof - - from divmod_rel have divmod_m_n: "divmod_rel m n (m div n, m mod n)" . + from divmod_nat_rel have divmod_nat_m_n: "divmod_nat_rel m n (m div n, m mod n)" . with assms have m_div_n: "m div n \ 1" - by (cases "m div n") (auto simp add: divmod_rel_def) - from assms divmod_m_n have "divmod_rel (m - n) n (m div n - Suc 0, m mod n)" - by (cases "m div n") (auto simp add: divmod_rel_def) - with divmod_eq have "divmod (m - n) n = (m div n - Suc 0, m mod n)" by simp - moreover from divmod_div_mod have "divmod (m - n) n = ((m - n) div n, (m - n) mod n)" . + by (cases "m div n") (auto simp add: divmod_nat_rel_def) + from assms divmod_nat_m_n have "divmod_nat_rel (m - n) n (m div n - Suc 0, m mod n)" + by (cases "m div n") (auto simp add: divmod_nat_rel_def) + with divmod_nat_eq have "divmod_nat (m - n) n = (m div n - Suc 0, m mod n)" by simp + moreover from divmod_nat_div_mod have "divmod_nat (m - n) n = ((m - n) div n, (m - n) mod n)" . ultimately have "m div n = Suc ((m - n) div n)" and "m mod n = (m - n) mod n" using m_div_n by simp_all - then show ?thesis using divmod_div_mod by simp + then show ?thesis using divmod_nat_div_mod by simp qed text {* The ''recursion'' equations for @{const div} and @{const mod} *} @@ -584,29 +576,29 @@ fixes m n :: nat assumes "m < n" shows "m div n = 0" - using assms divmod_base divmod_div_mod by simp + using assms divmod_nat_base divmod_nat_div_mod by simp lemma le_div_geq: fixes m n :: nat assumes "0 < n" and "n \ m" shows "m div n = Suc ((m - n) div n)" - using assms divmod_step divmod_div_mod by simp + using assms divmod_nat_step divmod_nat_div_mod by simp lemma mod_less [simp]: fixes m n :: nat assumes "m < n" shows "m mod n = m" - using assms divmod_base divmod_div_mod by simp + using assms divmod_nat_base divmod_nat_div_mod by simp lemma le_mod_geq: fixes m n :: nat assumes "n \ m" shows "m mod n = (m - n) mod n" - using assms divmod_step divmod_div_mod by (cases "n = 0") simp_all + using assms divmod_nat_step divmod_nat_div_mod by (cases "n = 0") simp_all instance proof - have [simp]: "\n::nat. n div 0 = 0" - by (simp add: div_nat_def divmod_zero) + by (simp add: div_nat_def divmod_nat_zero) have [simp]: "\n::nat. 0 div n = 0" proof - fix n :: nat @@ -616,7 +608,7 @@ show "OFCLASS(nat, semiring_div_class)" proof fix m n :: nat show "m div n * n + m mod n = m" - using divmod_rel [of m n] by (simp add: divmod_rel_def) + using divmod_nat_rel [of m n] by (simp add: divmod_nat_rel_def) next fix m n q :: nat assume "n \ 0" @@ -631,10 +623,10 @@ next case True with `m \ 0` have "m > 0" and "n > 0" and "q > 0" by auto - then have "\a b. divmod_rel n q (a, b) \ divmod_rel (m * n) (m * q) (a, m * b)" - by (auto simp add: divmod_rel_def) (simp_all add: algebra_simps) - moreover from divmod_rel have "divmod_rel n q (n div q, n mod q)" . - ultimately have "divmod_rel (m * n) (m * q) (n div q, m * (n mod q))" . + then have "\a b. divmod_nat_rel n q (a, b) \ divmod_nat_rel (m * n) (m * q) (a, m * b)" + by (auto simp add: divmod_nat_rel_def) (simp_all add: algebra_simps) + moreover from divmod_nat_rel have "divmod_nat_rel n q (n div q, n mod q)" . + ultimately have "divmod_nat_rel (m * n) (m * q) (n div q, m * (n mod q))" . then show ?thesis by (simp add: div_eq) qed qed simp_all @@ -642,6 +634,11 @@ end +lemma divmod_nat_if [code]: "divmod_nat m n = (if n = 0 \ m < n then (0, m) else + let (q, r) = divmod_nat (m - n) n in (Suc q, r))" +by (simp add: divmod_nat_zero divmod_nat_base divmod_nat_step) + (simp add: divmod_nat_div_mod) + text {* Simproc for cancelling @{const div} and @{const mod} *} ML {* @@ -674,22 +671,6 @@ end *} -text {* code generator setup *} - -lemma divmod_if [code]: "divmod m n = (if n = 0 \ m < n then (0, m) else - let (q, r) = divmod (m - n) n in (Suc q, r))" -by (simp add: divmod_zero divmod_base divmod_step) - (simp add: divmod_div_mod) - -code_modulename SML - Divides Nat - -code_modulename OCaml - Divides Nat - -code_modulename Haskell - Divides Nat - subsubsection {* Quotient *} @@ -712,7 +693,7 @@ fixes m n :: nat assumes "n > 0" shows "m mod n < (n::nat)" - using assms divmod_rel [of m n] unfolding divmod_rel_def by auto + using assms divmod_nat_rel [of m n] unfolding divmod_nat_rel_def by auto lemma mod_less_eq_dividend [simp]: fixes m n :: nat @@ -753,27 +734,27 @@ subsubsection {* Quotient and Remainder *} -lemma divmod_rel_mult1_eq: - "divmod_rel b c (q, r) \ c > 0 - \ divmod_rel (a * b) c (a * q + a * r div c, a * r mod c)" -by (auto simp add: split_ifs divmod_rel_def algebra_simps) +lemma divmod_nat_rel_mult1_eq: + "divmod_nat_rel b c (q, r) \ c > 0 + \ divmod_nat_rel (a * b) c (a * q + a * r div c, a * r mod c)" +by (auto simp add: split_ifs divmod_nat_rel_def algebra_simps) lemma div_mult1_eq: "(a * b) div c = a * (b div c) + a * (b mod c) div (c::nat)" apply (cases "c = 0", simp) -apply (blast intro: divmod_rel [THEN divmod_rel_mult1_eq, THEN div_eq]) +apply (blast intro: divmod_nat_rel [THEN divmod_nat_rel_mult1_eq, THEN div_eq]) done -lemma divmod_rel_add1_eq: - "divmod_rel a c (aq, ar) \ divmod_rel b c (bq, br) \ c > 0 - \ divmod_rel (a + b) c (aq + bq + (ar + br) div c, (ar + br) mod c)" -by (auto simp add: split_ifs divmod_rel_def algebra_simps) +lemma divmod_nat_rel_add1_eq: + "divmod_nat_rel a c (aq, ar) \ divmod_nat_rel b c (bq, br) \ c > 0 + \ divmod_nat_rel (a + b) c (aq + bq + (ar + br) div c, (ar + br) mod c)" +by (auto simp add: split_ifs divmod_nat_rel_def algebra_simps) (*NOT suitable for rewriting: the RHS has an instance of the LHS*) lemma div_add1_eq: "(a+b) div (c::nat) = a div c + b div c + ((a mod c + b mod c) div c)" apply (cases "c = 0", simp) -apply (blast intro: divmod_rel_add1_eq [THEN div_eq] divmod_rel) +apply (blast intro: divmod_nat_rel_add1_eq [THEN div_eq] divmod_nat_rel) done lemma mod_lemma: "[| (0::nat) < c; r < b |] ==> b * (q mod c) + r < b * c" @@ -783,21 +764,21 @@ apply (simp add: add_mult_distrib2) done -lemma divmod_rel_mult2_eq: - "divmod_rel a b (q, r) \ 0 < b \ 0 < c - \ divmod_rel a (b * c) (q div c, b *(q mod c) + r)" -by (auto simp add: mult_ac divmod_rel_def add_mult_distrib2 [symmetric] mod_lemma) +lemma divmod_nat_rel_mult2_eq: + "divmod_nat_rel a b (q, r) \ 0 < b \ 0 < c + \ divmod_nat_rel a (b * c) (q div c, b *(q mod c) + r)" +by (auto simp add: mult_ac divmod_nat_rel_def add_mult_distrib2 [symmetric] mod_lemma) lemma div_mult2_eq: "a div (b*c) = (a div b) div (c::nat)" apply (cases "b = 0", simp) apply (cases "c = 0", simp) - apply (force simp add: divmod_rel [THEN divmod_rel_mult2_eq, THEN div_eq]) + apply (force simp add: divmod_nat_rel [THEN divmod_nat_rel_mult2_eq, THEN div_eq]) done lemma mod_mult2_eq: "a mod (b*c) = b*(a div b mod c) + a mod (b::nat)" apply (cases "b = 0", simp) apply (cases "c = 0", simp) - apply (auto simp add: mult_commute divmod_rel [THEN divmod_rel_mult2_eq, THEN mod_eq]) + apply (auto simp add: mult_commute divmod_nat_rel [THEN divmod_nat_rel_mult2_eq, THEN mod_eq]) done @@ -944,9 +925,9 @@ from A B show ?lhs .. next assume P: ?lhs - then have "divmod_rel m n (q, m - n * q)" - unfolding divmod_rel_def by (auto simp add: mult_ac) - with divmod_rel_unique divmod_rel [of m n] + then have "divmod_nat_rel m n (q, m - n * q)" + unfolding divmod_nat_rel_def by (auto simp add: mult_ac) + with divmod_nat_rel_unique divmod_nat_rel [of m n] have "(q, m - n * q) = (m div n, m mod n)" by auto then show ?rhs by simp qed @@ -1145,113 +1126,1347 @@ declare Suc_mod_eq_add3_mod_number_of [simp] -subsection {* Proof Tools setup; Combination and Cancellation Simprocs *} - -declare split_div[of _ _ "number_of k", standard, arith_split] -declare split_mod[of _ _ "number_of k", standard, arith_split] - - -subsubsection{*For @{text combine_numerals}*} - -lemma left_add_mult_distrib: "i*u + (j*u + k) = (i+j)*u + (k::nat)" -by (simp add: add_mult_distrib) - - -subsubsection{*For @{text cancel_numerals}*} - -lemma nat_diff_add_eq1: - "j <= (i::nat) ==> ((i*u + m) - (j*u + n)) = (((i-j)*u + m) - n)" -by (simp split add: nat_diff_split add: add_mult_distrib) - -lemma nat_diff_add_eq2: - "i <= (j::nat) ==> ((i*u + m) - (j*u + n)) = (m - ((j-i)*u + n))" -by (simp split add: nat_diff_split add: add_mult_distrib) - -lemma nat_eq_add_iff1: - "j <= (i::nat) ==> (i*u + m = j*u + n) = ((i-j)*u + m = n)" -by (auto split add: nat_diff_split simp add: add_mult_distrib) - -lemma nat_eq_add_iff2: - "i <= (j::nat) ==> (i*u + m = j*u + n) = (m = (j-i)*u + n)" -by (auto split add: nat_diff_split simp add: add_mult_distrib) - -lemma nat_less_add_iff1: - "j <= (i::nat) ==> (i*u + m < j*u + n) = ((i-j)*u + m < n)" -by (auto split add: nat_diff_split simp add: add_mult_distrib) - -lemma nat_less_add_iff2: - "i <= (j::nat) ==> (i*u + m < j*u + n) = (m < (j-i)*u + n)" -by (auto split add: nat_diff_split simp add: add_mult_distrib) - -lemma nat_le_add_iff1: - "j <= (i::nat) ==> (i*u + m <= j*u + n) = ((i-j)*u + m <= n)" -by (auto split add: nat_diff_split simp add: add_mult_distrib) - -lemma nat_le_add_iff2: - "i <= (j::nat) ==> (i*u + m <= j*u + n) = (m <= (j-i)*u + n)" -by (auto split add: nat_diff_split simp add: add_mult_distrib) - - -subsubsection{*For @{text cancel_numeral_factors} *} - -lemma nat_mult_le_cancel1: "(0::nat) < k ==> (k*m <= k*n) = (m<=n)" +lemma Suc_times_mod_eq: "1 Suc (k * m) mod k = 1" +apply (induct "m") +apply (simp_all add: mod_Suc) +done + +declare Suc_times_mod_eq [of "number_of w", standard, simp] + +lemma [simp]: "n div k \ (Suc n) div k" +by (simp add: div_le_mono) + +lemma Suc_n_div_2_gt_zero [simp]: "(0::nat) < n ==> 0 < (n + 1) div 2" +by (cases n) simp_all + +lemma div_2_gt_zero [simp]: "(1::nat) < n ==> 0 < n div 2" +using Suc_n_div_2_gt_zero [of "n - 1"] by simp + + (* Potential use of algebra : Equality modulo n*) +lemma mod_mult_self3 [simp]: "(k*n + m) mod n = m mod (n::nat)" +by (simp add: mult_ac add_ac) + +lemma mod_mult_self4 [simp]: "Suc (k*n + m) mod n = Suc m mod n" +proof - + have "Suc (k * n + m) mod n = (k * n + Suc m) mod n" by simp + also have "... = Suc m mod n" by (rule mod_mult_self3) + finally show ?thesis . +qed + +lemma mod_Suc_eq_Suc_mod: "Suc m mod n = Suc (m mod n) mod n" +apply (subst mod_Suc [of m]) +apply (subst mod_Suc [of "m mod n"], simp) +done + + +subsection {* Division on @{typ int} *} + +definition divmod_int_rel :: "int \ int \ int \ int \ bool" where + --{*definition of quotient and remainder*} + [code]: "divmod_int_rel a b = (\(q, r). a = b * q + r \ + (if 0 < b then 0 \ r \ r < b else b < r \ r \ 0))" + +definition adjust :: "int \ int \ int \ int \ int" where + --{*for the division algorithm*} + [code]: "adjust b = (\(q, r). if 0 \ r - b then (2 * q + 1, r - b) + else (2 * q, r))" + +text{*algorithm for the case @{text "a\0, b>0"}*} +function posDivAlg :: "int \ int \ int \ int" where + "posDivAlg a b = (if a < b \ b \ 0 then (0, a) + else adjust b (posDivAlg a (2 * b)))" by auto - -lemma nat_mult_less_cancel1: "(0::nat) < k ==> (k*m < k*n) = (m(a, b). nat (a - b + 1))") + (auto simp add: mult_2) + +text{*algorithm for the case @{text "a<0, b>0"}*} +function negDivAlg :: "int \ int \ int \ int" where + "negDivAlg a b = (if 0 \a + b \ b \ 0 then (-1, a + b) + else adjust b (negDivAlg a (2 * b)))" by auto - -lemma nat_mult_eq_cancel1: "(0::nat) < k ==> (k*m = k*n) = (m=n)" -by auto - -lemma nat_mult_div_cancel1: "(0::nat) < k ==> (k*m) div (k*n) = (m div n)" -by auto - -lemma nat_mult_dvd_cancel_disj[simp]: - "(k*m) dvd (k*n) = (k=0 | m dvd (n::nat))" -by(auto simp: dvd_eq_mod_eq_0 mod_mult_distrib2[symmetric]) - -lemma nat_mult_dvd_cancel1: "0 < k \ (k*m) dvd (k*n::nat) = (m dvd n)" -by(auto) - - -subsubsection{*For @{text cancel_factor} *} - -lemma nat_mult_le_cancel_disj: "(k*m <= k*n) = ((0::nat) < k --> m<=n)" -by auto - -lemma nat_mult_less_cancel_disj: "(k*m < k*n) = ((0::nat) < k & m Lin_Arith.add_simps (@{thms ring_distribs} @ [@{thm Let_number_of}, @{thm Let_0}, @{thm Let_1}, - @{thm nat_0}, @{thm nat_1}, - @{thm add_nat_number_of}, @{thm diff_nat_number_of}, @{thm mult_nat_number_of}, - @{thm eq_nat_number_of}, @{thm less_nat_number_of}, @{thm le_number_of_eq_not_less}, - @{thm le_Suc_number_of}, @{thm le_number_of_Suc}, - @{thm less_Suc_number_of}, @{thm less_number_of_Suc}, - @{thm Suc_eq_number_of}, @{thm eq_number_of_Suc}, - @{thm mult_Suc}, @{thm mult_Suc_right}, - @{thm add_Suc}, @{thm add_Suc_right}, - @{thm eq_number_of_0}, @{thm eq_0_number_of}, @{thm less_0_number_of}, - @{thm of_int_number_of_eq}, @{thm of_nat_number_of_eq}, @{thm nat_number_of}, - @{thm if_True}, @{thm if_False}]) - #> Lin_Arith.add_simprocs (Numeral_Simprocs.assoc_fold_simproc - :: Numeral_Simprocs.combine_numerals - :: Numeral_Simprocs.cancel_numerals) - #> Lin_Arith.add_simprocs (Nat_Numeral_Simprocs.combine_numerals :: Nat_Numeral_Simprocs.cancel_numerals)) +termination by (relation "measure (\(a, b). nat (- a - b))") + (auto simp add: mult_2) + +text{*algorithm for the general case @{term "b\0"}*} +definition negateSnd :: "int \ int \ int \ int" where + [code_unfold]: "negateSnd = apsnd uminus" + +definition divmod_int :: "int \ int \ int \ int" where + --{*The full division algorithm considers all possible signs for a, b + including the special case @{text "a=0, b<0"} because + @{term negDivAlg} requires @{term "a<0"}.*} + "divmod_int a b = (if 0 \ a then if 0 \ b then posDivAlg a b + else if a = 0 then (0, 0) + else negateSnd (negDivAlg (-a) (-b)) + else + if 0 < b then negDivAlg a b + else negateSnd (posDivAlg (-a) (-b)))" + +instantiation int :: Divides.div +begin + +definition + "a div b = fst (divmod_int a b)" + +definition + "a mod b = snd (divmod_int a b)" + +instance .. + +end + +lemma divmod_int_mod_div: + "divmod_int p q = (p div q, p mod q)" + by (auto simp add: div_int_def mod_int_def) + +text{* +Here is the division algorithm in ML: + +\begin{verbatim} + fun posDivAlg (a,b) = + if ar-b then (2*q+1, r-b) else (2*q, r) + end + + fun negDivAlg (a,b) = + if 0\a+b then (~1,a+b) + else let val (q,r) = negDivAlg(a, 2*b) + in if 0\r-b then (2*q+1, r-b) else (2*q, r) + end; + + fun negateSnd (q,r:int) = (q,~r); + + fun divmod (a,b) = if 0\a then + if b>0 then posDivAlg (a,b) + else if a=0 then (0,0) + else negateSnd (negDivAlg (~a,~b)) + else + if 0 b*q + r; 0 \ r'; r' < b; r < b |] + ==> q' \ (q::int)" +apply (subgoal_tac "r' + b * (q'-q) \ r") + prefer 2 apply (simp add: right_diff_distrib) +apply (subgoal_tac "0 < b * (1 + q - q') ") +apply (erule_tac [2] order_le_less_trans) + prefer 2 apply (simp add: right_diff_distrib right_distrib) +apply (subgoal_tac "b * q' < b * (1 + q) ") + prefer 2 apply (simp add: right_diff_distrib right_distrib) +apply (simp add: mult_less_cancel_left) +done + +lemma unique_quotient_lemma_neg: + "[| b*q' + r' \ b*q + r; r \ 0; b < r; b < r' |] + ==> q \ (q'::int)" +by (rule_tac b = "-b" and r = "-r'" and r' = "-r" in unique_quotient_lemma, + auto) + +lemma unique_quotient: + "[| divmod_int_rel a b (q, r); divmod_int_rel a b (q', r'); b \ 0 |] + ==> q = q'" +apply (simp add: divmod_int_rel_def linorder_neq_iff split: split_if_asm) +apply (blast intro: order_antisym + dest: order_eq_refl [THEN unique_quotient_lemma] + order_eq_refl [THEN unique_quotient_lemma_neg] sym)+ +done + + +lemma unique_remainder: + "[| divmod_int_rel a b (q, r); divmod_int_rel a b (q', r'); b \ 0 |] + ==> r = r'" +apply (subgoal_tac "q = q'") + apply (simp add: divmod_int_rel_def) +apply (blast intro: unique_quotient) +done + + +subsubsection{*Correctness of @{term posDivAlg}, the Algorithm for Non-Negative Dividends*} + +text{*And positive divisors*} + +lemma adjust_eq [simp]: + "adjust b (q,r) = + (let diff = r-b in + if 0 \ diff then (2*q + 1, diff) + else (2*q, r))" +by (simp add: Let_def adjust_def) + +declare posDivAlg.simps [simp del] + +text{*use with a simproc to avoid repeatedly proving the premise*} +lemma posDivAlg_eqn: + "0 < b ==> + posDivAlg a b = (if a a" and "0 < b" + shows "divmod_int_rel a b (posDivAlg a b)" +using prems apply (induct a b rule: posDivAlg.induct) +apply auto +apply (simp add: divmod_int_rel_def) +apply (subst posDivAlg_eqn, simp add: right_distrib) +apply (case_tac "a < b") +apply simp_all +apply (erule splitE) +apply (auto simp add: right_distrib Let_def mult_ac mult_2_right) +done + + +subsubsection{*Correctness of @{term negDivAlg}, the Algorithm for Negative Dividends*} + +text{*And positive divisors*} + +declare negDivAlg.simps [simp del] + +text{*use with a simproc to avoid repeatedly proving the premise*} +lemma negDivAlg_eqn: + "0 < b ==> + negDivAlg a b = + (if 0\a+b then (-1,a+b) else adjust b (negDivAlg a (2*b)))" +by (rule negDivAlg.simps [THEN trans], simp) + +(*Correctness of negDivAlg: it computes quotients correctly + It doesn't work if a=0 because the 0/b equals 0, not -1*) +lemma negDivAlg_correct: + assumes "a < 0" and "b > 0" + shows "divmod_int_rel a b (negDivAlg a b)" +using prems apply (induct a b rule: negDivAlg.induct) +apply (auto simp add: linorder_not_le) +apply (simp add: divmod_int_rel_def) +apply (subst negDivAlg_eqn, assumption) +apply (case_tac "a + b < (0\int)") +apply simp_all +apply (erule splitE) +apply (auto simp add: right_distrib Let_def mult_ac mult_2_right) +done + + +subsubsection{*Existence Shown by Proving the Division Algorithm to be Correct*} + +(*the case a=0*) +lemma divmod_int_rel_0: "b \ 0 ==> divmod_int_rel 0 b (0, 0)" +by (auto simp add: divmod_int_rel_def linorder_neq_iff) + +lemma posDivAlg_0 [simp]: "posDivAlg 0 b = (0, 0)" +by (subst posDivAlg.simps, auto) + +lemma negDivAlg_minus1 [simp]: "negDivAlg -1 b = (-1, b - 1)" +by (subst negDivAlg.simps, auto) + +lemma negateSnd_eq [simp]: "negateSnd(q,r) = (q,-r)" +by (simp add: negateSnd_def) + +lemma divmod_int_rel_neg: "divmod_int_rel (-a) (-b) qr ==> divmod_int_rel a b (negateSnd qr)" +by (auto simp add: split_ifs divmod_int_rel_def) + +lemma divmod_int_correct: "b \ 0 ==> divmod_int_rel a b (divmod_int a b)" +by (force simp add: linorder_neq_iff divmod_int_rel_0 divmod_int_def divmod_int_rel_neg + posDivAlg_correct negDivAlg_correct) + +text{*Arbitrary definitions for division by zero. Useful to simplify + certain equations.*} + +lemma DIVISION_BY_ZERO [simp]: "a div (0::int) = 0 & a mod (0::int) = a" +by (simp add: div_int_def mod_int_def divmod_int_def posDivAlg.simps) + + +text{*Basic laws about division and remainder*} + +lemma zmod_zdiv_equality: "(a::int) = b * (a div b) + (a mod b)" +apply (case_tac "b = 0", simp) +apply (cut_tac a = a and b = b in divmod_int_correct) +apply (auto simp add: divmod_int_rel_def div_int_def mod_int_def) +done + +lemma zdiv_zmod_equality: "(b * (a div b) + (a mod b)) + k = (a::int)+k" +by(simp add: zmod_zdiv_equality[symmetric]) + +lemma zdiv_zmod_equality2: "((a div b) * b + (a mod b)) + k = (a::int)+k" +by(simp add: mult_commute zmod_zdiv_equality[symmetric]) + +text {* Tool setup *} + +ML {* +local + +structure CancelDivMod = CancelDivModFun(struct + + val div_name = @{const_name div}; + val mod_name = @{const_name mod}; + val mk_binop = HOLogic.mk_binop; + val mk_sum = Arith_Data.mk_sum HOLogic.intT; + val dest_sum = Arith_Data.dest_sum; + + val div_mod_eqs = map mk_meta_eq [@{thm zdiv_zmod_equality}, @{thm zdiv_zmod_equality2}]; + + val trans = trans; + + val prove_eq_sums = Arith_Data.prove_conv2 all_tac (Arith_Data.simp_all_tac + (@{thm diff_minus} :: @{thms add_0s} @ @{thms add_ac})) + +end) + +in + +val cancel_div_mod_int_proc = Simplifier.simproc @{theory} + "cancel_zdiv_zmod" ["(k::int) + l"] (K CancelDivMod.proc); + +val _ = Addsimprocs [cancel_div_mod_int_proc]; + end +*} + +lemma pos_mod_conj : "(0::int) < b ==> 0 \ a mod b & a mod b < b" +apply (cut_tac a = a and b = b in divmod_int_correct) +apply (auto simp add: divmod_int_rel_def mod_int_def) +done + +lemmas pos_mod_sign [simp] = pos_mod_conj [THEN conjunct1, standard] + and pos_mod_bound [simp] = pos_mod_conj [THEN conjunct2, standard] + +lemma neg_mod_conj : "b < (0::int) ==> a mod b \ 0 & b < a mod b" +apply (cut_tac a = a and b = b in divmod_int_correct) +apply (auto simp add: divmod_int_rel_def div_int_def mod_int_def) +done + +lemmas neg_mod_sign [simp] = neg_mod_conj [THEN conjunct1, standard] + and neg_mod_bound [simp] = neg_mod_conj [THEN conjunct2, standard] + + + +subsubsection{*General Properties of div and mod*} + +lemma divmod_int_rel_div_mod: "b \ 0 ==> divmod_int_rel a b (a div b, a mod b)" +apply (cut_tac a = a and b = b in zmod_zdiv_equality) +apply (force simp add: divmod_int_rel_def linorder_neq_iff) +done + +lemma divmod_int_rel_div: "[| divmod_int_rel a b (q, r); b \ 0 |] ==> a div b = q" +by (simp add: divmod_int_rel_div_mod [THEN unique_quotient]) + +lemma divmod_int_rel_mod: "[| divmod_int_rel a b (q, r); b \ 0 |] ==> a mod b = r" +by (simp add: divmod_int_rel_div_mod [THEN unique_remainder]) + +lemma div_pos_pos_trivial: "[| (0::int) \ a; a < b |] ==> a div b = 0" +apply (rule divmod_int_rel_div) +apply (auto simp add: divmod_int_rel_def) +done + +lemma div_neg_neg_trivial: "[| a \ (0::int); b < a |] ==> a div b = 0" +apply (rule divmod_int_rel_div) +apply (auto simp add: divmod_int_rel_def) +done + +lemma div_pos_neg_trivial: "[| (0::int) < a; a+b \ 0 |] ==> a div b = -1" +apply (rule divmod_int_rel_div) +apply (auto simp add: divmod_int_rel_def) +done + +(*There is no div_neg_pos_trivial because 0 div b = 0 would supersede it*) + +lemma mod_pos_pos_trivial: "[| (0::int) \ a; a < b |] ==> a mod b = a" +apply (rule_tac q = 0 in divmod_int_rel_mod) +apply (auto simp add: divmod_int_rel_def) +done + +lemma mod_neg_neg_trivial: "[| a \ (0::int); b < a |] ==> a mod b = a" +apply (rule_tac q = 0 in divmod_int_rel_mod) +apply (auto simp add: divmod_int_rel_def) +done + +lemma mod_pos_neg_trivial: "[| (0::int) < a; a+b \ 0 |] ==> a mod b = a+b" +apply (rule_tac q = "-1" in divmod_int_rel_mod) +apply (auto simp add: divmod_int_rel_def) +done + +text{*There is no @{text mod_neg_pos_trivial}.*} + + +(*Simpler laws such as -a div b = -(a div b) FAIL, but see just below*) +lemma zdiv_zminus_zminus [simp]: "(-a) div (-b) = a div (b::int)" +apply (case_tac "b = 0", simp) +apply (simp add: divmod_int_rel_div_mod [THEN divmod_int_rel_neg, simplified, + THEN divmod_int_rel_div, THEN sym]) + +done + +(*Simpler laws such as -a mod b = -(a mod b) FAIL, but see just below*) +lemma zmod_zminus_zminus [simp]: "(-a) mod (-b) = - (a mod (b::int))" +apply (case_tac "b = 0", simp) +apply (subst divmod_int_rel_div_mod [THEN divmod_int_rel_neg, simplified, THEN divmod_int_rel_mod], + auto) +done + + +subsubsection{*Laws for div and mod with Unary Minus*} + +lemma zminus1_lemma: + "divmod_int_rel a b (q, r) + ==> divmod_int_rel (-a) b (if r=0 then -q else -q - 1, + if r=0 then 0 else b-r)" +by (force simp add: split_ifs divmod_int_rel_def linorder_neq_iff right_diff_distrib) + + +lemma zdiv_zminus1_eq_if: + "b \ (0::int) + ==> (-a) div b = + (if a mod b = 0 then - (a div b) else - (a div b) - 1)" +by (blast intro: divmod_int_rel_div_mod [THEN zminus1_lemma, THEN divmod_int_rel_div]) + +lemma zmod_zminus1_eq_if: + "(-a::int) mod b = (if a mod b = 0 then 0 else b - (a mod b))" +apply (case_tac "b = 0", simp) +apply (blast intro: divmod_int_rel_div_mod [THEN zminus1_lemma, THEN divmod_int_rel_mod]) +done + +lemma zmod_zminus1_not_zero: + fixes k l :: int + shows "- k mod l \ 0 \ k mod l \ 0" + unfolding zmod_zminus1_eq_if by auto + +lemma zdiv_zminus2: "a div (-b) = (-a::int) div b" +by (cut_tac a = "-a" in zdiv_zminus_zminus, auto) + +lemma zmod_zminus2: "a mod (-b) = - ((-a::int) mod b)" +by (cut_tac a = "-a" and b = b in zmod_zminus_zminus, auto) + +lemma zdiv_zminus2_eq_if: + "b \ (0::int) + ==> a div (-b) = + (if a mod b = 0 then - (a div b) else - (a div b) - 1)" +by (simp add: zdiv_zminus1_eq_if zdiv_zminus2) + +lemma zmod_zminus2_eq_if: + "a mod (-b::int) = (if a mod b = 0 then 0 else (a mod b) - b)" +by (simp add: zmod_zminus1_eq_if zmod_zminus2) + +lemma zmod_zminus2_not_zero: + fixes k l :: int + shows "k mod - l \ 0 \ k mod l \ 0" + unfolding zmod_zminus2_eq_if by auto + + +subsubsection{*Division of a Number by Itself*} + +lemma self_quotient_aux1: "[| (0::int) < a; a = r + a*q; r < a |] ==> 1 \ q" +apply (subgoal_tac "0 < a*q") + apply (simp add: zero_less_mult_iff, arith) +done + +lemma self_quotient_aux2: "[| (0::int) < a; a = r + a*q; 0 \ r |] ==> q \ 1" +apply (subgoal_tac "0 \ a* (1-q) ") + apply (simp add: zero_le_mult_iff) +apply (simp add: right_diff_distrib) +done + +lemma self_quotient: "[| divmod_int_rel a a (q, r); a \ (0::int) |] ==> q = 1" +apply (simp add: split_ifs divmod_int_rel_def linorder_neq_iff) +apply (rule order_antisym, safe, simp_all) +apply (rule_tac [3] a = "-a" and r = "-r" in self_quotient_aux1) +apply (rule_tac a = "-a" and r = "-r" in self_quotient_aux2) +apply (force intro: self_quotient_aux1 self_quotient_aux2 simp add: add_commute)+ +done + +lemma self_remainder: "[| divmod_int_rel a a (q, r); a \ (0::int) |] ==> r = 0" +apply (frule self_quotient, assumption) +apply (simp add: divmod_int_rel_def) +done + +lemma zdiv_self [simp]: "a \ 0 ==> a div a = (1::int)" +by (simp add: divmod_int_rel_div_mod [THEN self_quotient]) + +(*Here we have 0 mod 0 = 0, also assumed by Knuth (who puts m mod 0 = 0) *) +lemma zmod_self [simp]: "a mod a = (0::int)" +apply (case_tac "a = 0", simp) +apply (simp add: divmod_int_rel_div_mod [THEN self_remainder]) +done + + +subsubsection{*Computation of Division and Remainder*} + +lemma zdiv_zero [simp]: "(0::int) div b = 0" +by (simp add: div_int_def divmod_int_def) + +lemma div_eq_minus1: "(0::int) < b ==> -1 div b = -1" +by (simp add: div_int_def divmod_int_def) + +lemma zmod_zero [simp]: "(0::int) mod b = 0" +by (simp add: mod_int_def divmod_int_def) + +lemma zmod_minus1: "(0::int) < b ==> -1 mod b = b - 1" +by (simp add: mod_int_def divmod_int_def) + +text{*a positive, b positive *} + +lemma div_pos_pos: "[| 0 < a; 0 \ b |] ==> a div b = fst (posDivAlg a b)" +by (simp add: div_int_def divmod_int_def) + +lemma mod_pos_pos: "[| 0 < a; 0 \ b |] ==> a mod b = snd (posDivAlg a b)" +by (simp add: mod_int_def divmod_int_def) + +text{*a negative, b positive *} + +lemma div_neg_pos: "[| a < 0; 0 < b |] ==> a div b = fst (negDivAlg a b)" +by (simp add: div_int_def divmod_int_def) + +lemma mod_neg_pos: "[| a < 0; 0 < b |] ==> a mod b = snd (negDivAlg a b)" +by (simp add: mod_int_def divmod_int_def) + +text{*a positive, b negative *} + +lemma div_pos_neg: + "[| 0 < a; b < 0 |] ==> a div b = fst (negateSnd (negDivAlg (-a) (-b)))" +by (simp add: div_int_def divmod_int_def) + +lemma mod_pos_neg: + "[| 0 < a; b < 0 |] ==> a mod b = snd (negateSnd (negDivAlg (-a) (-b)))" +by (simp add: mod_int_def divmod_int_def) + +text{*a negative, b negative *} + +lemma div_neg_neg: + "[| a < 0; b \ 0 |] ==> a div b = fst (negateSnd (posDivAlg (-a) (-b)))" +by (simp add: div_int_def divmod_int_def) + +lemma mod_neg_neg: + "[| a < 0; b \ 0 |] ==> a mod b = snd (negateSnd (posDivAlg (-a) (-b)))" +by (simp add: mod_int_def divmod_int_def) + +text {*Simplify expresions in which div and mod combine numerical constants*} + +lemma divmod_int_relI: + "\a == b * q + r; if 0 < b then 0 \ r \ r < b else b < r \ r \ 0\ + \ divmod_int_rel a b (q, r)" + unfolding divmod_int_rel_def by simp + +lemmas divmod_int_rel_div_eq = divmod_int_relI [THEN divmod_int_rel_div, THEN eq_reflection] +lemmas divmod_int_rel_mod_eq = divmod_int_relI [THEN divmod_int_rel_mod, THEN eq_reflection] +lemmas arithmetic_simps = + arith_simps + add_special + OrderedGroup.add_0_left + OrderedGroup.add_0_right + mult_zero_left + mult_zero_right + mult_1_left + mult_1_right + +(* simprocs adapted from HOL/ex/Binary.thy *) +ML {* +local + val mk_number = HOLogic.mk_number HOLogic.intT; + fun mk_cert u k l = @{term "plus :: int \ int \ int"} $ + (@{term "times :: int \ int \ int"} $ u $ mk_number k) $ + mk_number l; + fun prove ctxt prop = Goal.prove ctxt [] [] prop + (K (ALLGOALS (full_simp_tac (HOL_basic_ss addsimps @{thms arithmetic_simps})))); + fun binary_proc proc ss ct = + (case Thm.term_of ct of + _ $ t $ u => + (case try (pairself (`(snd o HOLogic.dest_number))) (t, u) of + SOME args => proc (Simplifier.the_context ss) args + | NONE => NONE) + | _ => NONE); +in + fun divmod_proc rule = binary_proc (fn ctxt => fn ((m, t), (n, u)) => + if n = 0 then NONE + else let val (k, l) = Integer.div_mod m n; + in SOME (rule OF [prove ctxt (Logic.mk_equals (t, mk_cert u k l))]) end); +end +*} + +simproc_setup binary_int_div ("number_of m div number_of n :: int") = + {* K (divmod_proc (@{thm divmod_int_rel_div_eq})) *} + +simproc_setup binary_int_mod ("number_of m mod number_of n :: int") = + {* K (divmod_proc (@{thm divmod_int_rel_mod_eq})) *} + +lemmas posDivAlg_eqn_number_of [simp] = + posDivAlg_eqn [of "number_of v" "number_of w", standard] + +lemmas negDivAlg_eqn_number_of [simp] = + negDivAlg_eqn [of "number_of v" "number_of w", standard] + + +text{*Special-case simplification *} + +lemma zmod_minus1_right [simp]: "a mod (-1::int) = 0" +apply (cut_tac a = a and b = "-1" in neg_mod_sign) +apply (cut_tac [2] a = a and b = "-1" in neg_mod_bound) +apply (auto simp del: neg_mod_sign neg_mod_bound) +done + +lemma zdiv_minus1_right [simp]: "a div (-1::int) = -a" +by (cut_tac a = a and b = "-1" in zmod_zdiv_equality, auto) + +(** The last remaining special cases for constant arithmetic: + 1 div z and 1 mod z **) + +lemmas div_pos_pos_1_number_of [simp] = + div_pos_pos [OF int_0_less_1, of "number_of w", standard] + +lemmas div_pos_neg_1_number_of [simp] = + div_pos_neg [OF int_0_less_1, of "number_of w", standard] + +lemmas mod_pos_pos_1_number_of [simp] = + mod_pos_pos [OF int_0_less_1, of "number_of w", standard] + +lemmas mod_pos_neg_1_number_of [simp] = + mod_pos_neg [OF int_0_less_1, of "number_of w", standard] + + +lemmas posDivAlg_eqn_1_number_of [simp] = + posDivAlg_eqn [of concl: 1 "number_of w", standard] + +lemmas negDivAlg_eqn_1_number_of [simp] = + negDivAlg_eqn [of concl: 1 "number_of w", standard] + + + +subsubsection{*Monotonicity in the First Argument (Dividend)*} + +lemma zdiv_mono1: "[| a \ a'; 0 < (b::int) |] ==> a div b \ a' div b" +apply (cut_tac a = a and b = b in zmod_zdiv_equality) +apply (cut_tac a = a' and b = b in zmod_zdiv_equality) +apply (rule unique_quotient_lemma) +apply (erule subst) +apply (erule subst, simp_all) +done + +lemma zdiv_mono1_neg: "[| a \ a'; (b::int) < 0 |] ==> a' div b \ a div b" +apply (cut_tac a = a and b = b in zmod_zdiv_equality) +apply (cut_tac a = a' and b = b in zmod_zdiv_equality) +apply (rule unique_quotient_lemma_neg) +apply (erule subst) +apply (erule subst, simp_all) +done + + +subsubsection{*Monotonicity in the Second Argument (Divisor)*} + +lemma q_pos_lemma: + "[| 0 \ b'*q' + r'; r' < b'; 0 < b' |] ==> 0 \ (q'::int)" +apply (subgoal_tac "0 < b'* (q' + 1) ") + apply (simp add: zero_less_mult_iff) +apply (simp add: right_distrib) +done + +lemma zdiv_mono2_lemma: + "[| b*q + r = b'*q' + r'; 0 \ b'*q' + r'; + r' < b'; 0 \ r; 0 < b'; b' \ b |] + ==> q \ (q'::int)" +apply (frule q_pos_lemma, assumption+) +apply (subgoal_tac "b*q < b* (q' + 1) ") + apply (simp add: mult_less_cancel_left) +apply (subgoal_tac "b*q = r' - r + b'*q'") + prefer 2 apply simp +apply (simp (no_asm_simp) add: right_distrib) +apply (subst add_commute, rule zadd_zless_mono, arith) +apply (rule mult_right_mono, auto) +done + +lemma zdiv_mono2: + "[| (0::int) \ a; 0 < b'; b' \ b |] ==> a div b \ a div b'" +apply (subgoal_tac "b \ 0") + prefer 2 apply arith +apply (cut_tac a = a and b = b in zmod_zdiv_equality) +apply (cut_tac a = a and b = b' in zmod_zdiv_equality) +apply (rule zdiv_mono2_lemma) +apply (erule subst) +apply (erule subst, simp_all) +done + +lemma q_neg_lemma: + "[| b'*q' + r' < 0; 0 \ r'; 0 < b' |] ==> q' \ (0::int)" +apply (subgoal_tac "b'*q' < 0") + apply (simp add: mult_less_0_iff, arith) +done + +lemma zdiv_mono2_neg_lemma: + "[| b*q + r = b'*q' + r'; b'*q' + r' < 0; + r < b; 0 \ r'; 0 < b'; b' \ b |] + ==> q' \ (q::int)" +apply (frule q_neg_lemma, assumption+) +apply (subgoal_tac "b*q' < b* (q + 1) ") + apply (simp add: mult_less_cancel_left) +apply (simp add: right_distrib) +apply (subgoal_tac "b*q' \ b'*q'") + prefer 2 apply (simp add: mult_right_mono_neg, arith) +done + +lemma zdiv_mono2_neg: + "[| a < (0::int); 0 < b'; b' \ b |] ==> a div b' \ a div b" +apply (cut_tac a = a and b = b in zmod_zdiv_equality) +apply (cut_tac a = a and b = b' in zmod_zdiv_equality) +apply (rule zdiv_mono2_neg_lemma) +apply (erule subst) +apply (erule subst, simp_all) +done + + +subsubsection{*More Algebraic Laws for div and mod*} + +text{*proving (a*b) div c = a * (b div c) + a * (b mod c) *} + +lemma zmult1_lemma: + "[| divmod_int_rel b c (q, r); c \ 0 |] + ==> divmod_int_rel (a * b) c (a*q + a*r div c, a*r mod c)" +by (auto simp add: split_ifs divmod_int_rel_def linorder_neq_iff right_distrib mult_ac) + +lemma zdiv_zmult1_eq: "(a*b) div c = a*(b div c) + a*(b mod c) div (c::int)" +apply (case_tac "c = 0", simp) +apply (blast intro: divmod_int_rel_div_mod [THEN zmult1_lemma, THEN divmod_int_rel_div]) +done + +lemma zmod_zmult1_eq: "(a*b) mod c = a*(b mod c) mod (c::int)" +apply (case_tac "c = 0", simp) +apply (blast intro: divmod_int_rel_div_mod [THEN zmult1_lemma, THEN divmod_int_rel_mod]) +done + +lemma zmod_zdiv_trivial: "(a mod b) div b = (0::int)" +apply (case_tac "b = 0", simp) +apply (auto simp add: linorder_neq_iff div_pos_pos_trivial div_neg_neg_trivial) +done + +text{*proving (a+b) div c = a div c + b div c + ((a mod c + b mod c) div c) *} + +lemma zadd1_lemma: + "[| divmod_int_rel a c (aq, ar); divmod_int_rel b c (bq, br); c \ 0 |] + ==> divmod_int_rel (a+b) c (aq + bq + (ar+br) div c, (ar+br) mod c)" +by (force simp add: split_ifs divmod_int_rel_def linorder_neq_iff right_distrib) + +(*NOT suitable for rewriting: the RHS has an instance of the LHS*) +lemma zdiv_zadd1_eq: + "(a+b) div (c::int) = a div c + b div c + ((a mod c + b mod c) div c)" +apply (case_tac "c = 0", simp) +apply (blast intro: zadd1_lemma [OF divmod_int_rel_div_mod divmod_int_rel_div_mod] divmod_int_rel_div) +done + +instance int :: ring_div +proof + fix a b c :: int + assume not0: "b \ 0" + show "(a + c * b) div b = c + a div b" + unfolding zdiv_zadd1_eq [of a "c * b"] using not0 + by (simp add: zmod_zmult1_eq zmod_zdiv_trivial zdiv_zmult1_eq) +next + fix a b c :: int + assume "a \ 0" + then show "(a * b) div (a * c) = b div c" + proof (cases "b \ 0 \ c \ 0") + case False then show ?thesis by auto + next + case True then have "b \ 0" and "c \ 0" by auto + with `a \ 0` + have "\q r. divmod_int_rel b c (q, r) \ divmod_int_rel (a * b) (a * c) (q, a * r)" + apply (auto simp add: divmod_int_rel_def) + apply (auto simp add: algebra_simps) + apply (auto simp add: zero_less_mult_iff zero_le_mult_iff mult_le_0_iff mult_commute [of a] mult_less_cancel_right) + done + moreover with `c \ 0` divmod_int_rel_div_mod have "divmod_int_rel b c (b div c, b mod c)" by auto + ultimately have "divmod_int_rel (a * b) (a * c) (b div c, a * (b mod c))" . + moreover from `a \ 0` `c \ 0` have "a * c \ 0" by simp + ultimately show ?thesis by (rule divmod_int_rel_div) + qed +qed auto + +lemma posDivAlg_div_mod: + assumes "k \ 0" + and "l \ 0" + shows "posDivAlg k l = (k div l, k mod l)" +proof (cases "l = 0") + case True then show ?thesis by (simp add: posDivAlg.simps) +next + case False with assms posDivAlg_correct + have "divmod_int_rel k l (fst (posDivAlg k l), snd (posDivAlg k l))" + by simp + from divmod_int_rel_div [OF this `l \ 0`] divmod_int_rel_mod [OF this `l \ 0`] + show ?thesis by simp +qed + +lemma negDivAlg_div_mod: + assumes "k < 0" + and "l > 0" + shows "negDivAlg k l = (k div l, k mod l)" +proof - + from assms have "l \ 0" by simp + from assms negDivAlg_correct + have "divmod_int_rel k l (fst (negDivAlg k l), snd (negDivAlg k l))" + by simp + from divmod_int_rel_div [OF this `l \ 0`] divmod_int_rel_mod [OF this `l \ 0`] + show ?thesis by simp +qed + +lemma zmod_eq_0_iff: "(m mod d = 0) = (EX q::int. m = d*q)" +by (simp add: dvd_eq_mod_eq_0 [symmetric] dvd_def) + +(* REVISIT: should this be generalized to all semiring_div types? *) +lemmas zmod_eq_0D [dest!] = zmod_eq_0_iff [THEN iffD1] + + +subsubsection{*Proving @{term "a div (b*c) = (a div b) div c"} *} + +(*The condition c>0 seems necessary. Consider that 7 div ~6 = ~2 but + 7 div 2 div ~3 = 3 div ~3 = ~1. The subcase (a div b) mod c = 0 seems + to cause particular problems.*) + +text{*first, four lemmas to bound the remainder for the cases b<0 and b>0 *} + +lemma zmult2_lemma_aux1: "[| (0::int) < c; b < r; r \ 0 |] ==> b*c < b*(q mod c) + r" +apply (subgoal_tac "b * (c - q mod c) < r * 1") + apply (simp add: algebra_simps) +apply (rule order_le_less_trans) + apply (erule_tac [2] mult_strict_right_mono) + apply (rule mult_left_mono_neg) + using add1_zle_eq[of "q mod c"]apply(simp add: algebra_simps pos_mod_bound) + apply (simp) +apply (simp) +done + +lemma zmult2_lemma_aux2: + "[| (0::int) < c; b < r; r \ 0 |] ==> b * (q mod c) + r \ 0" +apply (subgoal_tac "b * (q mod c) \ 0") + apply arith +apply (simp add: mult_le_0_iff) +done + +lemma zmult2_lemma_aux3: "[| (0::int) < c; 0 \ r; r < b |] ==> 0 \ b * (q mod c) + r" +apply (subgoal_tac "0 \ b * (q mod c) ") +apply arith +apply (simp add: zero_le_mult_iff) +done + +lemma zmult2_lemma_aux4: "[| (0::int) < c; 0 \ r; r < b |] ==> b * (q mod c) + r < b * c" +apply (subgoal_tac "r * 1 < b * (c - q mod c) ") + apply (simp add: right_diff_distrib) +apply (rule order_less_le_trans) + apply (erule mult_strict_right_mono) + apply (rule_tac [2] mult_left_mono) + apply simp + using add1_zle_eq[of "q mod c"] apply (simp add: algebra_simps pos_mod_bound) +apply simp +done + +lemma zmult2_lemma: "[| divmod_int_rel a b (q, r); b \ 0; 0 < c |] + ==> divmod_int_rel a (b * c) (q div c, b*(q mod c) + r)" +by (auto simp add: mult_ac divmod_int_rel_def linorder_neq_iff + zero_less_mult_iff right_distrib [symmetric] + zmult2_lemma_aux1 zmult2_lemma_aux2 zmult2_lemma_aux3 zmult2_lemma_aux4) + +lemma zdiv_zmult2_eq: "(0::int) < c ==> a div (b*c) = (a div b) div c" +apply (case_tac "b = 0", simp) +apply (force simp add: divmod_int_rel_div_mod [THEN zmult2_lemma, THEN divmod_int_rel_div]) +done + +lemma zmod_zmult2_eq: + "(0::int) < c ==> a mod (b*c) = b*(a div b mod c) + a mod b" +apply (case_tac "b = 0", simp) +apply (force simp add: divmod_int_rel_div_mod [THEN zmult2_lemma, THEN divmod_int_rel_mod]) +done + + +subsubsection {*Splitting Rules for div and mod*} + +text{*The proofs of the two lemmas below are essentially identical*} + +lemma split_pos_lemma: + "0 + P(n div k :: int)(n mod k) = (\i j. 0\j & j P i j)" +apply (rule iffI, clarify) + apply (erule_tac P="P ?x ?y" in rev_mp) + apply (subst mod_add_eq) + apply (subst zdiv_zadd1_eq) + apply (simp add: div_pos_pos_trivial mod_pos_pos_trivial) +txt{*converse direction*} +apply (drule_tac x = "n div k" in spec) +apply (drule_tac x = "n mod k" in spec, simp) +done + +lemma split_neg_lemma: + "k<0 ==> + P(n div k :: int)(n mod k) = (\i j. k0 & n = k*i + j --> P i j)" +apply (rule iffI, clarify) + apply (erule_tac P="P ?x ?y" in rev_mp) + apply (subst mod_add_eq) + apply (subst zdiv_zadd1_eq) + apply (simp add: div_neg_neg_trivial mod_neg_neg_trivial) +txt{*converse direction*} +apply (drule_tac x = "n div k" in spec) +apply (drule_tac x = "n mod k" in spec, simp) +done + +lemma split_zdiv: + "P(n div k :: int) = + ((k = 0 --> P 0) & + (0 (\i j. 0\j & j P i)) & + (k<0 --> (\i j. k0 & n = k*i + j --> P i)))" +apply (case_tac "k=0", simp) +apply (simp only: linorder_neq_iff) +apply (erule disjE) + apply (simp_all add: split_pos_lemma [of concl: "%x y. P x"] + split_neg_lemma [of concl: "%x y. P x"]) +done + +lemma split_zmod: + "P(n mod k :: int) = + ((k = 0 --> P n) & + (0 (\i j. 0\j & j P j)) & + (k<0 --> (\i j. k0 & n = k*i + j --> P j)))" +apply (case_tac "k=0", simp) +apply (simp only: linorder_neq_iff) +apply (erule disjE) + apply (simp_all add: split_pos_lemma [of concl: "%x y. P y"] + split_neg_lemma [of concl: "%x y. P y"]) +done + +(* Enable arith to deal with div 2 and mod 2: *) +declare split_zdiv [of _ _ "number_of k", simplified, standard, arith_split] +declare split_zmod [of _ _ "number_of k", simplified, standard, arith_split] + + +subsubsection{*Speeding up the Division Algorithm with Shifting*} + +text{*computing div by shifting *} + +lemma pos_zdiv_mult_2: "(0::int) \ a ==> (1 + 2*b) div (2*a) = b div a" +proof cases + assume "a=0" + thus ?thesis by simp +next + assume "a\0" and le_a: "0\a" + hence a_pos: "1 \ a" by arith + hence one_less_a2: "1 < 2 * a" by arith + hence le_2a: "2 * (1 + b mod a) \ 2 * a" + unfolding mult_le_cancel_left + by (simp add: add1_zle_eq add_commute [of 1]) + with a_pos have "0 \ b mod a" by simp + hence le_addm: "0 \ 1 mod (2*a) + 2*(b mod a)" + by (simp add: mod_pos_pos_trivial one_less_a2) + with le_2a + have "(1 mod (2*a) + 2*(b mod a)) div (2*a) = 0" + by (simp add: div_pos_pos_trivial le_addm mod_pos_pos_trivial one_less_a2 + right_distrib) + thus ?thesis + by (subst zdiv_zadd1_eq, + simp add: mod_mult_mult1 one_less_a2 + div_pos_pos_trivial) +qed + +lemma neg_zdiv_mult_2: "a \ (0::int) ==> (1 + 2*b) div (2*a) = (b+1) div a" +apply (subgoal_tac " (1 + 2* (-b - 1)) div (2 * (-a)) = (-b - 1) div (-a) ") +apply (rule_tac [2] pos_zdiv_mult_2) +apply (auto simp add: right_diff_distrib) +apply (subgoal_tac " (-1 - (2 * b)) = - (1 + (2 * b))") +apply (simp only: zdiv_zminus_zminus diff_minus minus_add_distrib [symmetric]) +apply (simp_all add: algebra_simps) +apply (simp only: ab_diff_minus minus_add_distrib [symmetric] number_of_Min zdiv_zminus_zminus) +done + +lemma zdiv_number_of_Bit0 [simp]: + "number_of (Int.Bit0 v) div number_of (Int.Bit0 w) = + number_of v div (number_of w :: int)" +by (simp only: number_of_eq numeral_simps) (simp add: mult_2 [symmetric]) + +lemma zdiv_number_of_Bit1 [simp]: + "number_of (Int.Bit1 v) div number_of (Int.Bit0 w) = + (if (0::int) \ number_of w + then number_of v div (number_of w) + else (number_of v + (1::int)) div (number_of w))" +apply (simp only: number_of_eq numeral_simps UNIV_I split: split_if) +apply (simp add: pos_zdiv_mult_2 neg_zdiv_mult_2 add_ac mult_2 [symmetric]) +done + + +subsubsection{*Computing mod by Shifting (proofs resemble those for div)*} + +lemma pos_zmod_mult_2: + fixes a b :: int + assumes "0 \ a" + shows "(1 + 2 * b) mod (2 * a) = 1 + 2 * (b mod a)" +proof (cases "0 < a") + case False with assms show ?thesis by simp +next + case True + then have "b mod a < a" by (rule pos_mod_bound) + then have "1 + b mod a \ a" by simp + then have A: "2 * (1 + b mod a) \ 2 * a" by simp + from `0 < a` have "0 \ b mod a" by (rule pos_mod_sign) + then have B: "0 \ 1 + 2 * (b mod a)" by simp + have "((1\int) mod ((2\int) * a) + (2\int) * b mod ((2\int) * a)) mod ((2\int) * a) = (1\int) + (2\int) * (b mod a)" + using `0 < a` and A + by (auto simp add: mod_mult_mult1 mod_pos_pos_trivial ring_distribs intro!: mod_pos_pos_trivial B) + then show ?thesis by (subst mod_add_eq) +qed + +lemma neg_zmod_mult_2: + fixes a b :: int + assumes "a \ 0" + shows "(1 + 2 * b) mod (2 * a) = 2 * ((b + 1) mod a) - 1" +proof - + from assms have "0 \ - a" by auto + then have "(1 + 2 * (- b - 1)) mod (2 * (- a)) = 1 + 2 * ((- b - 1) mod (- a))" + by (rule pos_zmod_mult_2) + then show ?thesis by (simp add: zmod_zminus2 algebra_simps) + (simp add: diff_minus add_ac) +qed + +lemma zmod_number_of_Bit0 [simp]: + "number_of (Int.Bit0 v) mod number_of (Int.Bit0 w) = + (2::int) * (number_of v mod number_of w)" +apply (simp only: number_of_eq numeral_simps) +apply (simp add: mod_mult_mult1 pos_zmod_mult_2 + neg_zmod_mult_2 add_ac mult_2 [symmetric]) +done + +lemma zmod_number_of_Bit1 [simp]: + "number_of (Int.Bit1 v) mod number_of (Int.Bit0 w) = + (if (0::int) \ number_of w + then 2 * (number_of v mod number_of w) + 1 + else 2 * ((number_of v + (1::int)) mod number_of w) - 1)" +apply (simp only: number_of_eq numeral_simps) +apply (simp add: mod_mult_mult1 pos_zmod_mult_2 + neg_zmod_mult_2 add_ac mult_2 [symmetric]) +done + + +subsubsection{*Quotients of Signs*} + +lemma div_neg_pos_less0: "[| a < (0::int); 0 < b |] ==> a div b < 0" +apply (subgoal_tac "a div b \ -1", force) +apply (rule order_trans) +apply (rule_tac a' = "-1" in zdiv_mono1) +apply (auto simp add: div_eq_minus1) +done + +lemma div_nonneg_neg_le0: "[| (0::int) \ a; b < 0 |] ==> a div b \ 0" +by (drule zdiv_mono1_neg, auto) + +lemma div_nonpos_pos_le0: "[| (a::int) \ 0; b > 0 |] ==> a div b \ 0" +by (drule zdiv_mono1, auto) + +lemma pos_imp_zdiv_nonneg_iff: "(0::int) < b ==> (0 \ a div b) = (0 \ a)" +apply auto +apply (drule_tac [2] zdiv_mono1) +apply (auto simp add: linorder_neq_iff) +apply (simp (no_asm_use) add: linorder_not_less [symmetric]) +apply (blast intro: div_neg_pos_less0) +done + +lemma neg_imp_zdiv_nonneg_iff: + "b < (0::int) ==> (0 \ a div b) = (a \ (0::int))" +apply (subst zdiv_zminus_zminus [symmetric]) +apply (subst pos_imp_zdiv_nonneg_iff, auto) +done + +(*But not (a div b \ 0 iff a\0); consider a=1, b=2 when a div b = 0.*) +lemma pos_imp_zdiv_neg_iff: "(0::int) < b ==> (a div b < 0) = (a < 0)" +by (simp add: linorder_not_le [symmetric] pos_imp_zdiv_nonneg_iff) + +(*Again the law fails for \: consider a = -1, b = -2 when a div b = 0*) +lemma neg_imp_zdiv_neg_iff: "b < (0::int) ==> (a div b < 0) = (0 < a)" +by (simp add: linorder_not_le [symmetric] neg_imp_zdiv_nonneg_iff) + + +subsubsection {* The Divides Relation *} + +lemmas zdvd_iff_zmod_eq_0_number_of [simp] = + dvd_eq_mod_eq_0 [of "number_of x::int" "number_of y::int", standard] + +lemma zdvd_zmod: "f dvd m ==> f dvd (n::int) ==> f dvd m mod n" + by (rule dvd_mod) (* TODO: remove *) + +lemma zdvd_zmod_imp_zdvd: "k dvd m mod n ==> k dvd n ==> k dvd (m::int)" + by (rule dvd_mod_imp_dvd) (* TODO: remove *) + +lemma zmult_div_cancel: "(n::int) * (m div n) = m - (m mod n)" + using zmod_zdiv_equality[where a="m" and b="n"] + by (simp add: algebra_simps) + +lemma zpower_zmod: "((x::int) mod m)^y mod m = x^y mod m" +apply (induct "y", auto) +apply (rule zmod_zmult1_eq [THEN trans]) +apply (simp (no_asm_simp)) +apply (rule mod_mult_eq [symmetric]) +done + +lemma zdiv_int: "int (a div b) = (int a) div (int b)" +apply (subst split_div, auto) +apply (subst split_zdiv, auto) +apply (rule_tac a="int (b * i) + int j" and b="int b" and r="int j" and r'=ja in unique_quotient) +apply (auto simp add: divmod_int_rel_def of_nat_mult) +done + +lemma zmod_int: "int (a mod b) = (int a) mod (int b)" +apply (subst split_mod, auto) +apply (subst split_zmod, auto) +apply (rule_tac a="int (b * i) + int j" and b="int b" and q="int i" and q'=ia + in unique_remainder) +apply (auto simp add: divmod_int_rel_def of_nat_mult) +done + +lemma abs_div: "(y::int) dvd x \ abs (x div y) = abs x div abs y" +by (unfold dvd_def, cases "y=0", auto simp add: abs_mult) + +lemma zdvd_mult_div_cancel:"(n::int) dvd m \ n * (m div n) = m" +apply (subgoal_tac "m mod n = 0") + apply (simp add: zmult_div_cancel) +apply (simp only: dvd_eq_mod_eq_0) +done + +text{*Suggested by Matthias Daum*} +lemma int_power_div_base: + "\0 < m; 0 < k\ \ k ^ m div k = (k::int) ^ (m - Suc 0)" +apply (subgoal_tac "k ^ m = k ^ ((m - Suc 0) + Suc 0)") + apply (erule ssubst) + apply (simp only: power_add) + apply simp_all +done + +text {* by Brian Huffman *} +lemma zminus_zmod: "- ((x::int) mod m) mod m = - x mod m" +by (rule mod_minus_eq [symmetric]) + +lemma zdiff_zmod_left: "(x mod m - y) mod m = (x - y) mod (m::int)" +by (rule mod_diff_left_eq [symmetric]) + +lemma zdiff_zmod_right: "(x - y mod m) mod m = (x - y) mod (m::int)" +by (rule mod_diff_right_eq [symmetric]) + +lemmas zmod_simps = + mod_add_left_eq [symmetric] + mod_add_right_eq [symmetric] + zmod_zmult1_eq [symmetric] + mod_mult_left_eq [symmetric] + zpower_zmod + zminus_zmod zdiff_zmod_left zdiff_zmod_right + +text {* Distributive laws for function @{text nat}. *} + +lemma nat_div_distrib: "0 \ x \ nat (x div y) = nat x div nat y" +apply (rule linorder_cases [of y 0]) +apply (simp add: div_nonneg_neg_le0) +apply simp +apply (simp add: nat_eq_iff pos_imp_zdiv_nonneg_iff zdiv_int) +done + +(*Fails if y<0: the LHS collapses to (nat z) but the RHS doesn't*) +lemma nat_mod_distrib: + "\0 \ x; 0 \ y\ \ nat (x mod y) = nat x mod nat y" +apply (case_tac "y = 0", simp) +apply (simp add: nat_eq_iff zmod_int) +done + +text {* transfer setup *} + +lemma transfer_nat_int_functions: + "(x::int) >= 0 \ y >= 0 \ (nat x) div (nat y) = nat (x div y)" + "(x::int) >= 0 \ y >= 0 \ (nat x) mod (nat y) = nat (x mod y)" + by (auto simp add: nat_div_distrib nat_mod_distrib) + +lemma transfer_nat_int_function_closures: + "(x::int) >= 0 \ y >= 0 \ x div y >= 0" + "(x::int) >= 0 \ y >= 0 \ x mod y >= 0" + apply (cases "y = 0") + apply (auto simp add: pos_imp_zdiv_nonneg_iff) + apply (cases "y = 0") + apply auto +done + +declare TransferMorphism_nat_int [transfer add return: + transfer_nat_int_functions + transfer_nat_int_function_closures +] + +lemma transfer_int_nat_functions: + "(int x) div (int y) = int (x div y)" + "(int x) mod (int y) = int (x mod y)" + by (auto simp add: zdiv_int zmod_int) + +lemma transfer_int_nat_function_closures: + "is_nat x \ is_nat y \ is_nat (x div y)" + "is_nat x \ is_nat y \ is_nat (x mod y)" + by (simp_all only: is_nat_def transfer_nat_int_function_closures) + +declare TransferMorphism_int_nat [transfer add return: + transfer_int_nat_functions + transfer_int_nat_function_closures +] + +text{*Suggested by Matthias Daum*} +lemma int_div_less_self: "\0 < x; 1 < k\ \ x div k < (x::int)" +apply (subgoal_tac "nat x div nat k < nat x") + apply (simp (asm_lr) add: nat_div_distrib [symmetric]) +apply (rule Divides.div_less_dividend, simp_all) +done + +text {* code generator setup *} + +context ring_1 +begin + +lemma of_int_num [code]: + "of_int k = (if k = 0 then 0 else if k < 0 then + - of_int (- k) else let + (l, m) = divmod_int k 2; + l' = of_int l + in if m = 0 then l' + l' else l' + l' + 1)" +proof - + have aux1: "k mod (2\int) \ (0\int) \ + of_int k = of_int (k div 2 * 2 + 1)" + proof - + have "k mod 2 < 2" by (auto intro: pos_mod_bound) + moreover have "0 \ k mod 2" by (auto intro: pos_mod_sign) + moreover assume "k mod 2 \ 0" + ultimately have "k mod 2 = 1" by arith + moreover have "of_int k = of_int (k div 2 * 2 + k mod 2)" by simp + ultimately show ?thesis by auto + qed + have aux2: "\x. of_int 2 * x = x + x" + proof - + fix x + have int2: "(2::int) = 1 + 1" by arith + show "of_int 2 * x = x + x" + unfolding int2 of_int_add left_distrib by simp + qed + have aux3: "\x. x * of_int 2 = x + x" + proof - + fix x + have int2: "(2::int) = 1 + 1" by arith + show "x * of_int 2 = x + x" + unfolding int2 of_int_add right_distrib by simp + qed + from aux1 show ?thesis by (auto simp add: divmod_int_mod_div Let_def aux2 aux3) +qed + +end + +lemma zmod_eq_dvd_iff: "(x::int) mod n = y mod n \ n dvd x - y" +proof + assume H: "x mod n = y mod n" + hence "x mod n - y mod n = 0" by simp + hence "(x mod n - y mod n) mod n = 0" by simp + hence "(x - y) mod n = 0" by (simp add: mod_diff_eq[symmetric]) + thus "n dvd x - y" by (simp add: dvd_eq_mod_eq_0) +next + assume H: "n dvd x - y" + then obtain k where k: "x-y = n*k" unfolding dvd_def by blast + hence "x = n*k + y" by simp + hence "x mod n = (n*k + y) mod n" by simp + thus "x mod n = y mod n" by (simp add: mod_add_left_eq) +qed + +lemma nat_mod_eq_lemma: assumes xyn: "(x::nat) mod n = y mod n" and xy:"y \ x" + shows "\q. x = y + n * q" +proof- + from xy have th: "int x - int y = int (x - y)" by simp + from xyn have "int x mod int n = int y mod int n" + by (simp add: zmod_int[symmetric]) + hence "int n dvd int x - int y" by (simp only: zmod_eq_dvd_iff[symmetric]) + hence "n dvd x - y" by (simp add: th zdvd_int) + then show ?thesis using xy unfolding dvd_def apply clarsimp apply (rule_tac x="k" in exI) by arith +qed + +lemma nat_mod_eq_iff: "(x::nat) mod n = y mod n \ (\q1 q2. x + n * q1 = y + n * q2)" + (is "?lhs = ?rhs") +proof + assume H: "x mod n = y mod n" + {assume xy: "x \ y" + from H have th: "y mod n = x mod n" by simp + from nat_mod_eq_lemma[OF th xy] have ?rhs + apply clarify apply (rule_tac x="q" in exI) by (rule exI[where x="0"], simp)} + moreover + {assume xy: "y \ x" + from nat_mod_eq_lemma[OF H xy] have ?rhs + apply clarify apply (rule_tac x="0" in exI) by (rule_tac x="q" in exI, simp)} + ultimately show ?rhs using linear[of x y] by blast +next + assume ?rhs then obtain q1 q2 where q12: "x + n * q1 = y + n * q2" by blast + hence "(x + n * q1) mod n = (y + n * q2) mod n" by simp + thus ?lhs by simp +qed + +lemma div_nat_number_of [simp]: + "(number_of v :: nat) div number_of v' = + (if neg (number_of v :: int) then 0 + else nat (number_of v div number_of v'))" + unfolding nat_number_of_def number_of_is_id neg_def + by (simp add: nat_div_distrib) + +lemma one_div_nat_number_of [simp]: + "Suc 0 div number_of v' = nat (1 div number_of v')" +by (simp del: nat_numeral_1_eq_1 add: numeral_1_eq_Suc_0 [symmetric]) + +lemma mod_nat_number_of [simp]: + "(number_of v :: nat) mod number_of v' = + (if neg (number_of v :: int) then 0 + else if neg (number_of v' :: int) then number_of v + else nat (number_of v mod number_of v'))" + unfolding nat_number_of_def number_of_is_id neg_def + by (simp add: nat_mod_distrib) + +lemma one_mod_nat_number_of [simp]: + "Suc 0 mod number_of v' = + (if neg (number_of v' :: int) then Suc 0 + else nat (1 mod number_of v'))" +by (simp del: nat_numeral_1_eq_1 add: numeral_1_eq_Suc_0 [symmetric]) + +lemmas dvd_eq_mod_eq_0_number_of = + dvd_eq_mod_eq_0 [of "number_of x" "number_of y", standard] + +declare dvd_eq_mod_eq_0_number_of [simp] + + +subsubsection {* Code generation *} + +definition pdivmod :: "int \ int \ int \ int" where + "pdivmod k l = (\k\ div \l\, \k\ mod \l\)" + +lemma pdivmod_posDivAlg [code]: + "pdivmod k l = (if l = 0 then (0, \k\) else posDivAlg \k\ \l\)" +by (subst posDivAlg_div_mod) (simp_all add: pdivmod_def) + +lemma divmod_int_pdivmod: "divmod_int k l = (if k = 0 then (0, 0) else if l = 0 then (0, k) else + apsnd ((op *) (sgn l)) (if 0 < l \ 0 \ k \ l < 0 \ k < 0 + then pdivmod k l + else (let (r, s) = pdivmod k l in + if s = 0 then (- r, 0) else (- r - 1, \l\ - s))))" +proof - + have aux: "\q::int. - k = l * q \ k = l * - q" by auto + show ?thesis + by (simp add: divmod_int_mod_div pdivmod_def) + (auto simp add: aux not_less not_le zdiv_zminus1_eq_if + zmod_zminus1_eq_if zdiv_zminus2_eq_if zmod_zminus2_eq_if) +qed + +lemma divmod_int_code [code]: "divmod_int k l = (if k = 0 then (0, 0) else if l = 0 then (0, k) else + apsnd ((op *) (sgn l)) (if sgn k = sgn l + then pdivmod k l + else (let (r, s) = pdivmod k l in + if s = 0 then (- r, 0) else (- r - 1, \l\ - s))))" +proof - + have "k \ 0 \ l \ 0 \ 0 < l \ 0 \ k \ l < 0 \ k < 0 \ sgn k = sgn l" + by (auto simp add: not_less sgn_if) + then show ?thesis by (simp add: divmod_int_pdivmod) +qed + +code_modulename SML + Divides Arith + +code_modulename OCaml + Divides Arith + +code_modulename Haskell + Divides Arith + +end diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/Finite_Set.thy --- a/src/HOL/Finite_Set.thy Thu Oct 29 16:23:57 2009 +0100 +++ b/src/HOL/Finite_Set.thy Wed Nov 04 11:40:59 2009 +0100 @@ -3261,15 +3261,15 @@ lemma finite_linorder_max_induct[consumes 1, case_names empty insert]: "finite A \ P {} \ - (!!A b. finite A \ ALL a:A. a < b \ P A \ P(insert b A)) + (!!b A. finite A \ ALL a:A. a < b \ P A \ P(insert b A)) \ P A" proof (induct rule: finite_psubset_induct) fix A :: "'a set" assume IH: "!! B. finite B \ B < A \ P {} \ - (!!A b. finite A \ (\a\A. a P A \ P (insert b A)) + (!!b A. finite A \ (\a\A. a P A \ P (insert b A)) \ P B" and "finite A" and "P {}" - and step: "!!A b. \finite A; \a\A. a < b; P A\ \ P (insert b A)" + and step: "!!b A. \finite A; \a\A. a < b; P A\ \ P (insert b A)" show "P A" proof (cases "A = {}") assume "A = {}" thus "P A" using `P {}` by simp @@ -3279,14 +3279,14 @@ with `finite A` have "Max A : A" by auto hence A: "?A = A" using insert_Diff_single insert_absorb by auto moreover have "finite ?B" using `finite A` by simp - ultimately have "P ?B" using `P {}` step IH by blast + ultimately have "P ?B" using `P {}` step IH[of ?B] by blast moreover have "\a\?B. a < Max A" using Max_ge [OF `finite A`] by fastsimp ultimately show "P A" using A insert_Diff_single step[OF `finite ?B`] by fastsimp qed qed lemma finite_linorder_min_induct[consumes 1, case_names empty insert]: - "\finite A; P {}; \A b. \finite A; \a\A. b < a; P A\ \ P (insert b A)\ \ P A" + "\finite A; P {}; \b A. \finite A; \a\A. b < a; P A\ \ P (insert b A)\ \ P A" by(rule linorder.finite_linorder_max_induct[OF dual_linorder]) end diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/FunDef.thy --- a/src/HOL/FunDef.thy Thu Oct 29 16:23:57 2009 +0100 +++ b/src/HOL/FunDef.thy Wed Nov 04 11:40:59 2009 +0100 @@ -11,7 +11,6 @@ "Tools/sat_solver.ML" ("Tools/Function/function_lib.ML") ("Tools/Function/function_common.ML") - ("Tools/Function/inductive_wrap.ML") ("Tools/Function/context_tree.ML") ("Tools/Function/function_core.ML") ("Tools/Function/sum_tree.ML") @@ -106,7 +105,6 @@ use "Tools/Function/function_lib.ML" use "Tools/Function/function_common.ML" -use "Tools/Function/inductive_wrap.ML" use "Tools/Function/context_tree.ML" use "Tools/Function/function_core.ML" use "Tools/Function/sum_tree.ML" diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/Groebner_Basis.thy --- a/src/HOL/Groebner_Basis.thy Thu Oct 29 16:23:57 2009 +0100 +++ b/src/HOL/Groebner_Basis.thy Wed Nov 04 11:40:59 2009 +0100 @@ -5,7 +5,7 @@ header {* Semiring normalization and Groebner Bases *} theory Groebner_Basis -imports IntDiv +imports Numeral_Simprocs uses "Tools/Groebner_Basis/misc.ML" "Tools/Groebner_Basis/normalizer_data.ML" diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/HOL.thy --- a/src/HOL/HOL.thy Thu Oct 29 16:23:57 2009 +0100 +++ b/src/HOL/HOL.thy Wed Nov 04 11:40:59 2009 +0100 @@ -24,6 +24,7 @@ "~~/src/Tools/coherent.ML" "~~/src/Tools/eqsubst.ML" "~~/src/Provers/quantifier1.ML" + "Tools/res_blacklist.ML" ("Tools/simpdata.ML") "~~/src/Tools/random_word.ML" "~~/src/Tools/atomize_elim.ML" @@ -35,6 +36,8 @@ setup {* Intuitionistic.method_setup @{binding iprover} *} +setup Res_Blacklist.setup + subsection {* Primitive logic *} @@ -833,19 +836,14 @@ val hyp_subst_tacs = [Hypsubst.hyp_subst_tac] end); -structure BasicClassical: BASIC_CLASSICAL = Classical; -open BasicClassical; +structure Basic_Classical: BASIC_CLASSICAL = Classical; +open Basic_Classical; ML_Antiquote.value "claset" (Scan.succeed "Classical.claset_of (ML_Context.the_local_context ())"); - -structure ResBlacklist = Named_Thms - (val name = "noatp" val description = "theorems blacklisted for ATP"); *} -text {*ResBlacklist holds theorems blacklisted to sledgehammer. - These theorems typically produce clauses that are prolific (match too many equality or - membership literals) and relate to seldom-used facts. Some duplicate other rules.*} +setup Classical.setup setup {* let @@ -855,9 +853,7 @@ (nth (Thm.prems_of thm) (i - 1)) then Hypsubst.hyp_subst_tac i thm else no_tac thm; in Hypsubst.hypsubst_setup - #> ContextRules.addSWrapper (fn tac => hyp_subst_tac' ORELSE' tac) - #> Classical.setup - #> ResBlacklist.setup + #> Context_Rules.addSWrapper (fn tac => hyp_subst_tac' ORELSE' tac) end *} @@ -1822,7 +1818,7 @@ code_datatype "TYPE('a\{})" -code_datatype Trueprop "prop" +code_datatype "prop" Trueprop text {* Code equations *} diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/Import/import_syntax.ML --- a/src/HOL/Import/import_syntax.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/HOL/Import/import_syntax.ML Wed Nov 04 11:40:59 2009 +0100 @@ -61,7 +61,7 @@ val thyname = get_generating_thy thy val segname = get_import_segment thy val (int_thms,facts) = Replay.setup_int_thms thyname thy - val thmnames = List.filter (not o should_ignore thyname thy) facts + val thmnames = filter_out (should_ignore thyname thy) facts fun replay thy = let val _ = ImportRecorder.load_history thyname diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/Import/lazy_seq.ML --- a/src/HOL/Import/lazy_seq.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/HOL/Import/lazy_seq.ML Wed Nov 04 11:40:59 2009 +0100 @@ -407,8 +407,8 @@ make (fn () => copy (f x)) end -fun EVERY fs = List.foldr (op THEN) succeed fs -fun FIRST fs = List.foldr (op ORELSE) fail fs +fun EVERY fs = fold_rev (curry op THEN) fs succeed +fun FIRST fs = fold_rev (curry op ORELSE) fs fail fun TRY f x = make (fn () => diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/Import/proof_kernel.ML --- a/src/HOL/Import/proof_kernel.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/HOL/Import/proof_kernel.ML Wed Nov 04 11:40:59 2009 +0100 @@ -776,7 +776,7 @@ val (c,asl) = case terms of [] => raise ERR "x2p" "Bad oracle description" | (hd::tl) => (hd,tl) - val tg = List.foldr (fn (oracle,tg) => Tag.merge (Tag.read oracle) tg) Tag.empty_tag ors + val tg = fold_rev (Tag.merge o Tag.read) ors Tag.empty_tag in mk_proof (POracle(tg,map xml_to_term asl,xml_to_term c)) end @@ -1160,7 +1160,7 @@ | replace_name n' _ = ERR "replace_name" (* map: old or fresh name -> old free, invmap: old free which has fresh name assigned to it -> fresh name *) - fun dis (v, mapping as (map,invmap)) = + fun dis v (mapping as (map,invmap)) = let val n = name_of v in case Symtab.lookup map n of NONE => (Symtab.update (n, v) map, invmap) @@ -1179,11 +1179,11 @@ else let val (_, invmap) = - List.foldl dis (Symtab.empty, Termtab.empty) frees - fun make_subst ((oldfree, newname), (intros, elims)) = + fold dis frees (Symtab.empty, Termtab.empty) + fun make_subst (oldfree, newname) (intros, elims) = (cterm_of thy oldfree :: intros, cterm_of thy (replace_name newname oldfree) :: elims) - val (intros, elims) = List.foldl make_subst ([], []) (Termtab.dest invmap) + val (intros, elims) = fold make_subst (Termtab.dest invmap) ([], []) in forall_elim_list elims (forall_intr_list intros thm) end @@ -1837,7 +1837,7 @@ | inst_type ty1 ty2 (ty as Type(name,tys)) = Type(name,map (inst_type ty1 ty2) tys) in - List.foldr (fn (v,th) => + fold_rev (fn v => fn th => let val cdom = fst (dom_rng (fst (dom_rng cty))) val vty = type_of v @@ -1845,11 +1845,11 @@ val cc = cterm_of thy (Const(cname,newcty)) in mk_COMB (mk_REFL cc) (mk_ABS v th thy) thy - end) th vlist' + end) vlist' th end | SOME _ => raise ERR "GEN_ABS" "Bad constant" | NONE => - List.foldr (fn (v,th) => mk_ABS v th thy) th vlist' + fold_rev (fn v => fn th => mk_ABS v th thy) vlist' th val res = HOLThm(rens_of info',th1) val _ = message "RESULT:" val _ = if_debug pth res @@ -2020,8 +2020,8 @@ Sign.add_consts_i (map (fn (c, T, mx) => (Binding.name c, T, mx)) consts) thy' end - val thy1 = List.foldr (fn(name,thy)=> - snd (get_defname thyname name thy)) thy1 names + val thy1 = fold_rev (fn name => fn thy => + snd (get_defname thyname name thy)) names thy1 fun new_name name = fst (get_defname thyname name thy1) val names' = map (fn name => (new_name name,name,false)) names val (thy',res) = Choice_Specification.add_specification NONE @@ -2041,12 +2041,12 @@ then quotename name else (quotename newname) ^ ": " ^ (quotename name),thy') end - val (new_names,thy') = List.foldr (fn(name,(names,thy)) => + val (new_names,thy') = fold_rev (fn name => fn (names, thy) => let val (name',thy') = handle_const (name,thy) in (name'::names,thy') - end) ([],thy') names + end) names ([], thy') val thy'' = add_dump ("specification (" ^ (spaces new_names) ^ ") " ^ thmname ^ ": " ^ (smart_string_of_thm th) ^ "\n by (import " ^ thyname ^ " " ^ thmname ^ ")") thy' diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/Import/shuffler.ML --- a/src/HOL/Import/shuffler.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/HOL/Import/shuffler.ML Wed Nov 04 11:40:59 2009 +0100 @@ -628,7 +628,7 @@ val all_thms = map (`Thm.get_name_hint) (maps #2 (Facts.dest_static [] (PureThy.facts_of thy))) in - List.filter (match_consts ignored t) all_thms + filter (match_consts ignored t) all_thms end fun gen_shuffle_tac ctxt search thms i st = diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/Int.thy --- a/src/HOL/Int.thy Thu Oct 29 16:23:57 2009 +0100 +++ b/src/HOL/Int.thy Wed Nov 04 11:40:59 2009 +0100 @@ -2105,6 +2105,14 @@ lemma nat_dvd_iff: "(nat z dvd m) = (if 0 \ z then (z dvd int m) else m = 0)" by (auto simp add: dvd_int_iff) +lemma eq_nat_nat_iff: + "0 \ z \ 0 \ z' \ nat z = nat z' \ z = z'" + by (auto elim!: nonneg_eq_int) + +lemma nat_power_eq: + "0 \ z \ nat (z ^ n) = nat z ^ n" + by (induct n) (simp_all add: nat_mult_distrib) + lemma zdvd_imp_le: "[| z dvd n; 0 < n |] ==> z \ (n::int)" apply (rule_tac z=n in int_cases) apply (auto simp add: dvd_int_iff) @@ -2238,13 +2246,13 @@ by simp code_modulename SML - Int Integer + Int Arith code_modulename OCaml - Int Integer + Int Arith code_modulename Haskell - Int Integer + Int Arith types_code "int" ("int") diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/IntDiv.thy --- a/src/HOL/IntDiv.thy Thu Oct 29 16:23:57 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,1314 +0,0 @@ -(* Title: HOL/IntDiv.thy - Author: Lawrence C Paulson, Cambridge University Computer Laboratory - Copyright 1999 University of Cambridge -*) - -header{* The Division Operators div and mod *} - -theory IntDiv -imports Int Divides FunDef -begin - -definition divmod_rel :: "int \ int \ int \ int \ bool" where - --{*definition of quotient and remainder*} - [code]: "divmod_rel a b = (\(q, r). a = b * q + r \ - (if 0 < b then 0 \ r \ r < b else b < r \ r \ 0))" - -definition adjust :: "int \ int \ int \ int \ int" where - --{*for the division algorithm*} - [code]: "adjust b = (\(q, r). if 0 \ r - b then (2 * q + 1, r - b) - else (2 * q, r))" - -text{*algorithm for the case @{text "a\0, b>0"}*} -function posDivAlg :: "int \ int \ int \ int" where - "posDivAlg a b = (if a < b \ b \ 0 then (0, a) - else adjust b (posDivAlg a (2 * b)))" -by auto -termination by (relation "measure (\(a, b). nat (a - b + 1))") auto - -text{*algorithm for the case @{text "a<0, b>0"}*} -function negDivAlg :: "int \ int \ int \ int" where - "negDivAlg a b = (if 0 \a + b \ b \ 0 then (-1, a + b) - else adjust b (negDivAlg a (2 * b)))" -by auto -termination by (relation "measure (\(a, b). nat (- a - b))") auto - -text{*algorithm for the general case @{term "b\0"}*} -definition negateSnd :: "int \ int \ int \ int" where - [code_unfold]: "negateSnd = apsnd uminus" - -definition divmod :: "int \ int \ int \ int" where - --{*The full division algorithm considers all possible signs for a, b - including the special case @{text "a=0, b<0"} because - @{term negDivAlg} requires @{term "a<0"}.*} - "divmod a b = (if 0 \ a then if 0 \ b then posDivAlg a b - else if a = 0 then (0, 0) - else negateSnd (negDivAlg (-a) (-b)) - else - if 0 < b then negDivAlg a b - else negateSnd (posDivAlg (-a) (-b)))" - -instantiation int :: Divides.div -begin - -definition - div_def: "a div b = fst (divmod a b)" - -definition - mod_def: "a mod b = snd (divmod a b)" - -instance .. - -end - -lemma divmod_mod_div: - "divmod p q = (p div q, p mod q)" - by (auto simp add: div_def mod_def) - -text{* -Here is the division algorithm in ML: - -\begin{verbatim} - fun posDivAlg (a,b) = - if ar-b then (2*q+1, r-b) else (2*q, r) - end - - fun negDivAlg (a,b) = - if 0\a+b then (~1,a+b) - else let val (q,r) = negDivAlg(a, 2*b) - in if 0\r-b then (2*q+1, r-b) else (2*q, r) - end; - - fun negateSnd (q,r:int) = (q,~r); - - fun divmod (a,b) = if 0\a then - if b>0 then posDivAlg (a,b) - else if a=0 then (0,0) - else negateSnd (negDivAlg (~a,~b)) - else - if 0 b*q + r; 0 \ r'; r' < b; r < b |] - ==> q' \ (q::int)" -apply (subgoal_tac "r' + b * (q'-q) \ r") - prefer 2 apply (simp add: right_diff_distrib) -apply (subgoal_tac "0 < b * (1 + q - q') ") -apply (erule_tac [2] order_le_less_trans) - prefer 2 apply (simp add: right_diff_distrib right_distrib) -apply (subgoal_tac "b * q' < b * (1 + q) ") - prefer 2 apply (simp add: right_diff_distrib right_distrib) -apply (simp add: mult_less_cancel_left) -done - -lemma unique_quotient_lemma_neg: - "[| b*q' + r' \ b*q + r; r \ 0; b < r; b < r' |] - ==> q \ (q'::int)" -by (rule_tac b = "-b" and r = "-r'" and r' = "-r" in unique_quotient_lemma, - auto) - -lemma unique_quotient: - "[| divmod_rel a b (q, r); divmod_rel a b (q', r'); b \ 0 |] - ==> q = q'" -apply (simp add: divmod_rel_def linorder_neq_iff split: split_if_asm) -apply (blast intro: order_antisym - dest: order_eq_refl [THEN unique_quotient_lemma] - order_eq_refl [THEN unique_quotient_lemma_neg] sym)+ -done - - -lemma unique_remainder: - "[| divmod_rel a b (q, r); divmod_rel a b (q', r'); b \ 0 |] - ==> r = r'" -apply (subgoal_tac "q = q'") - apply (simp add: divmod_rel_def) -apply (blast intro: unique_quotient) -done - - -subsection{*Correctness of @{term posDivAlg}, the Algorithm for Non-Negative Dividends*} - -text{*And positive divisors*} - -lemma adjust_eq [simp]: - "adjust b (q,r) = - (let diff = r-b in - if 0 \ diff then (2*q + 1, diff) - else (2*q, r))" -by (simp add: Let_def adjust_def) - -declare posDivAlg.simps [simp del] - -text{*use with a simproc to avoid repeatedly proving the premise*} -lemma posDivAlg_eqn: - "0 < b ==> - posDivAlg a b = (if a a" and "0 < b" - shows "divmod_rel a b (posDivAlg a b)" -using prems apply (induct a b rule: posDivAlg.induct) -apply auto -apply (simp add: divmod_rel_def) -apply (subst posDivAlg_eqn, simp add: right_distrib) -apply (case_tac "a < b") -apply simp_all -apply (erule splitE) -apply (auto simp add: right_distrib Let_def) -done - - -subsection{*Correctness of @{term negDivAlg}, the Algorithm for Negative Dividends*} - -text{*And positive divisors*} - -declare negDivAlg.simps [simp del] - -text{*use with a simproc to avoid repeatedly proving the premise*} -lemma negDivAlg_eqn: - "0 < b ==> - negDivAlg a b = - (if 0\a+b then (-1,a+b) else adjust b (negDivAlg a (2*b)))" -by (rule negDivAlg.simps [THEN trans], simp) - -(*Correctness of negDivAlg: it computes quotients correctly - It doesn't work if a=0 because the 0/b equals 0, not -1*) -lemma negDivAlg_correct: - assumes "a < 0" and "b > 0" - shows "divmod_rel a b (negDivAlg a b)" -using prems apply (induct a b rule: negDivAlg.induct) -apply (auto simp add: linorder_not_le) -apply (simp add: divmod_rel_def) -apply (subst negDivAlg_eqn, assumption) -apply (case_tac "a + b < (0\int)") -apply simp_all -apply (erule splitE) -apply (auto simp add: right_distrib Let_def) -done - - -subsection{*Existence Shown by Proving the Division Algorithm to be Correct*} - -(*the case a=0*) -lemma divmod_rel_0: "b \ 0 ==> divmod_rel 0 b (0, 0)" -by (auto simp add: divmod_rel_def linorder_neq_iff) - -lemma posDivAlg_0 [simp]: "posDivAlg 0 b = (0, 0)" -by (subst posDivAlg.simps, auto) - -lemma negDivAlg_minus1 [simp]: "negDivAlg -1 b = (-1, b - 1)" -by (subst negDivAlg.simps, auto) - -lemma negateSnd_eq [simp]: "negateSnd(q,r) = (q,-r)" -by (simp add: negateSnd_def) - -lemma divmod_rel_neg: "divmod_rel (-a) (-b) qr ==> divmod_rel a b (negateSnd qr)" -by (auto simp add: split_ifs divmod_rel_def) - -lemma divmod_correct: "b \ 0 ==> divmod_rel a b (divmod a b)" -by (force simp add: linorder_neq_iff divmod_rel_0 divmod_def divmod_rel_neg - posDivAlg_correct negDivAlg_correct) - -text{*Arbitrary definitions for division by zero. Useful to simplify - certain equations.*} - -lemma DIVISION_BY_ZERO [simp]: "a div (0::int) = 0 & a mod (0::int) = a" -by (simp add: div_def mod_def divmod_def posDivAlg.simps) - - -text{*Basic laws about division and remainder*} - -lemma zmod_zdiv_equality: "(a::int) = b * (a div b) + (a mod b)" -apply (case_tac "b = 0", simp) -apply (cut_tac a = a and b = b in divmod_correct) -apply (auto simp add: divmod_rel_def div_def mod_def) -done - -lemma zdiv_zmod_equality: "(b * (a div b) + (a mod b)) + k = (a::int)+k" -by(simp add: zmod_zdiv_equality[symmetric]) - -lemma zdiv_zmod_equality2: "((a div b) * b + (a mod b)) + k = (a::int)+k" -by(simp add: mult_commute zmod_zdiv_equality[symmetric]) - -text {* Tool setup *} - -ML {* -local - -structure CancelDivMod = CancelDivModFun(struct - - val div_name = @{const_name div}; - val mod_name = @{const_name mod}; - val mk_binop = HOLogic.mk_binop; - val mk_sum = Numeral_Simprocs.mk_sum HOLogic.intT; - val dest_sum = Numeral_Simprocs.dest_sum; - - val div_mod_eqs = map mk_meta_eq [@{thm zdiv_zmod_equality}, @{thm zdiv_zmod_equality2}]; - - val trans = trans; - - val prove_eq_sums = Arith_Data.prove_conv2 all_tac (Arith_Data.simp_all_tac - (@{thm diff_minus} :: @{thms add_0s} @ @{thms add_ac})) - -end) - -in - -val cancel_div_mod_int_proc = Simplifier.simproc @{theory} - "cancel_zdiv_zmod" ["(k::int) + l"] (K CancelDivMod.proc); - -val _ = Addsimprocs [cancel_div_mod_int_proc]; - -end -*} - -lemma pos_mod_conj : "(0::int) < b ==> 0 \ a mod b & a mod b < b" -apply (cut_tac a = a and b = b in divmod_correct) -apply (auto simp add: divmod_rel_def mod_def) -done - -lemmas pos_mod_sign [simp] = pos_mod_conj [THEN conjunct1, standard] - and pos_mod_bound [simp] = pos_mod_conj [THEN conjunct2, standard] - -lemma neg_mod_conj : "b < (0::int) ==> a mod b \ 0 & b < a mod b" -apply (cut_tac a = a and b = b in divmod_correct) -apply (auto simp add: divmod_rel_def div_def mod_def) -done - -lemmas neg_mod_sign [simp] = neg_mod_conj [THEN conjunct1, standard] - and neg_mod_bound [simp] = neg_mod_conj [THEN conjunct2, standard] - - - -subsection{*General Properties of div and mod*} - -lemma divmod_rel_div_mod: "b \ 0 ==> divmod_rel a b (a div b, a mod b)" -apply (cut_tac a = a and b = b in zmod_zdiv_equality) -apply (force simp add: divmod_rel_def linorder_neq_iff) -done - -lemma divmod_rel_div: "[| divmod_rel a b (q, r); b \ 0 |] ==> a div b = q" -by (simp add: divmod_rel_div_mod [THEN unique_quotient]) - -lemma divmod_rel_mod: "[| divmod_rel a b (q, r); b \ 0 |] ==> a mod b = r" -by (simp add: divmod_rel_div_mod [THEN unique_remainder]) - -lemma div_pos_pos_trivial: "[| (0::int) \ a; a < b |] ==> a div b = 0" -apply (rule divmod_rel_div) -apply (auto simp add: divmod_rel_def) -done - -lemma div_neg_neg_trivial: "[| a \ (0::int); b < a |] ==> a div b = 0" -apply (rule divmod_rel_div) -apply (auto simp add: divmod_rel_def) -done - -lemma div_pos_neg_trivial: "[| (0::int) < a; a+b \ 0 |] ==> a div b = -1" -apply (rule divmod_rel_div) -apply (auto simp add: divmod_rel_def) -done - -(*There is no div_neg_pos_trivial because 0 div b = 0 would supersede it*) - -lemma mod_pos_pos_trivial: "[| (0::int) \ a; a < b |] ==> a mod b = a" -apply (rule_tac q = 0 in divmod_rel_mod) -apply (auto simp add: divmod_rel_def) -done - -lemma mod_neg_neg_trivial: "[| a \ (0::int); b < a |] ==> a mod b = a" -apply (rule_tac q = 0 in divmod_rel_mod) -apply (auto simp add: divmod_rel_def) -done - -lemma mod_pos_neg_trivial: "[| (0::int) < a; a+b \ 0 |] ==> a mod b = a+b" -apply (rule_tac q = "-1" in divmod_rel_mod) -apply (auto simp add: divmod_rel_def) -done - -text{*There is no @{text mod_neg_pos_trivial}.*} - - -(*Simpler laws such as -a div b = -(a div b) FAIL, but see just below*) -lemma zdiv_zminus_zminus [simp]: "(-a) div (-b) = a div (b::int)" -apply (case_tac "b = 0", simp) -apply (simp add: divmod_rel_div_mod [THEN divmod_rel_neg, simplified, - THEN divmod_rel_div, THEN sym]) - -done - -(*Simpler laws such as -a mod b = -(a mod b) FAIL, but see just below*) -lemma zmod_zminus_zminus [simp]: "(-a) mod (-b) = - (a mod (b::int))" -apply (case_tac "b = 0", simp) -apply (subst divmod_rel_div_mod [THEN divmod_rel_neg, simplified, THEN divmod_rel_mod], - auto) -done - - -subsection{*Laws for div and mod with Unary Minus*} - -lemma zminus1_lemma: - "divmod_rel a b (q, r) - ==> divmod_rel (-a) b (if r=0 then -q else -q - 1, - if r=0 then 0 else b-r)" -by (force simp add: split_ifs divmod_rel_def linorder_neq_iff right_diff_distrib) - - -lemma zdiv_zminus1_eq_if: - "b \ (0::int) - ==> (-a) div b = - (if a mod b = 0 then - (a div b) else - (a div b) - 1)" -by (blast intro: divmod_rel_div_mod [THEN zminus1_lemma, THEN divmod_rel_div]) - -lemma zmod_zminus1_eq_if: - "(-a::int) mod b = (if a mod b = 0 then 0 else b - (a mod b))" -apply (case_tac "b = 0", simp) -apply (blast intro: divmod_rel_div_mod [THEN zminus1_lemma, THEN divmod_rel_mod]) -done - -lemma zmod_zminus1_not_zero: - fixes k l :: int - shows "- k mod l \ 0 \ k mod l \ 0" - unfolding zmod_zminus1_eq_if by auto - -lemma zdiv_zminus2: "a div (-b) = (-a::int) div b" -by (cut_tac a = "-a" in zdiv_zminus_zminus, auto) - -lemma zmod_zminus2: "a mod (-b) = - ((-a::int) mod b)" -by (cut_tac a = "-a" and b = b in zmod_zminus_zminus, auto) - -lemma zdiv_zminus2_eq_if: - "b \ (0::int) - ==> a div (-b) = - (if a mod b = 0 then - (a div b) else - (a div b) - 1)" -by (simp add: zdiv_zminus1_eq_if zdiv_zminus2) - -lemma zmod_zminus2_eq_if: - "a mod (-b::int) = (if a mod b = 0 then 0 else (a mod b) - b)" -by (simp add: zmod_zminus1_eq_if zmod_zminus2) - -lemma zmod_zminus2_not_zero: - fixes k l :: int - shows "k mod - l \ 0 \ k mod l \ 0" - unfolding zmod_zminus2_eq_if by auto - - -subsection{*Division of a Number by Itself*} - -lemma self_quotient_aux1: "[| (0::int) < a; a = r + a*q; r < a |] ==> 1 \ q" -apply (subgoal_tac "0 < a*q") - apply (simp add: zero_less_mult_iff, arith) -done - -lemma self_quotient_aux2: "[| (0::int) < a; a = r + a*q; 0 \ r |] ==> q \ 1" -apply (subgoal_tac "0 \ a* (1-q) ") - apply (simp add: zero_le_mult_iff) -apply (simp add: right_diff_distrib) -done - -lemma self_quotient: "[| divmod_rel a a (q, r); a \ (0::int) |] ==> q = 1" -apply (simp add: split_ifs divmod_rel_def linorder_neq_iff) -apply (rule order_antisym, safe, simp_all) -apply (rule_tac [3] a = "-a" and r = "-r" in self_quotient_aux1) -apply (rule_tac a = "-a" and r = "-r" in self_quotient_aux2) -apply (force intro: self_quotient_aux1 self_quotient_aux2 simp add: add_commute)+ -done - -lemma self_remainder: "[| divmod_rel a a (q, r); a \ (0::int) |] ==> r = 0" -apply (frule self_quotient, assumption) -apply (simp add: divmod_rel_def) -done - -lemma zdiv_self [simp]: "a \ 0 ==> a div a = (1::int)" -by (simp add: divmod_rel_div_mod [THEN self_quotient]) - -(*Here we have 0 mod 0 = 0, also assumed by Knuth (who puts m mod 0 = 0) *) -lemma zmod_self [simp]: "a mod a = (0::int)" -apply (case_tac "a = 0", simp) -apply (simp add: divmod_rel_div_mod [THEN self_remainder]) -done - - -subsection{*Computation of Division and Remainder*} - -lemma zdiv_zero [simp]: "(0::int) div b = 0" -by (simp add: div_def divmod_def) - -lemma div_eq_minus1: "(0::int) < b ==> -1 div b = -1" -by (simp add: div_def divmod_def) - -lemma zmod_zero [simp]: "(0::int) mod b = 0" -by (simp add: mod_def divmod_def) - -lemma zmod_minus1: "(0::int) < b ==> -1 mod b = b - 1" -by (simp add: mod_def divmod_def) - -text{*a positive, b positive *} - -lemma div_pos_pos: "[| 0 < a; 0 \ b |] ==> a div b = fst (posDivAlg a b)" -by (simp add: div_def divmod_def) - -lemma mod_pos_pos: "[| 0 < a; 0 \ b |] ==> a mod b = snd (posDivAlg a b)" -by (simp add: mod_def divmod_def) - -text{*a negative, b positive *} - -lemma div_neg_pos: "[| a < 0; 0 < b |] ==> a div b = fst (negDivAlg a b)" -by (simp add: div_def divmod_def) - -lemma mod_neg_pos: "[| a < 0; 0 < b |] ==> a mod b = snd (negDivAlg a b)" -by (simp add: mod_def divmod_def) - -text{*a positive, b negative *} - -lemma div_pos_neg: - "[| 0 < a; b < 0 |] ==> a div b = fst (negateSnd (negDivAlg (-a) (-b)))" -by (simp add: div_def divmod_def) - -lemma mod_pos_neg: - "[| 0 < a; b < 0 |] ==> a mod b = snd (negateSnd (negDivAlg (-a) (-b)))" -by (simp add: mod_def divmod_def) - -text{*a negative, b negative *} - -lemma div_neg_neg: - "[| a < 0; b \ 0 |] ==> a div b = fst (negateSnd (posDivAlg (-a) (-b)))" -by (simp add: div_def divmod_def) - -lemma mod_neg_neg: - "[| a < 0; b \ 0 |] ==> a mod b = snd (negateSnd (posDivAlg (-a) (-b)))" -by (simp add: mod_def divmod_def) - -text {*Simplify expresions in which div and mod combine numerical constants*} - -lemma divmod_relI: - "\a == b * q + r; if 0 < b then 0 \ r \ r < b else b < r \ r \ 0\ - \ divmod_rel a b (q, r)" - unfolding divmod_rel_def by simp - -lemmas divmod_rel_div_eq = divmod_relI [THEN divmod_rel_div, THEN eq_reflection] -lemmas divmod_rel_mod_eq = divmod_relI [THEN divmod_rel_mod, THEN eq_reflection] -lemmas arithmetic_simps = - arith_simps - add_special - OrderedGroup.add_0_left - OrderedGroup.add_0_right - mult_zero_left - mult_zero_right - mult_1_left - mult_1_right - -(* simprocs adapted from HOL/ex/Binary.thy *) -ML {* -local - val mk_number = HOLogic.mk_number HOLogic.intT; - fun mk_cert u k l = @{term "plus :: int \ int \ int"} $ - (@{term "times :: int \ int \ int"} $ u $ mk_number k) $ - mk_number l; - fun prove ctxt prop = Goal.prove ctxt [] [] prop - (K (ALLGOALS (full_simp_tac (HOL_basic_ss addsimps @{thms arithmetic_simps})))); - fun binary_proc proc ss ct = - (case Thm.term_of ct of - _ $ t $ u => - (case try (pairself (`(snd o HOLogic.dest_number))) (t, u) of - SOME args => proc (Simplifier.the_context ss) args - | NONE => NONE) - | _ => NONE); -in - fun divmod_proc rule = binary_proc (fn ctxt => fn ((m, t), (n, u)) => - if n = 0 then NONE - else let val (k, l) = Integer.div_mod m n; - in SOME (rule OF [prove ctxt (Logic.mk_equals (t, mk_cert u k l))]) end); -end -*} - -simproc_setup binary_int_div ("number_of m div number_of n :: int") = - {* K (divmod_proc (@{thm divmod_rel_div_eq})) *} - -simproc_setup binary_int_mod ("number_of m mod number_of n :: int") = - {* K (divmod_proc (@{thm divmod_rel_mod_eq})) *} - -lemmas posDivAlg_eqn_number_of [simp] = - posDivAlg_eqn [of "number_of v" "number_of w", standard] - -lemmas negDivAlg_eqn_number_of [simp] = - negDivAlg_eqn [of "number_of v" "number_of w", standard] - - -text{*Special-case simplification *} - -lemma zmod_minus1_right [simp]: "a mod (-1::int) = 0" -apply (cut_tac a = a and b = "-1" in neg_mod_sign) -apply (cut_tac [2] a = a and b = "-1" in neg_mod_bound) -apply (auto simp del: neg_mod_sign neg_mod_bound) -done - -lemma zdiv_minus1_right [simp]: "a div (-1::int) = -a" -by (cut_tac a = a and b = "-1" in zmod_zdiv_equality, auto) - -(** The last remaining special cases for constant arithmetic: - 1 div z and 1 mod z **) - -lemmas div_pos_pos_1_number_of [simp] = - div_pos_pos [OF int_0_less_1, of "number_of w", standard] - -lemmas div_pos_neg_1_number_of [simp] = - div_pos_neg [OF int_0_less_1, of "number_of w", standard] - -lemmas mod_pos_pos_1_number_of [simp] = - mod_pos_pos [OF int_0_less_1, of "number_of w", standard] - -lemmas mod_pos_neg_1_number_of [simp] = - mod_pos_neg [OF int_0_less_1, of "number_of w", standard] - - -lemmas posDivAlg_eqn_1_number_of [simp] = - posDivAlg_eqn [of concl: 1 "number_of w", standard] - -lemmas negDivAlg_eqn_1_number_of [simp] = - negDivAlg_eqn [of concl: 1 "number_of w", standard] - - - -subsection{*Monotonicity in the First Argument (Dividend)*} - -lemma zdiv_mono1: "[| a \ a'; 0 < (b::int) |] ==> a div b \ a' div b" -apply (cut_tac a = a and b = b in zmod_zdiv_equality) -apply (cut_tac a = a' and b = b in zmod_zdiv_equality) -apply (rule unique_quotient_lemma) -apply (erule subst) -apply (erule subst, simp_all) -done - -lemma zdiv_mono1_neg: "[| a \ a'; (b::int) < 0 |] ==> a' div b \ a div b" -apply (cut_tac a = a and b = b in zmod_zdiv_equality) -apply (cut_tac a = a' and b = b in zmod_zdiv_equality) -apply (rule unique_quotient_lemma_neg) -apply (erule subst) -apply (erule subst, simp_all) -done - - -subsection{*Monotonicity in the Second Argument (Divisor)*} - -lemma q_pos_lemma: - "[| 0 \ b'*q' + r'; r' < b'; 0 < b' |] ==> 0 \ (q'::int)" -apply (subgoal_tac "0 < b'* (q' + 1) ") - apply (simp add: zero_less_mult_iff) -apply (simp add: right_distrib) -done - -lemma zdiv_mono2_lemma: - "[| b*q + r = b'*q' + r'; 0 \ b'*q' + r'; - r' < b'; 0 \ r; 0 < b'; b' \ b |] - ==> q \ (q'::int)" -apply (frule q_pos_lemma, assumption+) -apply (subgoal_tac "b*q < b* (q' + 1) ") - apply (simp add: mult_less_cancel_left) -apply (subgoal_tac "b*q = r' - r + b'*q'") - prefer 2 apply simp -apply (simp (no_asm_simp) add: right_distrib) -apply (subst add_commute, rule zadd_zless_mono, arith) -apply (rule mult_right_mono, auto) -done - -lemma zdiv_mono2: - "[| (0::int) \ a; 0 < b'; b' \ b |] ==> a div b \ a div b'" -apply (subgoal_tac "b \ 0") - prefer 2 apply arith -apply (cut_tac a = a and b = b in zmod_zdiv_equality) -apply (cut_tac a = a and b = b' in zmod_zdiv_equality) -apply (rule zdiv_mono2_lemma) -apply (erule subst) -apply (erule subst, simp_all) -done - -lemma q_neg_lemma: - "[| b'*q' + r' < 0; 0 \ r'; 0 < b' |] ==> q' \ (0::int)" -apply (subgoal_tac "b'*q' < 0") - apply (simp add: mult_less_0_iff, arith) -done - -lemma zdiv_mono2_neg_lemma: - "[| b*q + r = b'*q' + r'; b'*q' + r' < 0; - r < b; 0 \ r'; 0 < b'; b' \ b |] - ==> q' \ (q::int)" -apply (frule q_neg_lemma, assumption+) -apply (subgoal_tac "b*q' < b* (q + 1) ") - apply (simp add: mult_less_cancel_left) -apply (simp add: right_distrib) -apply (subgoal_tac "b*q' \ b'*q'") - prefer 2 apply (simp add: mult_right_mono_neg, arith) -done - -lemma zdiv_mono2_neg: - "[| a < (0::int); 0 < b'; b' \ b |] ==> a div b' \ a div b" -apply (cut_tac a = a and b = b in zmod_zdiv_equality) -apply (cut_tac a = a and b = b' in zmod_zdiv_equality) -apply (rule zdiv_mono2_neg_lemma) -apply (erule subst) -apply (erule subst, simp_all) -done - - -subsection{*More Algebraic Laws for div and mod*} - -text{*proving (a*b) div c = a * (b div c) + a * (b mod c) *} - -lemma zmult1_lemma: - "[| divmod_rel b c (q, r); c \ 0 |] - ==> divmod_rel (a * b) c (a*q + a*r div c, a*r mod c)" -by (force simp add: split_ifs divmod_rel_def linorder_neq_iff right_distrib) - -lemma zdiv_zmult1_eq: "(a*b) div c = a*(b div c) + a*(b mod c) div (c::int)" -apply (case_tac "c = 0", simp) -apply (blast intro: divmod_rel_div_mod [THEN zmult1_lemma, THEN divmod_rel_div]) -done - -lemma zmod_zmult1_eq: "(a*b) mod c = a*(b mod c) mod (c::int)" -apply (case_tac "c = 0", simp) -apply (blast intro: divmod_rel_div_mod [THEN zmult1_lemma, THEN divmod_rel_mod]) -done - -lemma zmod_zdiv_trivial: "(a mod b) div b = (0::int)" -apply (case_tac "b = 0", simp) -apply (auto simp add: linorder_neq_iff div_pos_pos_trivial div_neg_neg_trivial) -done - -text{*proving (a+b) div c = a div c + b div c + ((a mod c + b mod c) div c) *} - -lemma zadd1_lemma: - "[| divmod_rel a c (aq, ar); divmod_rel b c (bq, br); c \ 0 |] - ==> divmod_rel (a+b) c (aq + bq + (ar+br) div c, (ar+br) mod c)" -by (force simp add: split_ifs divmod_rel_def linorder_neq_iff right_distrib) - -(*NOT suitable for rewriting: the RHS has an instance of the LHS*) -lemma zdiv_zadd1_eq: - "(a+b) div (c::int) = a div c + b div c + ((a mod c + b mod c) div c)" -apply (case_tac "c = 0", simp) -apply (blast intro: zadd1_lemma [OF divmod_rel_div_mod divmod_rel_div_mod] divmod_rel_div) -done - -instance int :: ring_div -proof - fix a b c :: int - assume not0: "b \ 0" - show "(a + c * b) div b = c + a div b" - unfolding zdiv_zadd1_eq [of a "c * b"] using not0 - by (simp add: zmod_zmult1_eq zmod_zdiv_trivial zdiv_zmult1_eq) -next - fix a b c :: int - assume "a \ 0" - then show "(a * b) div (a * c) = b div c" - proof (cases "b \ 0 \ c \ 0") - case False then show ?thesis by auto - next - case True then have "b \ 0" and "c \ 0" by auto - with `a \ 0` - have "\q r. divmod_rel b c (q, r) \ divmod_rel (a * b) (a * c) (q, a * r)" - apply (auto simp add: divmod_rel_def) - apply (auto simp add: algebra_simps) - apply (auto simp add: zero_less_mult_iff zero_le_mult_iff mult_le_0_iff) - done - moreover with `c \ 0` divmod_rel_div_mod have "divmod_rel b c (b div c, b mod c)" by auto - ultimately have "divmod_rel (a * b) (a * c) (b div c, a * (b mod c))" . - moreover from `a \ 0` `c \ 0` have "a * c \ 0" by simp - ultimately show ?thesis by (rule divmod_rel_div) - qed -qed auto - -lemma posDivAlg_div_mod: - assumes "k \ 0" - and "l \ 0" - shows "posDivAlg k l = (k div l, k mod l)" -proof (cases "l = 0") - case True then show ?thesis by (simp add: posDivAlg.simps) -next - case False with assms posDivAlg_correct - have "divmod_rel k l (fst (posDivAlg k l), snd (posDivAlg k l))" - by simp - from divmod_rel_div [OF this `l \ 0`] divmod_rel_mod [OF this `l \ 0`] - show ?thesis by simp -qed - -lemma negDivAlg_div_mod: - assumes "k < 0" - and "l > 0" - shows "negDivAlg k l = (k div l, k mod l)" -proof - - from assms have "l \ 0" by simp - from assms negDivAlg_correct - have "divmod_rel k l (fst (negDivAlg k l), snd (negDivAlg k l))" - by simp - from divmod_rel_div [OF this `l \ 0`] divmod_rel_mod [OF this `l \ 0`] - show ?thesis by simp -qed - -lemma zmod_eq_0_iff: "(m mod d = 0) = (EX q::int. m = d*q)" -by (simp add: dvd_eq_mod_eq_0 [symmetric] dvd_def) - -(* REVISIT: should this be generalized to all semiring_div types? *) -lemmas zmod_eq_0D [dest!] = zmod_eq_0_iff [THEN iffD1] - - -subsection{*Proving @{term "a div (b*c) = (a div b) div c"} *} - -(*The condition c>0 seems necessary. Consider that 7 div ~6 = ~2 but - 7 div 2 div ~3 = 3 div ~3 = ~1. The subcase (a div b) mod c = 0 seems - to cause particular problems.*) - -text{*first, four lemmas to bound the remainder for the cases b<0 and b>0 *} - -lemma zmult2_lemma_aux1: "[| (0::int) < c; b < r; r \ 0 |] ==> b*c < b*(q mod c) + r" -apply (subgoal_tac "b * (c - q mod c) < r * 1") - apply (simp add: algebra_simps) -apply (rule order_le_less_trans) - apply (erule_tac [2] mult_strict_right_mono) - apply (rule mult_left_mono_neg) - using add1_zle_eq[of "q mod c"]apply(simp add: algebra_simps pos_mod_bound) - apply (simp) -apply (simp) -done - -lemma zmult2_lemma_aux2: - "[| (0::int) < c; b < r; r \ 0 |] ==> b * (q mod c) + r \ 0" -apply (subgoal_tac "b * (q mod c) \ 0") - apply arith -apply (simp add: mult_le_0_iff) -done - -lemma zmult2_lemma_aux3: "[| (0::int) < c; 0 \ r; r < b |] ==> 0 \ b * (q mod c) + r" -apply (subgoal_tac "0 \ b * (q mod c) ") -apply arith -apply (simp add: zero_le_mult_iff) -done - -lemma zmult2_lemma_aux4: "[| (0::int) < c; 0 \ r; r < b |] ==> b * (q mod c) + r < b * c" -apply (subgoal_tac "r * 1 < b * (c - q mod c) ") - apply (simp add: right_diff_distrib) -apply (rule order_less_le_trans) - apply (erule mult_strict_right_mono) - apply (rule_tac [2] mult_left_mono) - apply simp - using add1_zle_eq[of "q mod c"] apply (simp add: algebra_simps pos_mod_bound) -apply simp -done - -lemma zmult2_lemma: "[| divmod_rel a b (q, r); b \ 0; 0 < c |] - ==> divmod_rel a (b * c) (q div c, b*(q mod c) + r)" -by (auto simp add: mult_ac divmod_rel_def linorder_neq_iff - zero_less_mult_iff right_distrib [symmetric] - zmult2_lemma_aux1 zmult2_lemma_aux2 zmult2_lemma_aux3 zmult2_lemma_aux4) - -lemma zdiv_zmult2_eq: "(0::int) < c ==> a div (b*c) = (a div b) div c" -apply (case_tac "b = 0", simp) -apply (force simp add: divmod_rel_div_mod [THEN zmult2_lemma, THEN divmod_rel_div]) -done - -lemma zmod_zmult2_eq: - "(0::int) < c ==> a mod (b*c) = b*(a div b mod c) + a mod b" -apply (case_tac "b = 0", simp) -apply (force simp add: divmod_rel_div_mod [THEN zmult2_lemma, THEN divmod_rel_mod]) -done - - -subsection {*Splitting Rules for div and mod*} - -text{*The proofs of the two lemmas below are essentially identical*} - -lemma split_pos_lemma: - "0 - P(n div k :: int)(n mod k) = (\i j. 0\j & j P i j)" -apply (rule iffI, clarify) - apply (erule_tac P="P ?x ?y" in rev_mp) - apply (subst mod_add_eq) - apply (subst zdiv_zadd1_eq) - apply (simp add: div_pos_pos_trivial mod_pos_pos_trivial) -txt{*converse direction*} -apply (drule_tac x = "n div k" in spec) -apply (drule_tac x = "n mod k" in spec, simp) -done - -lemma split_neg_lemma: - "k<0 ==> - P(n div k :: int)(n mod k) = (\i j. k0 & n = k*i + j --> P i j)" -apply (rule iffI, clarify) - apply (erule_tac P="P ?x ?y" in rev_mp) - apply (subst mod_add_eq) - apply (subst zdiv_zadd1_eq) - apply (simp add: div_neg_neg_trivial mod_neg_neg_trivial) -txt{*converse direction*} -apply (drule_tac x = "n div k" in spec) -apply (drule_tac x = "n mod k" in spec, simp) -done - -lemma split_zdiv: - "P(n div k :: int) = - ((k = 0 --> P 0) & - (0 (\i j. 0\j & j P i)) & - (k<0 --> (\i j. k0 & n = k*i + j --> P i)))" -apply (case_tac "k=0", simp) -apply (simp only: linorder_neq_iff) -apply (erule disjE) - apply (simp_all add: split_pos_lemma [of concl: "%x y. P x"] - split_neg_lemma [of concl: "%x y. P x"]) -done - -lemma split_zmod: - "P(n mod k :: int) = - ((k = 0 --> P n) & - (0 (\i j. 0\j & j P j)) & - (k<0 --> (\i j. k0 & n = k*i + j --> P j)))" -apply (case_tac "k=0", simp) -apply (simp only: linorder_neq_iff) -apply (erule disjE) - apply (simp_all add: split_pos_lemma [of concl: "%x y. P y"] - split_neg_lemma [of concl: "%x y. P y"]) -done - -(* Enable arith to deal with div 2 and mod 2: *) -declare split_zdiv [of _ _ "number_of k", simplified, standard, arith_split] -declare split_zmod [of _ _ "number_of k", simplified, standard, arith_split] - - -subsection{*Speeding up the Division Algorithm with Shifting*} - -text{*computing div by shifting *} - -lemma pos_zdiv_mult_2: "(0::int) \ a ==> (1 + 2*b) div (2*a) = b div a" -proof cases - assume "a=0" - thus ?thesis by simp -next - assume "a\0" and le_a: "0\a" - hence a_pos: "1 \ a" by arith - hence one_less_a2: "1 < 2 * a" by arith - hence le_2a: "2 * (1 + b mod a) \ 2 * a" - unfolding mult_le_cancel_left - by (simp add: add1_zle_eq add_commute [of 1]) - with a_pos have "0 \ b mod a" by simp - hence le_addm: "0 \ 1 mod (2*a) + 2*(b mod a)" - by (simp add: mod_pos_pos_trivial one_less_a2) - with le_2a - have "(1 mod (2*a) + 2*(b mod a)) div (2*a) = 0" - by (simp add: div_pos_pos_trivial le_addm mod_pos_pos_trivial one_less_a2 - right_distrib) - thus ?thesis - by (subst zdiv_zadd1_eq, - simp add: mod_mult_mult1 one_less_a2 - div_pos_pos_trivial) -qed - -lemma neg_zdiv_mult_2: "a \ (0::int) ==> (1 + 2*b) div (2*a) = (b+1) div a" -apply (subgoal_tac " (1 + 2* (-b - 1)) div (2 * (-a)) = (-b - 1) div (-a) ") -apply (rule_tac [2] pos_zdiv_mult_2) -apply (auto simp add: minus_mult_right [symmetric] right_diff_distrib) -apply (subgoal_tac " (-1 - (2 * b)) = - (1 + (2 * b))") -apply (simp only: zdiv_zminus_zminus diff_minus minus_add_distrib [symmetric], - simp) -done - -lemma zdiv_number_of_Bit0 [simp]: - "number_of (Int.Bit0 v) div number_of (Int.Bit0 w) = - number_of v div (number_of w :: int)" -by (simp only: number_of_eq numeral_simps) simp - -lemma zdiv_number_of_Bit1 [simp]: - "number_of (Int.Bit1 v) div number_of (Int.Bit0 w) = - (if (0::int) \ number_of w - then number_of v div (number_of w) - else (number_of v + (1::int)) div (number_of w))" -apply (simp only: number_of_eq numeral_simps UNIV_I split: split_if) -apply (simp add: pos_zdiv_mult_2 neg_zdiv_mult_2 add_ac) -done - - -subsection{*Computing mod by Shifting (proofs resemble those for div)*} - -lemma pos_zmod_mult_2: - "(0::int) \ a ==> (1 + 2*b) mod (2*a) = 1 + 2 * (b mod a)" -apply (case_tac "a = 0", simp) -apply (subgoal_tac "1 < a * 2") - prefer 2 apply arith -apply (subgoal_tac "2* (1 + b mod a) \ 2*a") - apply (rule_tac [2] mult_left_mono) -apply (auto simp add: add_commute [of 1] mult_commute add1_zle_eq - pos_mod_bound) -apply (subst mod_add_eq) -apply (simp add: mod_mult_mult2 mod_pos_pos_trivial) -apply (rule mod_pos_pos_trivial) -apply (auto simp add: mod_pos_pos_trivial ring_distribs) -apply (subgoal_tac "0 \ b mod a", arith, simp) -done - -lemma neg_zmod_mult_2: - "a \ (0::int) ==> (1 + 2*b) mod (2*a) = 2 * ((b+1) mod a) - 1" -apply (subgoal_tac "(1 + 2* (-b - 1)) mod (2* (-a)) = - 1 + 2* ((-b - 1) mod (-a))") -apply (rule_tac [2] pos_zmod_mult_2) -apply (auto simp add: right_diff_distrib) -apply (subgoal_tac " (-1 - (2 * b)) = - (1 + (2 * b))") - prefer 2 apply simp -apply (simp only: zmod_zminus_zminus diff_minus minus_add_distrib [symmetric]) -done - -lemma zmod_number_of_Bit0 [simp]: - "number_of (Int.Bit0 v) mod number_of (Int.Bit0 w) = - (2::int) * (number_of v mod number_of w)" -apply (simp only: number_of_eq numeral_simps) -apply (simp add: mod_mult_mult1 pos_zmod_mult_2 - neg_zmod_mult_2 add_ac) -done - -lemma zmod_number_of_Bit1 [simp]: - "number_of (Int.Bit1 v) mod number_of (Int.Bit0 w) = - (if (0::int) \ number_of w - then 2 * (number_of v mod number_of w) + 1 - else 2 * ((number_of v + (1::int)) mod number_of w) - 1)" -apply (simp only: number_of_eq numeral_simps) -apply (simp add: mod_mult_mult1 pos_zmod_mult_2 - neg_zmod_mult_2 add_ac) -done - - -subsection{*Quotients of Signs*} - -lemma div_neg_pos_less0: "[| a < (0::int); 0 < b |] ==> a div b < 0" -apply (subgoal_tac "a div b \ -1", force) -apply (rule order_trans) -apply (rule_tac a' = "-1" in zdiv_mono1) -apply (auto simp add: div_eq_minus1) -done - -lemma div_nonneg_neg_le0: "[| (0::int) \ a; b < 0 |] ==> a div b \ 0" -by (drule zdiv_mono1_neg, auto) - -lemma div_nonpos_pos_le0: "[| (a::int) \ 0; b > 0 |] ==> a div b \ 0" -by (drule zdiv_mono1, auto) - -lemma pos_imp_zdiv_nonneg_iff: "(0::int) < b ==> (0 \ a div b) = (0 \ a)" -apply auto -apply (drule_tac [2] zdiv_mono1) -apply (auto simp add: linorder_neq_iff) -apply (simp (no_asm_use) add: linorder_not_less [symmetric]) -apply (blast intro: div_neg_pos_less0) -done - -lemma neg_imp_zdiv_nonneg_iff: - "b < (0::int) ==> (0 \ a div b) = (a \ (0::int))" -apply (subst zdiv_zminus_zminus [symmetric]) -apply (subst pos_imp_zdiv_nonneg_iff, auto) -done - -(*But not (a div b \ 0 iff a\0); consider a=1, b=2 when a div b = 0.*) -lemma pos_imp_zdiv_neg_iff: "(0::int) < b ==> (a div b < 0) = (a < 0)" -by (simp add: linorder_not_le [symmetric] pos_imp_zdiv_nonneg_iff) - -(*Again the law fails for \: consider a = -1, b = -2 when a div b = 0*) -lemma neg_imp_zdiv_neg_iff: "b < (0::int) ==> (a div b < 0) = (0 < a)" -by (simp add: linorder_not_le [symmetric] neg_imp_zdiv_nonneg_iff) - - -subsection {* The Divides Relation *} - -lemmas zdvd_iff_zmod_eq_0_number_of [simp] = - dvd_eq_mod_eq_0 [of "number_of x::int" "number_of y::int", standard] - -lemma zdvd_zmod: "f dvd m ==> f dvd (n::int) ==> f dvd m mod n" - by (rule dvd_mod) (* TODO: remove *) - -lemma zdvd_zmod_imp_zdvd: "k dvd m mod n ==> k dvd n ==> k dvd (m::int)" - by (rule dvd_mod_imp_dvd) (* TODO: remove *) - -lemma zmult_div_cancel: "(n::int) * (m div n) = m - (m mod n)" - using zmod_zdiv_equality[where a="m" and b="n"] - by (simp add: algebra_simps) - -lemma zpower_zmod: "((x::int) mod m)^y mod m = x^y mod m" -apply (induct "y", auto) -apply (rule zmod_zmult1_eq [THEN trans]) -apply (simp (no_asm_simp)) -apply (rule mod_mult_eq [symmetric]) -done - -lemma zdiv_int: "int (a div b) = (int a) div (int b)" -apply (subst split_div, auto) -apply (subst split_zdiv, auto) -apply (rule_tac a="int (b * i) + int j" and b="int b" and r="int j" and r'=ja in IntDiv.unique_quotient) -apply (auto simp add: IntDiv.divmod_rel_def of_nat_mult) -done - -lemma zmod_int: "int (a mod b) = (int a) mod (int b)" -apply (subst split_mod, auto) -apply (subst split_zmod, auto) -apply (rule_tac a="int (b * i) + int j" and b="int b" and q="int i" and q'=ia - in unique_remainder) -apply (auto simp add: IntDiv.divmod_rel_def of_nat_mult) -done - -lemma abs_div: "(y::int) dvd x \ abs (x div y) = abs x div abs y" -by (unfold dvd_def, cases "y=0", auto simp add: abs_mult) - -lemma zdvd_mult_div_cancel:"(n::int) dvd m \ n * (m div n) = m" -apply (subgoal_tac "m mod n = 0") - apply (simp add: zmult_div_cancel) -apply (simp only: dvd_eq_mod_eq_0) -done - -text{*Suggested by Matthias Daum*} -lemma int_power_div_base: - "\0 < m; 0 < k\ \ k ^ m div k = (k::int) ^ (m - Suc 0)" -apply (subgoal_tac "k ^ m = k ^ ((m - Suc 0) + Suc 0)") - apply (erule ssubst) - apply (simp only: power_add) - apply simp_all -done - -text {* by Brian Huffman *} -lemma zminus_zmod: "- ((x::int) mod m) mod m = - x mod m" -by (rule mod_minus_eq [symmetric]) - -lemma zdiff_zmod_left: "(x mod m - y) mod m = (x - y) mod (m::int)" -by (rule mod_diff_left_eq [symmetric]) - -lemma zdiff_zmod_right: "(x - y mod m) mod m = (x - y) mod (m::int)" -by (rule mod_diff_right_eq [symmetric]) - -lemmas zmod_simps = - mod_add_left_eq [symmetric] - mod_add_right_eq [symmetric] - zmod_zmult1_eq [symmetric] - mod_mult_left_eq [symmetric] - zpower_zmod - zminus_zmod zdiff_zmod_left zdiff_zmod_right - -text {* Distributive laws for function @{text nat}. *} - -lemma nat_div_distrib: "0 \ x \ nat (x div y) = nat x div nat y" -apply (rule linorder_cases [of y 0]) -apply (simp add: div_nonneg_neg_le0) -apply simp -apply (simp add: nat_eq_iff pos_imp_zdiv_nonneg_iff zdiv_int) -done - -(*Fails if y<0: the LHS collapses to (nat z) but the RHS doesn't*) -lemma nat_mod_distrib: - "\0 \ x; 0 \ y\ \ nat (x mod y) = nat x mod nat y" -apply (case_tac "y = 0", simp add: DIVISION_BY_ZERO) -apply (simp add: nat_eq_iff zmod_int) -done - -text{*Suggested by Matthias Daum*} -lemma int_div_less_self: "\0 < x; 1 < k\ \ x div k < (x::int)" -apply (subgoal_tac "nat x div nat k < nat x") - apply (simp (asm_lr) add: nat_div_distrib [symmetric]) -apply (rule Divides.div_less_dividend, simp_all) -done - -text {* code generator setup *} - -context ring_1 -begin - -lemma of_int_num [code]: - "of_int k = (if k = 0 then 0 else if k < 0 then - - of_int (- k) else let - (l, m) = divmod k 2; - l' = of_int l - in if m = 0 then l' + l' else l' + l' + 1)" -proof - - have aux1: "k mod (2\int) \ (0\int) \ - of_int k = of_int (k div 2 * 2 + 1)" - proof - - have "k mod 2 < 2" by (auto intro: pos_mod_bound) - moreover have "0 \ k mod 2" by (auto intro: pos_mod_sign) - moreover assume "k mod 2 \ 0" - ultimately have "k mod 2 = 1" by arith - moreover have "of_int k = of_int (k div 2 * 2 + k mod 2)" by simp - ultimately show ?thesis by auto - qed - have aux2: "\x. of_int 2 * x = x + x" - proof - - fix x - have int2: "(2::int) = 1 + 1" by arith - show "of_int 2 * x = x + x" - unfolding int2 of_int_add left_distrib by simp - qed - have aux3: "\x. x * of_int 2 = x + x" - proof - - fix x - have int2: "(2::int) = 1 + 1" by arith - show "x * of_int 2 = x + x" - unfolding int2 of_int_add right_distrib by simp - qed - from aux1 show ?thesis by (auto simp add: divmod_mod_div Let_def aux2 aux3) -qed - -end - -lemma zmod_eq_dvd_iff: "(x::int) mod n = y mod n \ n dvd x - y" -proof - assume H: "x mod n = y mod n" - hence "x mod n - y mod n = 0" by simp - hence "(x mod n - y mod n) mod n = 0" by simp - hence "(x - y) mod n = 0" by (simp add: mod_diff_eq[symmetric]) - thus "n dvd x - y" by (simp add: dvd_eq_mod_eq_0) -next - assume H: "n dvd x - y" - then obtain k where k: "x-y = n*k" unfolding dvd_def by blast - hence "x = n*k + y" by simp - hence "x mod n = (n*k + y) mod n" by simp - thus "x mod n = y mod n" by (simp add: mod_add_left_eq) -qed - -lemma nat_mod_eq_lemma: assumes xyn: "(x::nat) mod n = y mod n" and xy:"y \ x" - shows "\q. x = y + n * q" -proof- - from xy have th: "int x - int y = int (x - y)" by simp - from xyn have "int x mod int n = int y mod int n" - by (simp add: zmod_int[symmetric]) - hence "int n dvd int x - int y" by (simp only: zmod_eq_dvd_iff[symmetric]) - hence "n dvd x - y" by (simp add: th zdvd_int) - then show ?thesis using xy unfolding dvd_def apply clarsimp apply (rule_tac x="k" in exI) by arith -qed - -lemma nat_mod_eq_iff: "(x::nat) mod n = y mod n \ (\q1 q2. x + n * q1 = y + n * q2)" - (is "?lhs = ?rhs") -proof - assume H: "x mod n = y mod n" - {assume xy: "x \ y" - from H have th: "y mod n = x mod n" by simp - from nat_mod_eq_lemma[OF th xy] have ?rhs - apply clarify apply (rule_tac x="q" in exI) by (rule exI[where x="0"], simp)} - moreover - {assume xy: "y \ x" - from nat_mod_eq_lemma[OF H xy] have ?rhs - apply clarify apply (rule_tac x="0" in exI) by (rule_tac x="q" in exI, simp)} - ultimately show ?rhs using linear[of x y] by blast -next - assume ?rhs then obtain q1 q2 where q12: "x + n * q1 = y + n * q2" by blast - hence "(x + n * q1) mod n = (y + n * q2) mod n" by simp - thus ?lhs by simp -qed - -lemma div_nat_number_of [simp]: - "(number_of v :: nat) div number_of v' = - (if neg (number_of v :: int) then 0 - else nat (number_of v div number_of v'))" - unfolding nat_number_of_def number_of_is_id neg_def - by (simp add: nat_div_distrib) - -lemma one_div_nat_number_of [simp]: - "Suc 0 div number_of v' = nat (1 div number_of v')" -by (simp del: nat_numeral_1_eq_1 add: numeral_1_eq_Suc_0 [symmetric]) - -lemma mod_nat_number_of [simp]: - "(number_of v :: nat) mod number_of v' = - (if neg (number_of v :: int) then 0 - else if neg (number_of v' :: int) then number_of v - else nat (number_of v mod number_of v'))" - unfolding nat_number_of_def number_of_is_id neg_def - by (simp add: nat_mod_distrib) - -lemma one_mod_nat_number_of [simp]: - "Suc 0 mod number_of v' = - (if neg (number_of v' :: int) then Suc 0 - else nat (1 mod number_of v'))" -by (simp del: nat_numeral_1_eq_1 add: numeral_1_eq_Suc_0 [symmetric]) - -lemmas dvd_eq_mod_eq_0_number_of = - dvd_eq_mod_eq_0 [of "number_of x" "number_of y", standard] - -declare dvd_eq_mod_eq_0_number_of [simp] - - -subsection {* Transfer setup *} - -lemma transfer_nat_int_functions: - "(x::int) >= 0 \ y >= 0 \ (nat x) div (nat y) = nat (x div y)" - "(x::int) >= 0 \ y >= 0 \ (nat x) mod (nat y) = nat (x mod y)" - by (auto simp add: nat_div_distrib nat_mod_distrib) - -lemma transfer_nat_int_function_closures: - "(x::int) >= 0 \ y >= 0 \ x div y >= 0" - "(x::int) >= 0 \ y >= 0 \ x mod y >= 0" - apply (cases "y = 0") - apply (auto simp add: pos_imp_zdiv_nonneg_iff) - apply (cases "y = 0") - apply auto -done - -declare TransferMorphism_nat_int[transfer add return: - transfer_nat_int_functions - transfer_nat_int_function_closures -] - -lemma transfer_int_nat_functions: - "(int x) div (int y) = int (x div y)" - "(int x) mod (int y) = int (x mod y)" - by (auto simp add: zdiv_int zmod_int) - -lemma transfer_int_nat_function_closures: - "is_nat x \ is_nat y \ is_nat (x div y)" - "is_nat x \ is_nat y \ is_nat (x mod y)" - by (simp_all only: is_nat_def transfer_nat_int_function_closures) - -declare TransferMorphism_int_nat[transfer add return: - transfer_int_nat_functions - transfer_int_nat_function_closures -] - - -subsection {* Code generation *} - -definition pdivmod :: "int \ int \ int \ int" where - "pdivmod k l = (\k\ div \l\, \k\ mod \l\)" - -lemma pdivmod_posDivAlg [code]: - "pdivmod k l = (if l = 0 then (0, \k\) else posDivAlg \k\ \l\)" -by (subst posDivAlg_div_mod) (simp_all add: pdivmod_def) - -lemma divmod_pdivmod: "divmod k l = (if k = 0 then (0, 0) else if l = 0 then (0, k) else - apsnd ((op *) (sgn l)) (if 0 < l \ 0 \ k \ l < 0 \ k < 0 - then pdivmod k l - else (let (r, s) = pdivmod k l in - if s = 0 then (- r, 0) else (- r - 1, \l\ - s))))" -proof - - have aux: "\q::int. - k = l * q \ k = l * - q" by auto - show ?thesis - by (simp add: divmod_mod_div pdivmod_def) - (auto simp add: aux not_less not_le zdiv_zminus1_eq_if - zmod_zminus1_eq_if zdiv_zminus2_eq_if zmod_zminus2_eq_if) -qed - -lemma divmod_code [code]: "divmod k l = (if k = 0 then (0, 0) else if l = 0 then (0, k) else - apsnd ((op *) (sgn l)) (if sgn k = sgn l - then pdivmod k l - else (let (r, s) = pdivmod k l in - if s = 0 then (- r, 0) else (- r - 1, \l\ - s))))" -proof - - have "k \ 0 \ l \ 0 \ 0 < l \ 0 \ k \ l < 0 \ k < 0 \ sgn k = sgn l" - by (auto simp add: not_less sgn_if) - then show ?thesis by (simp add: divmod_pdivmod) -qed - -code_modulename SML - IntDiv Integer - -code_modulename OCaml - IntDiv Integer - -code_modulename Haskell - IntDiv Integer - -end diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/IsaMakefile --- a/src/HOL/IsaMakefile Thu Oct 29 16:23:57 2009 +0100 +++ b/src/HOL/IsaMakefile Wed Nov 04 11:40:59 2009 +0100 @@ -10,6 +10,7 @@ images: \ HOL \ HOL-Algebra \ + HOL-Boogie \ HOL-Base \ HOL-Main \ HOL-Multivariate_Analysis \ @@ -27,6 +28,7 @@ HOL-ex \ HOL-Auth \ HOL-Bali \ + HOL-Boogie_Examples \ HOL-Decision_Procs \ HOL-Extraction \ HOL-Hahn_Banach \ @@ -54,7 +56,7 @@ HOL-Probability \ HOL-Prolog \ HOL-SET_Protocol \ - HOL-SMT-Examples \ + HOL-SMT_Examples \ HOL-SizeChange \ HOL-Statespace \ HOL-Subst \ @@ -180,7 +182,6 @@ Tools/Function/function.ML \ Tools/Function/fun.ML \ Tools/Function/induction_scheme.ML \ - Tools/Function/inductive_wrap.ML \ Tools/Function/lexicographic_order.ML \ Tools/Function/measure_functions.ML \ Tools/Function/mutual.ML \ @@ -249,12 +250,12 @@ Groebner_Basis.thy \ Hilbert_Choice.thy \ Int.thy \ - IntDiv.thy \ List.thy \ Main.thy \ Map.thy \ Nat_Numeral.thy \ Nat_Transfer.thy \ + Numeral_Simprocs.thy \ Presburger.thy \ Predicate_Compile.thy \ Quickcheck.thy \ @@ -302,6 +303,7 @@ Tools/recdef.ML \ Tools/res_atp.ML \ Tools/res_axioms.ML \ + Tools/res_blacklist.ML \ Tools/res_clause.ML \ Tools/res_hol_clause.ML \ Tools/res_reconstruct.ML \ @@ -382,7 +384,6 @@ Library/While_Combinator.thy Library/Product_ord.thy \ Library/Char_nat.thy Library/Char_ord.thy Library/Option_ord.thy \ Library/Sublist_Order.thy Library/List_lexord.thy \ - Library/Commutative_Ring.thy Library/comm_ring.ML \ Library/Coinductive_List.thy Library/AssocList.thy \ Library/Formal_Power_Series.thy Library/Binomial.thy \ Library/Eval_Witness.thy Library/Code_Char.thy \ @@ -785,6 +786,9 @@ $(LOG)/HOL-Decision_Procs.gz: $(OUT)/HOL \ Decision_Procs/Approximation.thy \ + Decision_Procs/Commutative_Ring.thy \ + Decision_Procs/Commutative_Ring_Complete.thy \ + Decision_Procs/commutative_ring_tac.ML \ Decision_Procs/Cooper.thy \ Decision_Procs/cooper_tac.ML \ Decision_Procs/Dense_Linear_Order.thy \ @@ -795,6 +799,7 @@ Decision_Procs/Decision_Procs.thy \ Decision_Procs/ex/Dense_Linear_Order_Ex.thy \ Decision_Procs/ex/Approximation_Ex.thy \ + Decision_Procs/ex/Commutative_Ring_Ex.thy \ Decision_Procs/ROOT.ML @$(ISABELLE_TOOL) usedir $(OUT)/HOL Decision_Procs @@ -937,7 +942,7 @@ HOL-ex: HOL $(LOG)/HOL-ex.gz -$(LOG)/HOL-ex.gz: $(OUT)/HOL Library/Commutative_Ring.thy \ +$(LOG)/HOL-ex.gz: $(OUT)/HOL Decision_Procs/Commutative_Ring.thy \ Number_Theory/Primes.thy \ Tools/Predicate_Compile/predicate_compile_core.ML \ ex/Abstract_NAT.thy ex/Antiquote.thy ex/Arith_Examples.thy \ @@ -945,8 +950,8 @@ ex/Binary.thy ex/CTL.thy ex/Chinese.thy ex/Classical.thy \ ex/CodegenSML_Test.thy ex/Codegenerator_Candidates.thy \ ex/Codegenerator_Pretty.thy ex/Codegenerator_Pretty_Test.thy \ - ex/Codegenerator_Test.thy ex/Coherent.thy ex/Commutative_RingEx.thy \ - ex/Commutative_Ring_Complete.thy ex/Efficient_Nat_examples.thy \ + ex/Codegenerator_Test.thy ex/Coherent.thy \ + ex/Efficient_Nat_examples.thy \ ex/Eval_Examples.thy ex/Fundefs.thy ex/Groebner_Examples.thy \ ex/Guess.thy ex/HarmonicSeries.thy ex/Hebrew.thy \ ex/Hex_Bin_Examples.thy ex/Higher_Order_Logic.thy \ @@ -1209,7 +1214,7 @@ HOL-SMT: HOL-Word $(OUT)/HOL-SMT -$(OUT)/HOL-SMT: $(OUT)/HOL-Word SMT/SMT_Base.thy SMT/Z3.thy \ +$(OUT)/HOL-SMT: $(OUT)/HOL-Word SMT/ROOT.ML SMT/SMT_Base.thy SMT/Z3.thy \ SMT/SMT.thy SMT/Tools/smt_normalize.ML SMT/Tools/smt_monomorph.ML \ SMT/Tools/smt_translate.ML SMT/Tools/smtlib_interface.ML \ SMT/Tools/smt_solver.ML SMT/Tools/cvc3_solver.ML \ @@ -1219,11 +1224,11 @@ @cd SMT; $(ISABELLE_TOOL) usedir -b -g true $(OUT)/HOL-Word HOL-SMT -## HOL-SMT-Examples +## HOL-SMT_Examples -HOL-SMT-Examples: HOL-SMT $(LOG)/HOL-SMT-Examples.gz +HOL-SMT_Examples: HOL-SMT $(LOG)/HOL-SMT_Examples.gz -$(LOG)/HOL-SMT-Examples.gz: $(OUT)/HOL-SMT SMT/Examples/ROOT.ML \ +$(LOG)/HOL-SMT_Examples.gz: $(OUT)/HOL-SMT SMT/Examples/ROOT.ML \ SMT/Examples/SMT_Examples.thy SMT/Examples/cert/z3_arith_quant_01 \ SMT/Examples/cert/z3_arith_quant_01.proof \ SMT/Examples/cert/z3_arith_quant_02 \ @@ -1378,6 +1383,33 @@ @cd SMT; $(ISABELLE_TOOL) usedir $(OUT)/HOL-SMT Examples +## HOL-Boogie + +HOL-Boogie: HOL-SMT $(OUT)/HOL-Boogie + +$(OUT)/HOL-Boogie: $(OUT)/HOL-SMT Boogie/ROOT.ML Boogie/Boogie.thy \ + Boogie/Tools/boogie_vcs.ML Boogie/Tools/boogie_loader.ML \ + Boogie/Tools/boogie_commands.ML Boogie/Tools/boogie_split.ML + @cd Boogie; $(ISABELLE_TOOL) usedir -b -g true $(OUT)/HOL-SMT HOL-Boogie + + +## HOL-Boogie_Examples + +HOL-Boogie_Examples: HOL-Boogie $(LOG)/HOL-Boogie_Examples.gz + +$(LOG)/HOL-Boogie_Examples.gz: $(OUT)/HOL-SMT Boogie/Examples/ROOT.ML \ + Boogie/Examples/Boogie_Max.thy Boogie/Examples/Boogie_Max.b2i \ + Boogie/Examples/Boogie_Dijkstra.thy Boogie/Examples/VCC_Max.thy \ + Boogie/Examples/Boogie_Dijkstra.b2i Boogie/Examples/VCC_Max.b2i \ + Boogie/Examples/cert/Boogie_b_max \ + Boogie/Examples/cert/Boogie_b_max.proof \ + Boogie/Examples/cert/Boogie_b_Dijkstra \ + Boogie/Examples/cert/Boogie_b_Dijkstra.proof \ + Boogie/Examples/cert/VCC_b_maximum \ + Boogie/Examples/cert/VCC_b_maximum.proof + @cd Boogie; $(ISABELLE_TOOL) usedir $(OUT)/HOL-Boogie Examples + + ## clean clean: diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/Library/Commutative_Ring.thy --- a/src/HOL/Library/Commutative_Ring.thy Thu Oct 29 16:23:57 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,319 +0,0 @@ -(* Author: Bernhard Haeupler - -Proving equalities in commutative rings done "right" in Isabelle/HOL. -*) - -header {* Proving equalities in commutative rings *} - -theory Commutative_Ring -imports List Parity Main -uses ("comm_ring.ML") -begin - -text {* Syntax of multivariate polynomials (pol) and polynomial expressions. *} - -datatype 'a pol = - Pc 'a - | Pinj nat "'a pol" - | PX "'a pol" nat "'a pol" - -datatype 'a polex = - Pol "'a pol" - | Add "'a polex" "'a polex" - | Sub "'a polex" "'a polex" - | Mul "'a polex" "'a polex" - | Pow "'a polex" nat - | Neg "'a polex" - -text {* Interpretation functions for the shadow syntax. *} - -primrec - Ipol :: "'a::{comm_ring_1} list \ 'a pol \ 'a" -where - "Ipol l (Pc c) = c" - | "Ipol l (Pinj i P) = Ipol (drop i l) P" - | "Ipol l (PX P x Q) = Ipol l P * (hd l)^x + Ipol (drop 1 l) Q" - -primrec - Ipolex :: "'a::{comm_ring_1} list \ 'a polex \ 'a" -where - "Ipolex l (Pol P) = Ipol l P" - | "Ipolex l (Add P Q) = Ipolex l P + Ipolex l Q" - | "Ipolex l (Sub P Q) = Ipolex l P - Ipolex l Q" - | "Ipolex l (Mul P Q) = Ipolex l P * Ipolex l Q" - | "Ipolex l (Pow p n) = Ipolex l p ^ n" - | "Ipolex l (Neg P) = - Ipolex l P" - -text {* Create polynomial normalized polynomials given normalized inputs. *} - -definition - mkPinj :: "nat \ 'a pol \ 'a pol" where - "mkPinj x P = (case P of - Pc c \ Pc c | - Pinj y P \ Pinj (x + y) P | - PX p1 y p2 \ Pinj x P)" - -definition - mkPX :: "'a::{comm_ring} pol \ nat \ 'a pol \ 'a pol" where - "mkPX P i Q = (case P of - Pc c \ (if (c = 0) then (mkPinj 1 Q) else (PX P i Q)) | - Pinj j R \ PX P i Q | - PX P2 i2 Q2 \ (if (Q2 = (Pc 0)) then (PX P2 (i+i2) Q) else (PX P i Q)) )" - -text {* Defining the basic ring operations on normalized polynomials *} - -function - add :: "'a::{comm_ring} pol \ 'a pol \ 'a pol" (infixl "\" 65) -where - "Pc a \ Pc b = Pc (a + b)" - | "Pc c \ Pinj i P = Pinj i (P \ Pc c)" - | "Pinj i P \ Pc c = Pinj i (P \ Pc c)" - | "Pc c \ PX P i Q = PX P i (Q \ Pc c)" - | "PX P i Q \ Pc c = PX P i (Q \ Pc c)" - | "Pinj x P \ Pinj y Q = - (if x = y then mkPinj x (P \ Q) - else (if x > y then mkPinj y (Pinj (x - y) P \ Q) - else mkPinj x (Pinj (y - x) Q \ P)))" - | "Pinj x P \ PX Q y R = - (if x = 0 then P \ PX Q y R - else (if x = 1 then PX Q y (R \ P) - else PX Q y (R \ Pinj (x - 1) P)))" - | "PX P x R \ Pinj y Q = - (if y = 0 then PX P x R \ Q - else (if y = 1 then PX P x (R \ Q) - else PX P x (R \ Pinj (y - 1) Q)))" - | "PX P1 x P2 \ PX Q1 y Q2 = - (if x = y then mkPX (P1 \ Q1) x (P2 \ Q2) - else (if x > y then mkPX (PX P1 (x - y) (Pc 0) \ Q1) y (P2 \ Q2) - else mkPX (PX Q1 (y-x) (Pc 0) \ P1) x (P2 \ Q2)))" -by pat_completeness auto -termination by (relation "measure (\(x, y). size x + size y)") auto - -function - mul :: "'a::{comm_ring} pol \ 'a pol \ 'a pol" (infixl "\" 70) -where - "Pc a \ Pc b = Pc (a * b)" - | "Pc c \ Pinj i P = - (if c = 0 then Pc 0 else mkPinj i (P \ Pc c))" - | "Pinj i P \ Pc c = - (if c = 0 then Pc 0 else mkPinj i (P \ Pc c))" - | "Pc c \ PX P i Q = - (if c = 0 then Pc 0 else mkPX (P \ Pc c) i (Q \ Pc c))" - | "PX P i Q \ Pc c = - (if c = 0 then Pc 0 else mkPX (P \ Pc c) i (Q \ Pc c))" - | "Pinj x P \ Pinj y Q = - (if x = y then mkPinj x (P \ Q) else - (if x > y then mkPinj y (Pinj (x-y) P \ Q) - else mkPinj x (Pinj (y - x) Q \ P)))" - | "Pinj x P \ PX Q y R = - (if x = 0 then P \ PX Q y R else - (if x = 1 then mkPX (Pinj x P \ Q) y (R \ P) - else mkPX (Pinj x P \ Q) y (R \ Pinj (x - 1) P)))" - | "PX P x R \ Pinj y Q = - (if y = 0 then PX P x R \ Q else - (if y = 1 then mkPX (Pinj y Q \ P) x (R \ Q) - else mkPX (Pinj y Q \ P) x (R \ Pinj (y - 1) Q)))" - | "PX P1 x P2 \ PX Q1 y Q2 = - mkPX (P1 \ Q1) (x + y) (P2 \ Q2) \ - (mkPX (P1 \ mkPinj 1 Q2) x (Pc 0) \ - (mkPX (Q1 \ mkPinj 1 P2) y (Pc 0)))" -by pat_completeness auto -termination by (relation "measure (\(x, y). size x + size y)") - (auto simp add: mkPinj_def split: pol.split) - -text {* Negation*} -primrec - neg :: "'a::{comm_ring} pol \ 'a pol" -where - "neg (Pc c) = Pc (-c)" - | "neg (Pinj i P) = Pinj i (neg P)" - | "neg (PX P x Q) = PX (neg P) x (neg Q)" - -text {* Substraction *} -definition - sub :: "'a::{comm_ring} pol \ 'a pol \ 'a pol" (infixl "\" 65) -where - "sub P Q = P \ neg Q" - -text {* Square for Fast Exponentation *} -primrec - sqr :: "'a::{comm_ring_1} pol \ 'a pol" -where - "sqr (Pc c) = Pc (c * c)" - | "sqr (Pinj i P) = mkPinj i (sqr P)" - | "sqr (PX A x B) = mkPX (sqr A) (x + x) (sqr B) \ - mkPX (Pc (1 + 1) \ A \ mkPinj 1 B) x (Pc 0)" - -text {* Fast Exponentation *} -fun - pow :: "nat \ 'a::{comm_ring_1} pol \ 'a pol" -where - "pow 0 P = Pc 1" - | "pow n P = (if even n then pow (n div 2) (sqr P) - else P \ pow (n div 2) (sqr P))" - -lemma pow_if: - "pow n P = - (if n = 0 then Pc 1 else if even n then pow (n div 2) (sqr P) - else P \ pow (n div 2) (sqr P))" - by (cases n) simp_all - - -text {* Normalization of polynomial expressions *} - -primrec - norm :: "'a::{comm_ring_1} polex \ 'a pol" -where - "norm (Pol P) = P" - | "norm (Add P Q) = norm P \ norm Q" - | "norm (Sub P Q) = norm P \ norm Q" - | "norm (Mul P Q) = norm P \ norm Q" - | "norm (Pow P n) = pow n (norm P)" - | "norm (Neg P) = neg (norm P)" - -text {* mkPinj preserve semantics *} -lemma mkPinj_ci: "Ipol l (mkPinj a B) = Ipol l (Pinj a B)" - by (induct B) (auto simp add: mkPinj_def algebra_simps) - -text {* mkPX preserves semantics *} -lemma mkPX_ci: "Ipol l (mkPX A b C) = Ipol l (PX A b C)" - by (cases A) (auto simp add: mkPX_def mkPinj_ci power_add algebra_simps) - -text {* Correctness theorems for the implemented operations *} - -text {* Negation *} -lemma neg_ci: "Ipol l (neg P) = -(Ipol l P)" - by (induct P arbitrary: l) auto - -text {* Addition *} -lemma add_ci: "Ipol l (P \ Q) = Ipol l P + Ipol l Q" -proof (induct P Q arbitrary: l rule: add.induct) - case (6 x P y Q) - show ?case - proof (rule linorder_cases) - assume "x < y" - with 6 show ?case by (simp add: mkPinj_ci algebra_simps) - next - assume "x = y" - with 6 show ?case by (simp add: mkPinj_ci) - next - assume "x > y" - with 6 show ?case by (simp add: mkPinj_ci algebra_simps) - qed -next - case (7 x P Q y R) - have "x = 0 \ x = 1 \ x > 1" by arith - moreover - { assume "x = 0" with 7 have ?case by simp } - moreover - { assume "x = 1" with 7 have ?case by (simp add: algebra_simps) } - moreover - { assume "x > 1" from 7 have ?case by (cases x) simp_all } - ultimately show ?case by blast -next - case (8 P x R y Q) - have "y = 0 \ y = 1 \ y > 1" by arith - moreover - { assume "y = 0" with 8 have ?case by simp } - moreover - { assume "y = 1" with 8 have ?case by simp } - moreover - { assume "y > 1" with 8 have ?case by simp } - ultimately show ?case by blast -next - case (9 P1 x P2 Q1 y Q2) - show ?case - proof (rule linorder_cases) - assume a: "x < y" hence "EX d. d + x = y" by arith - with 9 a show ?case by (auto simp add: mkPX_ci power_add algebra_simps) - next - assume a: "y < x" hence "EX d. d + y = x" by arith - with 9 a show ?case by (auto simp add: power_add mkPX_ci algebra_simps) - next - assume "x = y" - with 9 show ?case by (simp add: mkPX_ci algebra_simps) - qed -qed (auto simp add: algebra_simps) - -text {* Multiplication *} -lemma mul_ci: "Ipol l (P \ Q) = Ipol l P * Ipol l Q" - by (induct P Q arbitrary: l rule: mul.induct) - (simp_all add: mkPX_ci mkPinj_ci algebra_simps add_ci power_add) - -text {* Substraction *} -lemma sub_ci: "Ipol l (P \ Q) = Ipol l P - Ipol l Q" - by (simp add: add_ci neg_ci sub_def) - -text {* Square *} -lemma sqr_ci: "Ipol ls (sqr P) = Ipol ls P * Ipol ls P" - by (induct P arbitrary: ls) - (simp_all add: add_ci mkPinj_ci mkPX_ci mul_ci algebra_simps power_add) - -text {* Power *} -lemma even_pow:"even n \ pow n P = pow (n div 2) (sqr P)" - by (induct n) simp_all - -lemma pow_ci: "Ipol ls (pow n P) = Ipol ls P ^ n" -proof (induct n arbitrary: P rule: nat_less_induct) - case (1 k) - show ?case - proof (cases k) - case 0 - then show ?thesis by simp - next - case (Suc l) - show ?thesis - proof cases - assume "even l" - then have "Suc l div 2 = l div 2" - by (simp add: nat_number even_nat_plus_one_div_two) - moreover - from Suc have "l < k" by simp - with 1 have "\P. Ipol ls (pow l P) = Ipol ls P ^ l" by simp - moreover - note Suc `even l` even_nat_plus_one_div_two - ultimately show ?thesis by (auto simp add: mul_ci power_Suc even_pow) - next - assume "odd l" - { - fix p - have "Ipol ls (sqr P) ^ (Suc l div 2) = Ipol ls P ^ Suc l" - proof (cases l) - case 0 - with `odd l` show ?thesis by simp - next - case (Suc w) - with `odd l` have "even w" by simp - have two_times: "2 * (w div 2) = w" - by (simp only: numerals even_nat_div_two_times_two [OF `even w`]) - have "Ipol ls P * Ipol ls P = Ipol ls P ^ Suc (Suc 0)" - by (simp add: power_Suc) - then have "Ipol ls P * Ipol ls P = Ipol ls P ^ 2" - by (simp add: numerals) - with Suc show ?thesis - by (auto simp add: power_mult [symmetric, of _ 2 _] two_times mul_ci sqr_ci - simp del: power_Suc) - qed - } with 1 Suc `odd l` show ?thesis by simp - qed - qed -qed - -text {* Normalization preserves semantics *} -lemma norm_ci: "Ipolex l Pe = Ipol l (norm Pe)" - by (induct Pe) (simp_all add: add_ci sub_ci mul_ci neg_ci pow_ci) - -text {* Reflection lemma: Key to the (incomplete) decision procedure *} -lemma norm_eq: - assumes "norm P1 = norm P2" - shows "Ipolex l P1 = Ipolex l P2" -proof - - from prems have "Ipol l (norm P1) = Ipol l (norm P2)" by simp - then show ?thesis by (simp only: norm_ci) -qed - - -use "comm_ring.ML" -setup CommRing.setup - -end diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/Library/Efficient_Nat.thy --- a/src/HOL/Library/Efficient_Nat.thy Thu Oct 29 16:23:57 2009 +0100 +++ b/src/HOL/Library/Efficient_Nat.thy Wed Nov 04 11:40:59 2009 +0100 @@ -54,15 +54,15 @@ and @{term "op mod \ nat \ nat \ nat"} operations. *} definition divmod_aux :: "nat \ nat \ nat \ nat" where - [code del]: "divmod_aux = Divides.divmod" + [code del]: "divmod_aux = divmod_nat" lemma [code]: - "Divides.divmod n m = (if m = 0 then (0, n) else divmod_aux n m)" - unfolding divmod_aux_def divmod_div_mod by simp + "divmod_nat n m = (if m = 0 then (0, n) else divmod_aux n m)" + unfolding divmod_aux_def divmod_nat_div_mod by simp lemma divmod_aux_code [code]: "divmod_aux n m = (nat (of_nat n div of_nat m), nat (of_nat n mod of_nat m))" - unfolding divmod_aux_def divmod_div_mod zdiv_int [symmetric] zmod_int [symmetric] by simp + unfolding divmod_aux_def divmod_nat_div_mod zdiv_int [symmetric] zmod_int [symmetric] by simp lemma eq_nat_code [code]: "eq_class.eq n m \ eq_class.eq (of_nat n \ int) (of_nat m)" @@ -424,22 +424,13 @@ text {* Module names *} code_modulename SML - Nat Integer - Divides Integer - Ring_and_Field Integer - Efficient_Nat Integer + Efficient_Nat Arith code_modulename OCaml - Nat Integer - Divides Integer - Ring_and_Field Integer - Efficient_Nat Integer + Efficient_Nat Arith code_modulename Haskell - Nat Integer - Divides Integer - Ring_and_Field Integer - Efficient_Nat Integer + Efficient_Nat Arith hide const int diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/Library/Library.thy --- a/src/HOL/Library/Library.thy Thu Oct 29 16:23:57 2009 +0100 +++ b/src/HOL/Library/Library.thy Wed Nov 04 11:40:59 2009 +0100 @@ -11,7 +11,6 @@ Code_Char_chr Code_Integer Coinductive_List - Commutative_Ring Continuity ContNotDenum Countable diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/Library/Numeral_Type.thy --- a/src/HOL/Library/Numeral_Type.thy Thu Oct 29 16:23:57 2009 +0100 +++ b/src/HOL/Library/Numeral_Type.thy Wed Nov 04 11:40:59 2009 +0100 @@ -188,7 +188,7 @@ by (rule type_definition.Abs_inverse [OF type]) lemma Rep_Abs_mod: "Rep (Abs (m mod n)) = m mod n" -by (simp add: Abs_inverse IntDiv.pos_mod_conj [OF size0]) +by (simp add: Abs_inverse pos_mod_conj [OF size0]) lemma Rep_Abs_0: "Rep (Abs 0) = 0" by (simp add: Abs_inverse size0) diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/Library/OptionalSugar.thy --- a/src/HOL/Library/OptionalSugar.thy Thu Oct 29 16:23:57 2009 +0100 +++ b/src/HOL/Library/OptionalSugar.thy Wed Nov 04 11:40:59 2009 +0100 @@ -50,7 +50,7 @@ (* type constraints with spacing *) setup {* let - val typ = SimpleSyntax.read_typ; + val typ = Simple_Syntax.read_typ; val typeT = Syntax.typeT; val spropT = Syntax.spropT; in diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/Library/Sublist_Order.thy --- a/src/HOL/Library/Sublist_Order.thy Thu Oct 29 16:23:57 2009 +0100 +++ b/src/HOL/Library/Sublist_Order.thy Wed Nov 04 11:40:59 2009 +0100 @@ -1,6 +1,6 @@ (* Title: HOL/Library/Sublist_Order.thy Authors: Peter Lammich, Uni Muenster - Florian Haftmann, TU Muenchen + Florian Haftmann, Tobias Nipkow, TU Muenchen *) header {* Sublist Ordering *} @@ -17,7 +17,7 @@ subsection {* Definitions and basic lemmas *} -instantiation list :: (type) order +instantiation list :: (type) ord begin inductive less_eq_list where @@ -25,162 +25,237 @@ | drop: "ys \ xs \ ys \ x # xs" | take: "ys \ xs \ x # ys \ x # xs" -lemmas ileq_empty = empty -lemmas ileq_drop = drop -lemmas ileq_take = take - -lemma ileq_cases [cases set, case_names empty drop take]: - assumes "xs \ ys" - and "xs = [] \ P" - and "\z zs. ys = z # zs \ xs \ zs \ P" - and "\x zs ws. xs = x # zs \ ys = x # ws \ zs \ ws \ P" - shows P - using assms by (blast elim: less_eq_list.cases) - -lemma ileq_induct [induct set, case_names empty drop take]: - assumes "xs \ ys" - and "\zs. P [] zs" - and "\z zs ws. ws \ zs \ P ws zs \ P ws (z # zs)" - and "\z zs ws. ws \ zs \ P ws zs \ P (z # ws) (z # zs)" - shows "P xs ys" - using assms by (induct rule: less_eq_list.induct) blast+ - definition [code del]: "(xs \ 'a list) < ys \ xs \ ys \ \ ys \ xs" -lemma ileq_length: "xs \ ys \ length xs \ length ys" - by (induct rule: ileq_induct) auto -lemma ileq_below_empty [simp]: "xs \ [] \ xs = []" - by (auto dest: ileq_length) +instance proof qed + +end + +lemma le_list_length: "xs \ ys \ length xs \ length ys" +by (induct rule: less_eq_list.induct) auto + +lemma le_list_same_length: "xs \ ys \ length xs = length ys \ xs = ys" +by (induct rule: less_eq_list.induct) (auto dest: le_list_length) + +lemma not_le_list_length[simp]: "length ys < length xs \ ~ xs <= ys" +by (metis le_list_length linorder_not_less) + +lemma le_list_below_empty [simp]: "xs \ [] \ xs = []" +by (auto dest: le_list_length) + +lemma le_list_drop_many: "xs \ ys \ xs \ zs @ ys" +by (induct zs) (auto intro: drop) + +lemma [code]: "[] <= xs \ True" +by(metis less_eq_list.empty) + +lemma [code]: "(x#xs) <= [] \ False" +by simp + +lemma le_list_drop_Cons: assumes "x#xs <= ys" shows "xs <= ys" +proof- + { fix xs' ys' + assume "xs' <= ys" + hence "ALL x xs. xs' = x#xs \ xs <= ys" + proof induct + case empty thus ?case by simp + next + case drop thus ?case by (metis less_eq_list.drop) + next + case take thus ?case by (simp add: drop) + qed } + from this[OF assms] show ?thesis by simp +qed + +lemma le_list_drop_Cons2: +assumes "x#xs <= x#ys" shows "xs <= ys" +using assms +proof cases + case drop thus ?thesis by (metis le_list_drop_Cons list.inject) +qed simp_all + +lemma le_list_drop_Cons_neq: assumes "x # xs <= y # ys" +shows "x ~= y \ x # xs <= ys" +using assms proof cases qed auto + +lemma le_list_Cons2_iff[simp,code]: "(x#xs) <= (y#ys) \ + (if x=y then xs <= ys else (x#xs) <= ys)" +by (metis drop take le_list_drop_Cons2 le_list_drop_Cons_neq) + +lemma le_list_take_many_iff: "zs @ xs \ zs @ ys \ xs \ ys" +by (induct zs) (auto intro: take) + +lemma le_list_Cons_EX: + assumes "x # ys <= zs" shows "EX us vs. zs = us @ x # vs & ys <= vs" +proof- + { fix xys zs :: "'a list" assume "xys <= zs" + hence "ALL x ys. xys = x#ys \ (EX us vs. zs = us @ x # vs & ys <= vs)" + proof induct + case empty show ?case by simp + next + case take thus ?case by (metis list.inject self_append_conv2) + next + case drop thus ?case by (metis append_eq_Cons_conv) + qed + } with assms show ?thesis by blast +qed + +instantiation list :: (type) order +begin instance proof fix xs ys :: "'a list" show "xs < ys \ xs \ ys \ \ ys \ xs" unfolding less_list_def .. next fix xs :: "'a list" - show "xs \ xs" by (induct xs) (auto intro!: ileq_empty ileq_drop ileq_take) + show "xs \ xs" by (induct xs) (auto intro!: less_eq_list.drop) next fix xs ys :: "'a list" - (* TODO: Is there a simpler proof ? *) - { fix n - have "!!l l'. \l\l'; l'\l; n=length l + length l'\ \ l=l'" - proof (induct n rule: nat_less_induct) - case (1 n l l') from "1.prems"(1) show ?case proof (cases rule: ileq_cases) - case empty with "1.prems"(2) show ?thesis by auto - next - case (drop a l2') with "1.prems"(2) have "length l'\length l" "length l \ length l2'" "1+length l2' = length l'" by (auto dest: ileq_length) - hence False by simp thus ?thesis .. - next - case (take a l1' l2') hence LEN': "length l1' + length l2' < length l + length l'" by simp - from "1.prems" have LEN: "length l' = length l" by (auto dest!: ileq_length) - from "1.prems"(2) show ?thesis proof (cases rule: ileq_cases[case_names empty' drop' take']) - case empty' with take LEN show ?thesis by simp - next - case (drop' ah l2h) with take LEN have "length l1' \ length l2h" "1+length l2h = length l2'" "length l2' = length l1'" by (auto dest: ileq_length) - hence False by simp thus ?thesis .. - next - case (take' ah l1h l2h) - with take have 2: "ah=a" "l1h=l2'" "l2h=l1'" "l1' \ l2'" "l2' \ l1'" by auto - with LEN' "1.hyps" "1.prems"(3) have "l1'=l2'" by blast - with take 2 show ?thesis by simp - qed - qed - qed - } - moreover assume "xs \ ys" "ys \ xs" + assume "xs <= ys" + hence "ys <= xs \ xs = ys" + proof induct + case empty show ?case by simp + next + case take thus ?case by simp + next + case drop thus ?case + by(metis le_list_drop_Cons le_list_length Suc_length_conv Suc_n_not_le_n) + qed + moreover assume "ys <= xs" ultimately show "xs = ys" by blast next fix xs ys zs :: "'a list" - { - fix n - have "!!x y z. \x \ y; y \ z; n=length x + length y + length z\ \ x \ z" - proof (induct rule: nat_less_induct[case_names I]) - case (I n x y z) - from I.prems(2) show ?case proof (cases rule: ileq_cases) - case empty with I.prems(1) show ?thesis by auto - next - case (drop a z') hence "length x + length y + length z' < length x + length y + length z" by simp - with I.hyps I.prems(3,1) drop(2) have "x\z'" by blast - with drop(1) show ?thesis by (auto intro: ileq_drop) - next - case (take a y' z') from I.prems(1) show ?thesis proof (cases rule: ileq_cases[case_names empty' drop' take']) - case empty' thus ?thesis by auto - next - case (drop' ah y'h) with take have "x\y'" "y'\z'" "length x + length y' + length z' < length x + length y + length z" by auto - with I.hyps I.prems(3) have "x\z'" by (blast) - with take(2) show ?thesis by (auto intro: ileq_drop) - next - case (take' ah x' y'h) with take have 2: "x=a#x'" "x'\y'" "y'\z'" "length x' + length y' + length z' < length x + length y + length z" by auto - with I.hyps I.prems(3) have "x'\z'" by blast - with 2 take(2) show ?thesis by (auto intro: ileq_take) - qed - qed + assume "xs <= ys" + hence "ys <= zs \ xs <= zs" + proof (induct arbitrary:zs) + case empty show ?case by simp + next + case (take xs ys x) show ?case + proof + assume "x # ys <= zs" + with take show "x # xs <= zs" + by(metis le_list_Cons_EX le_list_drop_many less_eq_list.take local.take(2)) qed - } - moreover assume "xs \ ys" "ys \ zs" - ultimately show "xs \ zs" by blast + next + case drop thus ?case by (metis le_list_drop_Cons) + qed + moreover assume "ys <= zs" + ultimately show "xs <= zs" by blast qed end -lemmas ileq_intros = ileq_empty ileq_drop ileq_take +lemma le_list_append_le_same_iff: "xs @ ys <= ys \ xs=[]" +by (auto dest: le_list_length) -lemma ileq_drop_many: "xs \ ys \ xs \ zs @ ys" - by (induct zs) (auto intro: ileq_drop) -lemma ileq_take_many: "xs \ ys \ zs @ xs \ zs @ ys" - by (induct zs) (auto intro: ileq_take) +lemma le_list_append_mono: "\ xs <= xs'; ys <= ys' \ \ xs@ys <= xs'@ys'" +apply (induct rule:less_eq_list.induct) + apply (metis eq_Nil_appendI le_list_drop_many) + apply (metis Cons_eq_append_conv le_list_drop_Cons order_eq_refl order_trans) +apply simp +done -lemma ileq_same_length: "xs \ ys \ length xs = length ys \ xs = ys" - by (induct rule: ileq_induct) (auto dest: ileq_length) -lemma ileq_same_append [simp]: "x # xs \ xs \ False" - by (auto dest: ileq_length) +lemma less_list_length: "xs < ys \ length xs < length ys" +by (metis le_list_length le_list_same_length le_neq_implies_less less_list_def) -lemma ilt_length [intro]: - assumes "xs < ys" - shows "length xs < length ys" -proof - - from assms have "xs \ ys" and "xs \ ys" by (simp_all add: less_le) - moreover with ileq_length have "length xs \ length ys" by auto - ultimately show ?thesis by (auto intro: ileq_same_length) -qed +lemma less_list_empty [simp]: "[] < xs \ xs \ []" +by (metis empty order_less_le) + +lemma less_list_below_empty[simp]: "xs < [] \ False" +by (metis empty less_list_def) + +lemma less_list_drop: "xs < ys \ xs < x # ys" +by (unfold less_le) (auto intro: less_eq_list.drop) -lemma ilt_empty [simp]: "[] < xs \ xs \ []" - by (unfold less_list_def, auto) -lemma ilt_emptyI: "xs \ [] \ [] < xs" - by (unfold less_list_def, auto) -lemma ilt_emptyD: "[] < xs \ xs \ []" - by (unfold less_list_def, auto) -lemma ilt_below_empty[simp]: "xs < [] \ False" - by (auto dest: ilt_length) -lemma ilt_drop: "xs < ys \ xs < x # ys" - by (unfold less_le) (auto intro: ileq_intros) -lemma ilt_take: "xs < ys \ x # xs < x # ys" - by (unfold less_le) (auto intro: ileq_intros) -lemma ilt_drop_many: "xs < ys \ xs < zs @ ys" - by (induct zs) (auto intro: ilt_drop) -lemma ilt_take_many: "xs < ys \ zs @ xs < zs @ ys" - by (induct zs) (auto intro: ilt_take) +lemma less_list_take_iff: "x # xs < x # ys \ xs < ys" +by (metis le_list_Cons2_iff less_list_def) + +lemma less_list_drop_many: "xs < ys \ xs < zs @ ys" +by(metis le_list_append_le_same_iff le_list_drop_many order_less_le self_append_conv2) + +lemma less_list_take_many_iff: "zs @ xs < zs @ ys \ xs < ys" +by (metis le_list_take_many_iff less_list_def) subsection {* Appending elements *} -lemma ileq_rev_take: "xs \ ys \ xs @ [x] \ ys @ [x]" - by (induct rule: ileq_induct) (auto intro: ileq_intros ileq_drop_many) -lemma ilt_rev_take: "xs < ys \ xs @ [x] < ys @ [x]" - by (unfold less_le) (auto dest: ileq_rev_take) -lemma ileq_rev_drop: "xs \ ys \ xs \ ys @ [x]" - by (induct rule: ileq_induct) (auto intro: ileq_intros) -lemma ileq_rev_drop_many: "xs \ ys \ xs \ ys @ zs" - by (induct zs rule: rev_induct) (auto dest: ileq_rev_drop) +lemma le_list_rev_take_iff[simp]: "xs @ zs \ ys @ zs \ xs \ ys" (is "?L = ?R") +proof + { fix xs' ys' xs ys zs :: "'a list" assume "xs' <= ys'" + hence "xs' = xs @ zs & ys' = ys @ zs \ xs <= ys" + proof (induct arbitrary: xs ys zs) + case empty show ?case by simp + next + case (drop xs' ys' x) + { assume "ys=[]" hence ?case using drop(1) by auto } + moreover + { fix us assume "ys = x#us" + hence ?case using drop(2) by(simp add: less_eq_list.drop) } + ultimately show ?case by (auto simp:Cons_eq_append_conv) + next + case (take xs' ys' x) + { assume "xs=[]" hence ?case using take(1) by auto } + moreover + { fix us vs assume "xs=x#us" "ys=x#vs" hence ?case using take(2) by auto} + moreover + { fix us assume "xs=x#us" "ys=[]" hence ?case using take(2) by bestsimp } + ultimately show ?case by (auto simp:Cons_eq_append_conv) + qed } + moreover assume ?L + ultimately show ?R by blast +next + assume ?R thus ?L by(metis le_list_append_mono order_refl) +qed + +lemma less_list_rev_take: "xs @ zs < ys @ zs \ xs < ys" +by (unfold less_le) auto + +lemma le_list_rev_drop_many: "xs \ ys \ xs \ ys @ zs" +by (metis append_Nil2 empty le_list_append_mono) subsection {* Relation to standard list operations *} -lemma ileq_map: "xs \ ys \ map f xs \ map f ys" - by (induct rule: ileq_induct) (auto intro: ileq_intros) -lemma ileq_filter_left[simp]: "filter f xs \ xs" - by (induct xs) (auto intro: ileq_intros) -lemma ileq_filter: "xs \ ys \ filter f xs \ filter f ys" - by (induct rule: ileq_induct) (auto intro: ileq_intros) +lemma le_list_map: "xs \ ys \ map f xs \ map f ys" +by (induct rule: less_eq_list.induct) (auto intro: less_eq_list.drop) + +lemma le_list_filter_left[simp]: "filter f xs \ xs" +by (induct xs) (auto intro: less_eq_list.drop) + +lemma le_list_filter: "xs \ ys \ filter f xs \ filter f ys" +by (induct rule: less_eq_list.induct) (auto intro: less_eq_list.drop) + + +lemma "xs <= ys \ (EX N. xs = sublist ys N)" (is "?L = ?R") +proof + assume ?L + thus ?R + proof induct + case empty show ?case by (metis sublist_empty) + next + case (drop xs ys x) + then obtain N where "xs = sublist ys N" by blast + hence "xs = sublist (x#ys) (Suc ` N)" + by (clarsimp simp add:sublist_Cons inj_image_mem_iff) + thus ?case by blast + next + case (take xs ys x) + then obtain N where "xs = sublist ys N" by blast + hence "x#xs = sublist (x#ys) (insert 0 (Suc ` N))" + by (clarsimp simp add:sublist_Cons inj_image_mem_iff) + thus ?case by blast + qed +next + assume ?R + then obtain N where "xs = sublist ys N" .. + moreover have "sublist ys N <= ys" + proof (induct ys arbitrary:N) + case Nil show ?case by simp + next + case Cons thus ?case by (auto simp add:sublist_Cons drop) + qed + ultimately show ?L by simp +qed end diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/Library/Sum_Of_Squares/positivstellensatz_tools.ML --- a/src/HOL/Library/Sum_Of_Squares/positivstellensatz_tools.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/HOL/Library/Sum_Of_Squares/positivstellensatz_tools.ML Wed Nov 04 11:40:59 2009 +0100 @@ -100,7 +100,7 @@ (fn (x, k) => (cterm_of (ProofContext.theory_of ctxt) (Free (x, @{typ real})), k)) fun parse_monomial ctxt = repeat_sep "*" (parse_varpow ctxt) >> - List.foldl (uncurry FuncUtil.Ctermfunc.update) FuncUtil.Ctermfunc.empty + (fn xs => fold FuncUtil.Ctermfunc.update xs FuncUtil.Ctermfunc.empty) fun parse_cmonomial ctxt = rat_int --| str "*" -- (parse_monomial ctxt) >> swap || @@ -108,7 +108,7 @@ rat_int >> (fn r => (FuncUtil.Ctermfunc.empty, r)) fun parse_poly ctxt = repeat_sep "+" (parse_cmonomial ctxt) >> - List.foldl (uncurry FuncUtil.Monomialfunc.update) FuncUtil.Monomialfunc.empty + (fn xs => fold FuncUtil.Monomialfunc.update xs FuncUtil.Monomialfunc.empty) (* positivstellensatz parser *) diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/Library/Word.thy --- a/src/HOL/Library/Word.thy Thu Oct 29 16:23:57 2009 +0100 +++ b/src/HOL/Library/Word.thy Wed Nov 04 11:40:59 2009 +0100 @@ -5,7 +5,7 @@ header {* Binary Words *} theory Word -imports "~~/src/HOL/Main" +imports Main begin subsection {* Auxilary Lemmas *} @@ -561,35 +561,17 @@ shows "nat_to_bv (2 * bv_to_nat w + bitval x) = norm_unsigned (w @ [x])" proof (cases x) assume [simp]: "x = \" - show ?thesis - apply (simp add: nat_to_bv_non0) - apply safe - proof - - fix q - assume "Suc (2 * bv_to_nat w) = 2 * q" - hence orig: "(2 * bv_to_nat w + 1) mod 2 = 2 * q mod 2" (is "?lhs = ?rhs") - by simp - have "?lhs = (1 + 2 * bv_to_nat w) mod 2" - by (simp add: add_commute) - also have "... = 1" - by (subst mod_add_eq) simp - finally have eq1: "?lhs = 1" . - have "?rhs = 0" by simp - with orig and eq1 - show "nat_to_bv (Suc (2 * bv_to_nat w) div 2) @ [\] = norm_unsigned (w @ [\])" - by simp - next - have "nat_to_bv (Suc (2 * bv_to_nat w) div 2) @ [\] = - nat_to_bv ((1 + 2 * bv_to_nat w) div 2) @ [\]" - by (simp add: add_commute) - also have "... = nat_to_bv (bv_to_nat w) @ [\]" - by (subst div_add1_eq) simp - also have "... = norm_unsigned w @ [\]" - by (subst ass) (rule refl) - also have "... = norm_unsigned (w @ [\])" - by (cases "norm_unsigned w") simp_all - finally show "nat_to_bv (Suc (2 * bv_to_nat w) div 2) @ [\] = norm_unsigned (w @ [\])" . - qed + have "nat_to_bv (Suc (2 * bv_to_nat w) div 2) @ [\] = + nat_to_bv ((1 + 2 * bv_to_nat w) div 2) @ [\]" + by (simp add: add_commute) + also have "... = nat_to_bv (bv_to_nat w) @ [\]" + by (subst div_add1_eq) simp + also have "... = norm_unsigned w @ [\]" + by (subst ass) (rule refl) + also have "... = norm_unsigned (w @ [\])" + by (cases "norm_unsigned w") simp_all + finally have "nat_to_bv (Suc (2 * bv_to_nat w) div 2) @ [\] = norm_unsigned (w @ [\])" . + then show ?thesis by (simp add: nat_to_bv_non0) next assume [simp]: "x = \" show ?thesis diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/Library/comm_ring.ML --- a/src/HOL/Library/comm_ring.ML Thu Oct 29 16:23:57 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,106 +0,0 @@ -(* Author: Amine Chaieb - -Tactic for solving equalities over commutative rings. -*) - -signature COMM_RING = -sig - val comm_ring_tac : Proof.context -> int -> tactic - val setup : theory -> theory -end - -structure CommRing: COMM_RING = -struct - -(* Zero and One of the commutative ring *) -fun cring_zero T = Const (@{const_name HOL.zero}, T); -fun cring_one T = Const (@{const_name HOL.one}, T); - -(* reification functions *) -(* add two polynom expressions *) -fun polT t = Type ("Commutative_Ring.pol", [t]); -fun polexT t = Type ("Commutative_Ring.polex", [t]); - -(* pol *) -fun pol_Pc t = Const ("Commutative_Ring.pol.Pc", t --> polT t); -fun pol_Pinj t = Const ("Commutative_Ring.pol.Pinj", HOLogic.natT --> polT t --> polT t); -fun pol_PX t = Const ("Commutative_Ring.pol.PX", polT t --> HOLogic.natT --> polT t --> polT t); - -(* polex *) -fun polex_add t = Const ("Commutative_Ring.polex.Add", polexT t --> polexT t --> polexT t); -fun polex_sub t = Const ("Commutative_Ring.polex.Sub", polexT t --> polexT t --> polexT t); -fun polex_mul t = Const ("Commutative_Ring.polex.Mul", polexT t --> polexT t --> polexT t); -fun polex_neg t = Const ("Commutative_Ring.polex.Neg", polexT t --> polexT t); -fun polex_pol t = Const ("Commutative_Ring.polex.Pol", polT t --> polexT t); -fun polex_pow t = Const ("Commutative_Ring.polex.Pow", polexT t --> HOLogic.natT --> polexT t); - -(* reification of polynoms : primitive cring expressions *) -fun reif_pol T vs (t as Free _) = - let - val one = @{term "1::nat"}; - val i = find_index (fn t' => t' = t) vs - in if i = 0 - then pol_PX T $ (pol_Pc T $ cring_one T) - $ one $ (pol_Pc T $ cring_zero T) - else pol_Pinj T $ HOLogic.mk_nat i - $ (pol_PX T $ (pol_Pc T $ cring_one T) - $ one $ (pol_Pc T $ cring_zero T)) - end - | reif_pol T vs t = pol_Pc T $ t; - -(* reification of polynom expressions *) -fun reif_polex T vs (Const (@{const_name HOL.plus}, _) $ a $ b) = - polex_add T $ reif_polex T vs a $ reif_polex T vs b - | reif_polex T vs (Const (@{const_name HOL.minus}, _) $ a $ b) = - polex_sub T $ reif_polex T vs a $ reif_polex T vs b - | reif_polex T vs (Const (@{const_name HOL.times}, _) $ a $ b) = - polex_mul T $ reif_polex T vs a $ reif_polex T vs b - | reif_polex T vs (Const (@{const_name HOL.uminus}, _) $ a) = - polex_neg T $ reif_polex T vs a - | reif_polex T vs (Const (@{const_name Power.power}, _) $ a $ n) = - polex_pow T $ reif_polex T vs a $ n - | reif_polex T vs t = polex_pol T $ reif_pol T vs t; - -(* reification of the equation *) -val cr_sort = @{sort "comm_ring_1"}; - -fun reif_eq thy (eq as Const("op =", Type("fun", [T, _])) $ lhs $ rhs) = - if Sign.of_sort thy (T, cr_sort) then - let - val fs = OldTerm.term_frees eq; - val cvs = cterm_of thy (HOLogic.mk_list T fs); - val clhs = cterm_of thy (reif_polex T fs lhs); - val crhs = cterm_of thy (reif_polex T fs rhs); - val ca = ctyp_of thy T; - in (ca, cvs, clhs, crhs) end - else error ("reif_eq: not an equation over " ^ Syntax.string_of_sort_global thy cr_sort) - | reif_eq _ _ = error "reif_eq: not an equation"; - -(* The cring tactic *) -(* Attention: You have to make sure that no t^0 is in the goal!! *) -(* Use simply rewriting t^0 = 1 *) -val cring_simps = - [@{thm mkPX_def}, @{thm mkPinj_def}, @{thm sub_def}, @{thm power_add}, - @{thm even_def}, @{thm pow_if}, sym OF [@{thm power_add}]]; - -fun comm_ring_tac ctxt = SUBGOAL (fn (g, i) => - let - val thy = ProofContext.theory_of ctxt; - val cring_ss = Simplifier.simpset_of ctxt (*FIXME really the full simpset!?*) - addsimps cring_simps; - val (ca, cvs, clhs, crhs) = reif_eq thy (HOLogic.dest_Trueprop g) - val norm_eq_th = - simplify cring_ss (instantiate' [SOME ca] [SOME clhs, SOME crhs, SOME cvs] @{thm norm_eq}) - in - cut_rules_tac [norm_eq_th] i - THEN (simp_tac cring_ss i) - THEN (simp_tac cring_ss i) - end); - -val setup = - Method.setup @{binding comm_ring} (Scan.succeed (SIMPLE_METHOD' o comm_ring_tac)) - "reflective decision procedure for equalities over commutative rings" #> - Method.setup @{binding algebra} (Scan.succeed (SIMPLE_METHOD' o comm_ring_tac)) - "method for proving algebraic properties (same as comm_ring)"; - -end; diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/Matrix/ComputeNumeral.thy --- a/src/HOL/Matrix/ComputeNumeral.thy Thu Oct 29 16:23:57 2009 +0100 +++ b/src/HOL/Matrix/ComputeNumeral.thy Wed Nov 04 11:40:59 2009 +0100 @@ -153,16 +153,16 @@ lemma negateSnd: "negateSnd (q, r) = (q, -r)" by (simp add: negateSnd_def) -lemma divmod: "IntDiv.divmod a b = (if 0\a then +lemma divmod: "divmod_int a b = (if 0\a then if 0\b then posDivAlg a b else if a=0 then (0, 0) else negateSnd (negDivAlg (-a) (-b)) else if 0 ("trivial", SH_OK (0, 0, [])) + handle Res_HOL_Clause.TOO_TRIVIAL => ("trivial", SH_OK (0, 0, [])) | ERROR msg => ("error: " ^ msg, SH_ERROR) | TimeLimit.TimeOut => ("timeout", SH_ERROR) diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy --- a/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy Thu Oct 29 16:23:57 2009 +0100 +++ b/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy Wed Nov 04 11:40:59 2009 +0100 @@ -507,7 +507,7 @@ shows "x islimpt S \ (\e>0. \x'\ S. x' \ x \ dist x' x <= e)" unfolding islimpt_approachable using approachable_lt_le[where f="\x'. dist x' x" and P="\x'. \ (x'\S \ x'\x)"] - by metis (* FIXME: VERY slow! *) + by metis class perfect_space = (* FIXME: perfect_space should inherit from topological_space *) @@ -746,7 +746,7 @@ { fix x have "x \ interior S \ x \ UNIV - (closure (UNIV - S))" unfolding interior_def closure_def islimpt_def - by blast (* FIXME: VERY slow! *) + by auto } thus ?thesis by blast @@ -911,7 +911,7 @@ hence ?rhse using `?lhs` unfolding frontier_closures closure_def islimpt_def using open_ball[of a e] `e > 0` - by (auto, erule_tac x = "ball a e" in allE, auto) (* FIXME: VERY slow! *) + by simp (metis centre_in_ball mem_ball open_ball) } ultimately have ?rhse by auto } @@ -1052,10 +1052,7 @@ lemma eventually_within_le: "eventually P (at a within S) \ (\d>0. \x\S. 0 < dist x a \ dist x a <= d \ P x)" (is "?lhs = ?rhs") unfolding eventually_within -apply safe -apply (rule_tac x="d/2" in exI, simp) -apply (rule_tac x="d" in exI, simp) -done +by auto (metis Rats_dense_in_nn_real order_le_less_trans order_refl) lemma eventually_happens: "eventually P net ==> trivial_limit net \ (\x. P x)" unfolding eventually_def trivial_limit_def @@ -4404,28 +4401,20 @@ hence "norm (x - y) \ 2 * a" using norm_triangle_ineq[of x "-y", unfolded norm_minus_cancel] a[THEN bspec[where x=x]] a[THEN bspec[where x=y]] by (auto simp add: ring_simps) } note * = this { fix x y assume "x\s" "y\s" hence "s \ {}" by auto - have lub:"isLub UNIV ?D (Sup ?D)" using * Sup[of ?D] using `s\{}` unfolding setle_def - apply auto (*FIXME: something horrible has happened here!*) - apply atomize - apply safe - apply metis + - done - have "norm(x - y) \ diameter s" unfolding diameter_def using `s\{}` *[OF `x\s` `y\s`] `x\s` `y\s` isLubD1[OF lub] unfolding setle_def by auto } + have "norm(x - y) \ diameter s" unfolding diameter_def using `s\{}` *[OF `x\s` `y\s`] `x\s` `y\s` + by simp (blast intro!: Sup_upper *) } moreover { fix d::real assume "d>0" "d < diameter s" hence "s\{}" unfolding diameter_def by auto - hence lub:"isLub UNIV ?D (Sup ?D)" using * Sup[of ?D] unfolding setle_def - apply auto (*FIXME: something horrible has happened here!*) - apply atomize - apply safe - apply metis + - done have "\d' \ ?D. d' > d" proof(rule ccontr) assume "\ (\d'\{norm (x - y) |x y. x \ s \ y \ s}. d < d')" - hence as:"\d'\?D. d' \ d" apply auto apply(erule_tac x="norm (x - y)" in allE) by auto - hence "isUb UNIV ?D d" unfolding isUb_def unfolding setle_def by auto - thus False using `d < diameter s` `s\{}` isLub_le_isUb[OF lub, of d] unfolding diameter_def by auto + hence "\d'\?D. d' \ d" by auto (metis not_leE) + thus False using `d < diameter s` `s\{}` + apply (auto simp add: diameter_def) + apply (drule Sup_real_iff [THEN [2] rev_iffD2]) + apply (auto, force) + done qed hence "\x\s. \y\s. norm(x - y) > d" by auto } ultimately show "\x\s. \y\s. norm(x - y) \ diameter s" @@ -4445,7 +4434,8 @@ then obtain x y where xys:"x\s" "y\s" and xy:"\u\s. \v\s. norm (u - v) \ norm (x - y)" using compact_sup_maxdistance[OF assms] by auto hence "diameter s \ norm (x - y)" by (force simp add: diameter_def intro!: Sup_least) - thus ?thesis using diameter_bounded(1)[OF b, THEN bspec[where x=x], THEN bspec[where x=y], OF xys] and xys by auto + thus ?thesis + by (metis b diameter_bounded_bound order_antisym xys) qed text{* Related results with closure as the conclusion. *} @@ -5291,7 +5281,7 @@ "s homeomorphic t \ t homeomorphic s" unfolding homeomorphic_def unfolding homeomorphism_def -by blast (* FIXME: slow *) +by blast lemma homeomorphic_trans: assumes "s homeomorphic t" "t homeomorphic u" shows "s homeomorphic u" @@ -5459,7 +5449,7 @@ shows "complete(f ` s)" proof- { fix g assume as:"\n::nat. g n \ f ` s" and cfg:"Cauchy g" - then obtain x where "\n. x n \ s \ g n = f (x n)" unfolding image_iff and Bex_def + then obtain x where "\n. x n \ s \ g n = f (x n)" using choice[of "\ n xa. xa \ s \ g n = f xa"] by auto hence x:"\n. x n \ s" "\n. g n = f (x n)" by auto hence "f \ x = g" unfolding expand_fun_eq by auto diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/Nat.thy --- a/src/HOL/Nat.thy Thu Oct 29 16:23:57 2009 +0100 +++ b/src/HOL/Nat.thy Wed Nov 04 11:40:59 2009 +0100 @@ -1674,4 +1674,16 @@ class size = fixes size :: "'a \ nat" -- {* see further theory @{text Wellfounded} *} + +subsection {* code module namespace *} + +code_modulename SML + Nat Arith + +code_modulename OCaml + Nat Arith + +code_modulename Haskell + Nat Arith + end diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/Nat_Numeral.thy --- a/src/HOL/Nat_Numeral.thy Thu Oct 29 16:23:57 2009 +0100 +++ b/src/HOL/Nat_Numeral.thy Wed Nov 04 11:40:59 2009 +0100 @@ -420,10 +420,6 @@ subsubsection{*Equals (=) *} -lemma eq_nat_nat_iff: - "[| (0::int) <= z; 0 <= z' |] ==> (nat z = nat z') = (z=z')" -by (auto elim!: nonneg_eq_int) - lemma eq_nat_number_of [simp]: "((number_of v :: nat) = number_of v') = (if neg (number_of v :: int) then (number_of v' :: int) \ 0 @@ -625,11 +621,6 @@ subsection{*Literal arithmetic involving powers*} -lemma nat_power_eq: "(0::int) <= z ==> nat (z^n) = nat z ^ n" -apply (induct "n") -apply (simp_all (no_asm_simp) add: nat_mult_distrib) -done - lemma power_nat_number_of: "(number_of v :: nat) ^ n = (if neg (number_of v :: int) then 0^n else nat ((number_of v :: int) ^ n))" diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/Nat_Transfer.thy --- a/src/HOL/Nat_Transfer.thy Thu Oct 29 16:23:57 2009 +0100 +++ b/src/HOL/Nat_Transfer.thy Wed Nov 04 11:40:59 2009 +0100 @@ -67,8 +67,7 @@ "(2::int) >= 0" "(3::int) >= 0" "int z >= 0" - apply (auto simp add: zero_le_mult_iff tsub_def) -done + by (auto simp add: zero_le_mult_iff tsub_def) lemma transfer_nat_int_relations: "x >= 0 \ y >= 0 \ diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/Nominal/nominal_datatype.ML --- a/src/HOL/Nominal/nominal_datatype.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/HOL/Nominal/nominal_datatype.ML Wed Nov 04 11:40:59 2009 +0100 @@ -62,11 +62,11 @@ in -fun mk_case_names_induct descr = RuleCases.case_names (induct_cases descr); +fun mk_case_names_induct descr = Rule_Cases.case_names (induct_cases descr); fun mk_case_names_exhausts descr new = - map (RuleCases.case_names o exhaust_cases descr o #1) - (List.filter (fn ((_, (name, _, _))) => name mem_string new) descr); + map (Rule_Cases.case_names o exhaust_cases descr o #1) + (filter (fn ((_, (name, _, _))) => name mem_string new) descr); end; @@ -153,7 +153,7 @@ fun projections rule = Project_Rule.projections (ProofContext.init (Thm.theory_of_thm rule)) rule - |> map (Drule.standard #> RuleCases.save rule); + |> map (Drule.standard #> Rule_Cases.save rule); val supp_prod = thm "supp_prod"; val fresh_prod = thm "fresh_prod"; @@ -526,7 +526,7 @@ fun make_intr s T (cname, cargs) = let - fun mk_prem (dt, (j, j', prems, ts)) = + fun mk_prem dt (j, j', prems, ts) = let val (dts, dt') = strip_option dt; val (dts', dt'') = strip_dtyp dt'; @@ -535,7 +535,7 @@ val T = typ_of_dtyp descr sorts dt''; val free = mk_Free "x" (Us ---> T) j; val free' = app_bnds free (length Us); - fun mk_abs_fun (T, (i, t)) = + fun mk_abs_fun T (i, t) = let val U = fastype_of t in (i + 1, Const ("Nominal.abs_fun", [T, U, T] ---> Type ("Nominal.noption", [U])) $ mk_Free "y" T i $ t) @@ -546,10 +546,10 @@ HOLogic.mk_Trueprop (Free (List.nth (rep_set_names, k), T --> HOLogic.boolT) $ free')) :: prems | _ => prems, - snd (List.foldr mk_abs_fun (j', free) Ts) :: ts) + snd (fold_rev mk_abs_fun Ts (j', free)) :: ts) end; - val (_, _, prems, ts) = List.foldr mk_prem (1, 1, [], []) cargs; + val (_, _, prems, ts) = fold_rev mk_prem cargs (1, 1, [], []); val concl = HOLogic.mk_Trueprop (Free (s, T --> HOLogic.boolT) $ list_comb (Const (cname, map fastype_of ts ---> T), ts)) in Logic.list_implies (prems, concl) @@ -710,7 +710,7 @@ (**** constructors ****) - fun mk_abs_fun (x, t) = + fun mk_abs_fun x t = let val T = fastype_of x; val U = fastype_of t @@ -776,7 +776,7 @@ fun make_constr_def tname T T' (((cname_rep, _), (cname, cargs)), (cname', mx)) (thy, defs, eqns) = let - fun constr_arg ((dts, dt), (j, l_args, r_args)) = + fun constr_arg (dts, dt) (j, l_args, r_args) = let val xs = map (fn (dt, i) => mk_Free "x" (typ_of_dtyp descr'' sorts dt) i) (dts ~~ (j upto j + length dts - 1)) @@ -784,16 +784,16 @@ in (j + length dts + 1, xs @ x :: l_args, - List.foldr mk_abs_fun + fold_rev mk_abs_fun xs (case dt of DtRec k => if k < length new_type_names then Const (List.nth (rep_names, k), typ_of_dtyp descr'' sorts dt --> typ_of_dtyp descr sorts dt) $ x else error "nested recursion not (yet) supported" - | _ => x) xs :: r_args) + | _ => x) :: r_args) end - val (_, l_args, r_args) = List.foldr constr_arg (1, [], []) cargs; + val (_, l_args, r_args) = fold_rev constr_arg cargs (1, [], []); val abs_name = Sign.intern_const thy ("Abs_" ^ tname); val rep_name = Sign.intern_const thy ("Rep_" ^ tname); val constrT = map fastype_of l_args ---> T; @@ -902,7 +902,7 @@ let val T = fastype_of t in Const ("Nominal.perm", permT --> T --> T) $ pi $ t end; - fun constr_arg ((dts, dt), (j, l_args, r_args)) = + fun constr_arg (dts, dt) (j, l_args, r_args) = let val Ts = map (typ_of_dtyp descr'' sorts) dts; val xs = map (fn (T, i) => mk_Free "x" T i) @@ -914,7 +914,7 @@ map perm (xs @ [x]) @ r_args) end - val (_, l_args, r_args) = List.foldr constr_arg (1, [], []) dts; + val (_, l_args, r_args) = fold_rev constr_arg dts (1, [], []); val c = Const (cname, map fastype_of l_args ---> T) in Goal.prove_global thy8 [] [] @@ -952,7 +952,7 @@ val cname = Sign.intern_const thy8 (Long_Name.append tname (Long_Name.base_name cname)); - fun make_inj ((dts, dt), (j, args1, args2, eqs)) = + fun make_inj (dts, dt) (j, args1, args2, eqs) = let val Ts_idx = map (typ_of_dtyp descr'' sorts) dts ~~ (j upto j + length dts - 1); val xs = map (fn (T, i) => mk_Free "x" T i) Ts_idx; @@ -963,10 +963,10 @@ (j + length dts + 1, xs @ (x :: args1), ys @ (y :: args2), HOLogic.mk_eq - (List.foldr mk_abs_fun x xs, List.foldr mk_abs_fun y ys) :: eqs) + (fold_rev mk_abs_fun xs x, fold_rev mk_abs_fun ys y) :: eqs) end; - val (_, args1, args2, eqs) = List.foldr make_inj (1, [], [], []) dts; + val (_, args1, args2, eqs) = fold_rev make_inj dts (1, [], [], []); val Ts = map fastype_of args1; val c = Const (cname, Ts ---> T) in @@ -995,17 +995,17 @@ (Long_Name.append tname (Long_Name.base_name cname)); val atomT = Type (atom, []); - fun process_constr ((dts, dt), (j, args1, args2)) = + fun process_constr (dts, dt) (j, args1, args2) = let val Ts_idx = map (typ_of_dtyp descr'' sorts) dts ~~ (j upto j + length dts - 1); val xs = map (fn (T, i) => mk_Free "x" T i) Ts_idx; val x = mk_Free "x" (typ_of_dtyp descr'' sorts dt) (j + length dts) in (j + length dts + 1, - xs @ (x :: args1), List.foldr mk_abs_fun x xs :: args2) + xs @ (x :: args1), fold_rev mk_abs_fun xs x :: args2) end; - val (_, args1, args2) = List.foldr process_constr (1, [], []) dts; + val (_, args1, args2) = fold_rev process_constr dts (1, [], []); val Ts = map fastype_of args1; val c = list_comb (Const (cname, Ts ---> T), args1); fun supp t = @@ -1166,11 +1166,11 @@ fun make_ind_prem fsT f k T ((cname, cargs), idxs) = let - val recs = List.filter is_rec_type cargs; + val recs = filter is_rec_type cargs; val Ts = map (typ_of_dtyp descr'' sorts) cargs; val recTs' = map (typ_of_dtyp descr'' sorts) recs; val tnames = Name.variant_list pnames (DatatypeProp.make_tnames Ts); - val rec_tnames = map fst (List.filter (is_rec_type o snd) (tnames ~~ cargs)); + val rec_tnames = map fst (filter (is_rec_type o snd) (tnames ~~ cargs)); val frees = tnames ~~ Ts; val frees' = partition_cargs idxs frees; val z = (Name.variant tnames "z", fsT); diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/Nominal/nominal_induct.ML --- a/src/HOL/Nominal/nominal_induct.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/HOL/Nominal/nominal_induct.ML Wed Nov 04 11:40:59 2009 +0100 @@ -7,7 +7,7 @@ sig val nominal_induct_tac: Proof.context -> (binding option * term) option list list -> (string * typ) list -> (string * typ) list list -> thm list -> - thm list -> int -> RuleCases.cases_tactic + thm list -> int -> Rule_Cases.cases_tactic val nominal_induct_method: (Proof.context -> Proof.method) context_parser end = struct @@ -31,9 +31,9 @@ fun inst_mutual_rule ctxt insts avoiding rules = let - val (nconcls, joined_rule) = RuleCases.strict_mutual_rule ctxt rules; + val (nconcls, joined_rule) = Rule_Cases.strict_mutual_rule ctxt rules; val concls = Logic.dest_conjunctions (Thm.concl_of joined_rule); - val (cases, consumes) = RuleCases.get joined_rule; + val (cases, consumes) = Rule_Cases.get joined_rule; val l = length rules; val _ = @@ -93,12 +93,12 @@ split_all_tuples #> rename_params_rule true (map (Name.clean o ProofContext.revert_skolem defs_ctxt o fst) avoiding); - fun rule_cases r = RuleCases.make_nested true (Thm.prop_of r) (Induct.rulified_term r); + fun rule_cases r = Rule_Cases.make_nested true (Thm.prop_of r) (Induct.rulified_term r); in (fn i => fn st => rules |> inst_mutual_rule ctxt insts avoiding - |> RuleCases.consume (flat defs) facts + |> Rule_Cases.consume (flat defs) facts |> Seq.maps (fn (((cases, concls), (more_consumes, more_facts)), rule) => (PRECISE_CONJUNCTS (length concls) (ALLGOALS (fn j => (CONJUNCTS (ALLGOALS diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/Nominal/nominal_inductive.ML --- a/src/HOL/Nominal/nominal_inductive.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/HOL/Nominal/nominal_inductive.ML Wed Nov 04 11:40:59 2009 +0100 @@ -158,7 +158,7 @@ [] => () | xs => error ("Missing equivariance theorem for predicate(s): " ^ commas_quote xs)); - val induct_cases = map fst (fst (RuleCases.get (the + val induct_cases = map fst (fst (Rule_Cases.get (the (Induct.lookup_inductP ctxt (hd names))))); val ([raw_induct'], ctxt') = Variable.import_terms false [prop_of raw_induct] ctxt; val concls = raw_induct' |> Logic.strip_imp_concl |> HOLogic.dest_Trueprop |> @@ -547,7 +547,7 @@ let val rec_name = space_implode "_" (map Long_Name.base_name names); val rec_qualified = Binding.qualify false rec_name; - val ind_case_names = RuleCases.case_names induct_cases; + val ind_case_names = Rule_Cases.case_names induct_cases; val induct_cases' = Inductive.partition_rules' raw_induct (intrs ~~ induct_cases); val thss' = map (map atomize_intr) thss; @@ -558,9 +558,9 @@ (thsss ~~ elims ~~ cases_prems ~~ cases_prems'); val strong_induct = if length names > 1 then - (strong_raw_induct, [ind_case_names, RuleCases.consumes 0]) + (strong_raw_induct, [ind_case_names, Rule_Cases.consumes 0]) else (strong_raw_induct RSN (2, rev_mp), - [ind_case_names, RuleCases.consumes 1]); + [ind_case_names, Rule_Cases.consumes 1]); val ((_, [strong_induct']), ctxt') = LocalTheory.note Thm.generatedK ((rec_qualified (Binding.name "strong_induct"), map (Attrib.internal o K) (#2 strong_induct)), [#1 strong_induct]) @@ -572,12 +572,12 @@ LocalTheory.note Thm.generatedK ((rec_qualified (Binding.name "strong_inducts"), [Attrib.internal (K ind_case_names), - Attrib.internal (K (RuleCases.consumes 1))]), + Attrib.internal (K (Rule_Cases.consumes 1))]), strong_inducts) |> snd |> LocalTheory.notes Thm.generatedK (map (fn ((name, elim), (_, cases)) => ((Binding.qualified_name (Long_Name.qualify (Long_Name.base_name name) "strong_cases"), - [Attrib.internal (K (RuleCases.case_names (map snd cases))), - Attrib.internal (K (RuleCases.consumes 1))]), [([elim], [])])) + [Attrib.internal (K (Rule_Cases.case_names (map snd cases))), + Attrib.internal (K (Rule_Cases.consumes 1))]), [([elim], [])])) (strong_cases ~~ induct_cases')) |> snd end) (map (map (rulify_term thy #> rpair [])) vc_compat) diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/Nominal/nominal_inductive2.ML --- a/src/HOL/Nominal/nominal_inductive2.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/HOL/Nominal/nominal_inductive2.ML Wed Nov 04 11:40:59 2009 +0100 @@ -164,7 +164,7 @@ [] => () | xs => error ("Missing equivariance theorem for predicate(s): " ^ commas_quote xs)); - val induct_cases = map fst (fst (RuleCases.get (the + val induct_cases = map fst (fst (Rule_Cases.get (the (Induct.lookup_inductP ctxt (hd names))))); val induct_cases' = if null induct_cases then replicate (length intrs) "" else induct_cases; @@ -449,7 +449,7 @@ let val rec_name = space_implode "_" (map Long_Name.base_name names); val rec_qualified = Binding.qualify false rec_name; - val ind_case_names = RuleCases.case_names induct_cases; + val ind_case_names = Rule_Cases.case_names induct_cases; val induct_cases' = Inductive.partition_rules' raw_induct (intrs ~~ induct_cases); val thss' = map (map atomize_intr) thss; @@ -458,9 +458,9 @@ mk_ind_proof ctxt thss' |> Inductive.rulify; val strong_induct = if length names > 1 then - (strong_raw_induct, [ind_case_names, RuleCases.consumes 0]) + (strong_raw_induct, [ind_case_names, Rule_Cases.consumes 0]) else (strong_raw_induct RSN (2, rev_mp), - [ind_case_names, RuleCases.consumes 1]); + [ind_case_names, Rule_Cases.consumes 1]); val (induct_name, inducts_name) = case alt_name of NONE => (rec_qualified (Binding.name "strong_induct"), @@ -477,7 +477,7 @@ LocalTheory.note Thm.generatedK ((inducts_name, [Attrib.internal (K ind_case_names), - Attrib.internal (K (RuleCases.consumes 1))]), + Attrib.internal (K (Rule_Cases.consumes 1))]), strong_inducts) |> snd end) (map (map (rulify_term thy #> rpair [])) vc_compat) diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/Nominal/nominal_primrec.ML --- a/src/HOL/Nominal/nominal_primrec.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/HOL/Nominal/nominal_primrec.ML Wed Nov 04 11:40:59 2009 +0100 @@ -193,7 +193,7 @@ NONE => let val dummy_fns = map (fn (_, cargs) => Const (@{const_name undefined}, - replicate ((length cargs) + (length (List.filter is_rec_type cargs))) + replicate (length cargs + length (filter is_rec_type cargs)) dummyT ---> HOLogic.unitT)) constrs; val _ = warning ("No function definition for datatype " ^ quote tname) in diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/Numeral_Simprocs.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Numeral_Simprocs.thy Wed Nov 04 11:40:59 2009 +0100 @@ -0,0 +1,120 @@ +(* Author: Various *) + +header {* Combination and Cancellation Simprocs for Numeral Expressions *} + +theory Numeral_Simprocs +imports Divides +uses + "~~/src/Provers/Arith/assoc_fold.ML" + "~~/src/Provers/Arith/cancel_numerals.ML" + "~~/src/Provers/Arith/combine_numerals.ML" + "~~/src/Provers/Arith/cancel_numeral_factor.ML" + "~~/src/Provers/Arith/extract_common_term.ML" + ("Tools/numeral_simprocs.ML") + ("Tools/nat_numeral_simprocs.ML") +begin + +declare split_div [of _ _ "number_of k", standard, arith_split] +declare split_mod [of _ _ "number_of k", standard, arith_split] + +text {* For @{text combine_numerals} *} + +lemma left_add_mult_distrib: "i*u + (j*u + k) = (i+j)*u + (k::nat)" +by (simp add: add_mult_distrib) + +text {* For @{text cancel_numerals} *} + +lemma nat_diff_add_eq1: + "j <= (i::nat) ==> ((i*u + m) - (j*u + n)) = (((i-j)*u + m) - n)" +by (simp split add: nat_diff_split add: add_mult_distrib) + +lemma nat_diff_add_eq2: + "i <= (j::nat) ==> ((i*u + m) - (j*u + n)) = (m - ((j-i)*u + n))" +by (simp split add: nat_diff_split add: add_mult_distrib) + +lemma nat_eq_add_iff1: + "j <= (i::nat) ==> (i*u + m = j*u + n) = ((i-j)*u + m = n)" +by (auto split add: nat_diff_split simp add: add_mult_distrib) + +lemma nat_eq_add_iff2: + "i <= (j::nat) ==> (i*u + m = j*u + n) = (m = (j-i)*u + n)" +by (auto split add: nat_diff_split simp add: add_mult_distrib) + +lemma nat_less_add_iff1: + "j <= (i::nat) ==> (i*u + m < j*u + n) = ((i-j)*u + m < n)" +by (auto split add: nat_diff_split simp add: add_mult_distrib) + +lemma nat_less_add_iff2: + "i <= (j::nat) ==> (i*u + m < j*u + n) = (m < (j-i)*u + n)" +by (auto split add: nat_diff_split simp add: add_mult_distrib) + +lemma nat_le_add_iff1: + "j <= (i::nat) ==> (i*u + m <= j*u + n) = ((i-j)*u + m <= n)" +by (auto split add: nat_diff_split simp add: add_mult_distrib) + +lemma nat_le_add_iff2: + "i <= (j::nat) ==> (i*u + m <= j*u + n) = (m <= (j-i)*u + n)" +by (auto split add: nat_diff_split simp add: add_mult_distrib) + +text {* For @{text cancel_numeral_factors} *} + +lemma nat_mult_le_cancel1: "(0::nat) < k ==> (k*m <= k*n) = (m<=n)" +by auto + +lemma nat_mult_less_cancel1: "(0::nat) < k ==> (k*m < k*n) = (m (k*m = k*n) = (m=n)" +by auto + +lemma nat_mult_div_cancel1: "(0::nat) < k ==> (k*m) div (k*n) = (m div n)" +by auto + +lemma nat_mult_dvd_cancel_disj[simp]: + "(k*m) dvd (k*n) = (k=0 | m dvd (n::nat))" +by(auto simp: dvd_eq_mod_eq_0 mod_mult_distrib2[symmetric]) + +lemma nat_mult_dvd_cancel1: "0 < k \ (k*m) dvd (k*n::nat) = (m dvd n)" +by(auto) + +text {* For @{text cancel_factor} *} + +lemma nat_mult_le_cancel_disj: "(k*m <= k*n) = ((0::nat) < k --> m<=n)" +by auto + +lemma nat_mult_less_cancel_disj: "(k*m < k*n) = ((0::nat) < k & m Lin_Arith.add_simps (@{thms ring_distribs} @ [@{thm Let_number_of}, @{thm Let_0}, @{thm Let_1}, + @{thm nat_0}, @{thm nat_1}, + @{thm add_nat_number_of}, @{thm diff_nat_number_of}, @{thm mult_nat_number_of}, + @{thm eq_nat_number_of}, @{thm less_nat_number_of}, @{thm le_number_of_eq_not_less}, + @{thm le_Suc_number_of}, @{thm le_number_of_Suc}, + @{thm less_Suc_number_of}, @{thm less_number_of_Suc}, + @{thm Suc_eq_number_of}, @{thm eq_number_of_Suc}, + @{thm mult_Suc}, @{thm mult_Suc_right}, + @{thm add_Suc}, @{thm add_Suc_right}, + @{thm eq_number_of_0}, @{thm eq_0_number_of}, @{thm less_0_number_of}, + @{thm of_int_number_of_eq}, @{thm of_nat_number_of_eq}, @{thm nat_number_of}, + @{thm if_True}, @{thm if_False}]) + #> Lin_Arith.add_simprocs (Numeral_Simprocs.assoc_fold_simproc + :: Numeral_Simprocs.combine_numerals + :: Numeral_Simprocs.cancel_numerals) + #> Lin_Arith.add_simprocs (Nat_Numeral_Simprocs.combine_numerals :: Nat_Numeral_Simprocs.cancel_numerals)) +*} + +end \ No newline at end of file diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/OrderedGroup.thy --- a/src/HOL/OrderedGroup.thy Thu Oct 29 16:23:57 2009 +0100 +++ b/src/HOL/OrderedGroup.thy Wed Nov 04 11:40:59 2009 +0100 @@ -1409,4 +1409,13 @@ Addsimprocs [ab_group_add_cancel.sum_conv, ab_group_add_cancel.rel_conv]; *} +code_modulename SML + OrderedGroup Arith + +code_modulename OCaml + OrderedGroup Arith + +code_modulename Haskell + OrderedGroup Arith + end diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/Parity.thy --- a/src/HOL/Parity.thy Thu Oct 29 16:23:57 2009 +0100 +++ b/src/HOL/Parity.thy Wed Nov 04 11:40:59 2009 +0100 @@ -315,42 +315,6 @@ qed qed -subsection {* General Lemmas About Division *} - -(*FIXME move to Divides.thy*) - -lemma Suc_times_mod_eq: "1 Suc (k * m) mod k = 1" -apply (induct "m") -apply (simp_all add: mod_Suc) -done - -declare Suc_times_mod_eq [of "number_of w", standard, simp] - -lemma [simp]: "n div k \ (Suc n) div k" -by (simp add: div_le_mono) - -lemma Suc_n_div_2_gt_zero [simp]: "(0::nat) < n ==> 0 < (n + 1) div 2" -by arith - -lemma div_2_gt_zero [simp]: "(1::nat) < n ==> 0 < n div 2" -by arith - - (* Potential use of algebra : Equality modulo n*) -lemma mod_mult_self3 [simp]: "(k*n + m) mod n = m mod (n::nat)" -by (simp add: mult_ac add_ac) - -lemma mod_mult_self4 [simp]: "Suc (k*n + m) mod n = Suc m mod n" -proof - - have "Suc (k * n + m) mod n = (k * n + Suc m) mod n" by simp - also have "... = Suc m mod n" by (rule mod_mult_self3) - finally show ?thesis . -qed - -lemma mod_Suc_eq_Suc_mod: "Suc m mod n = Suc (m mod n) mod n" -apply (subst mod_Suc [of m]) -apply (subst mod_Suc [of "m mod n"], simp) -done - subsection {* More Even/Odd Results *} diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/Power.thy --- a/src/HOL/Power.thy Thu Oct 29 16:23:57 2009 +0100 +++ b/src/HOL/Power.thy Wed Nov 04 11:40:59 2009 +0100 @@ -467,4 +467,13 @@ declare power.power.simps [code] +code_modulename SML + Power Arith + +code_modulename OCaml + Power Arith + +code_modulename Haskell + Power Arith + end diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/Ring_and_Field.thy --- a/src/HOL/Ring_and_Field.thy Thu Oct 29 16:23:57 2009 +0100 +++ b/src/HOL/Ring_and_Field.thy Wed Nov 04 11:40:59 2009 +0100 @@ -2374,4 +2374,14 @@ then show ?thesis by simp qed + +code_modulename SML + Ring_and_Field Arith + +code_modulename OCaml + Ring_and_Field Arith + +code_modulename Haskell + Ring_and_Field Arith + end diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/SMT/Examples/cert/z3_arith_quant_18 --- a/src/HOL/SMT/Examples/cert/z3_arith_quant_18 Thu Oct 29 16:23:57 2009 +0100 +++ b/src/HOL/SMT/Examples/cert/z3_arith_quant_18 Wed Nov 04 11:40:59 2009 +0100 @@ -2,6 +2,6 @@ :extrafuns ( (uf_1 Int) ) -:assumption (not (forall (?x1 Int) (implies (< ?x1 uf_1) (< (* 2 ?x1) (* 2 uf_1))) :pat{ ?x1 })) +:assumption (not (forall (?x1 Int) (implies (< ?x1 uf_1) (< (* 2 ?x1) (* 2 uf_1))) :pat { ?x1 })) :formula true ) diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/SMT/Tools/cvc3_solver.ML --- a/src/HOL/SMT/Tools/cvc3_solver.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/HOL/SMT/Tools/cvc3_solver.ML Wed Nov 04 11:40:59 2009 +0100 @@ -15,21 +15,13 @@ val solver_name = "cvc3" val env_var = "CVC3_SOLVER" -val options = - ["+counterexample", "-lang", "smtlib", "-output-lang", "presentation"] +val options = ["-lang", "smtlib", "-output-lang", "presentation"] val is_sat = String.isPrefix "Satisfiable." val is_unsat = String.isPrefix "Unsatisfiable." val is_unknown = String.isPrefix "Unknown." -fun cex_kind true = "Counterexample" - | cex_kind false = "Possible counterexample" - -fun raise_cex real ctxt recon ls = - let - val ls' = filter_out (String.isPrefix "%") ls - val p = Pretty.big_list (cex_kind real ^ " found:") (map Pretty.str ls') - in error (Pretty.string_of p) end +fun raise_cex real = raise SMT_Solver.SMT_COUNTEREXAMPLE (real, []) fun core_oracle ({context, output, recon, ...} : SMT_Solver.proof_data) = let @@ -38,9 +30,9 @@ val (l, ls) = split_first (dropwhile empty_line output) in if is_unsat l then @{cprop False} - else if is_sat l then raise_cex true context recon ls - else if is_unknown l then raise_cex false context recon ls - else error (solver_name ^ " failed") + else if is_sat l then raise_cex true + else if is_unknown l then raise_cex false + else raise SMT_Solver.SMT (solver_name ^ " failed") end fun smtlib_solver oracle _ = { diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/SMT/Tools/smt_normalize.ML --- a/src/HOL/SMT/Tools/smt_normalize.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/HOL/SMT/Tools/smt_normalize.ML Wed Nov 04 11:40:59 2009 +0100 @@ -318,7 +318,7 @@ fun replace ctxt cvs ct (cx as (nctxt, defs)) = let val cvs' = used_vars cvs ct - val ct' = fold Thm.cabs cvs' ct + val ct' = fold_rev Thm.cabs cvs' ct val mk_repl = fold (fn ct => fn cu => Thm.capply cu ct) cvs' in (case Termtab.lookup defs (Thm.term_of ct') of diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/SMT/Tools/smt_solver.ML --- a/src/HOL/SMT/Tools/smt_solver.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/HOL/SMT/Tools/smt_solver.ML Wed Nov 04 11:40:59 2009 +0100 @@ -6,6 +6,7 @@ signature SMT_SOLVER = sig + exception SMT of string exception SMT_COUNTEREXAMPLE of bool * term list type interface = { @@ -53,6 +54,7 @@ structure SMT_Solver: SMT_SOLVER = struct +exception SMT of string exception SMT_COUNTEREXAMPLE of bool * term list @@ -79,7 +81,7 @@ fun with_timeout ctxt f x = TimeLimit.timeLimit (Time.fromSeconds (Config.get ctxt timeout)) f x - handle TimeLimit.TimeOut => error ("SMT: timeout") + handle TimeLimit.TimeOut => raise SMT "timeout" val (trace, setup_trace) = Attrib.config_bool "smt_trace" false @@ -112,13 +114,17 @@ fun run in_path out_path (ctxt, cmd, output) = let + fun pretty tag ls = Pretty.string_of (Pretty.big_list tag + (map Pretty.str ls)) + val x = File.open_output output in_path - val _ = trace_msg ctxt File.read in_path + val _ = trace_msg ctxt (pretty "SMT problem:" o split_lines o File.read) + in_path val _ = with_timeout ctxt system_out (cmd in_path out_path) fun lines_of path = the_default [] (try (File.fold_lines cons out_path) []) val ls = rev (dropwhile (equal "") (lines_of out_path)) - val _ = trace_msg ctxt cat_lines ls + val _ = trace_msg ctxt (pretty "SMT result:") ls in (x, ls) end in @@ -206,28 +212,34 @@ (* SMT tactic *) -fun smt_unsat_tac solver ctxt rules = +local + fun pretty_cex ctxt (real, ex) = + let + val msg = if real then "SMT: counterexample found" + else "SMT: potential counterexample found" + in + if null ex then msg ^ "." + else Pretty.string_of (Pretty.big_list (msg ^ ":") + (map (Syntax.pretty_term ctxt) ex)) + end + + fun fail_tac ctxt msg st = (trace_msg ctxt I msg; Tactical.no_tac st) + + fun SAFE pass_exns tac ctxt i st = + if pass_exns then tac ctxt i st + else (tac ctxt i st + handle SMT msg => fail_tac ctxt ("SMT: " ^ msg) st + | SMT_COUNTEREXAMPLE cex => fail_tac ctxt (pretty_cex ctxt cex) st) + + fun smt_solver rules ctxt = solver_of (Context.Proof ctxt) ctxt rules +in +fun smt_tac' pass_exns ctxt rules = Tactic.rtac @{thm ccontr} THEN' SUBPROOF (fn {context, prems, ...} => - Tactic.rtac (solver context (rules @ prems)) 1) ctxt - -fun pretty_counterex ctxt (real, ex) = - let - val msg = if real then "SMT: counterexample found:" - else "SMT: potential counterexample found:" - val cex = if null ex then [Pretty.str "(no assignments)"] - else map (Syntax.pretty_term ctxt) ex - in Pretty.string_of (Pretty.big_list msg cex) end - -fun smt_tac' pass_smt_exns ctxt = - let - val solver = solver_of (Context.Proof ctxt) - fun safe_solver ctxt thms = solver ctxt thms - handle SMT_COUNTEREXAMPLE cex => error (pretty_counterex ctxt cex) - val solver' = if pass_smt_exns then solver else safe_solver - in smt_unsat_tac solver' ctxt end + SAFE pass_exns (Tactic.rtac o smt_solver (rules @ prems)) context 1) ctxt val smt_tac = smt_tac' false +end val smt_method = Scan.optional Attrib.thms [] >> diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/SMT/Tools/smtlib_interface.ML --- a/src/HOL/SMT/Tools/smtlib_interface.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/HOL/SMT/Tools/smtlib_interface.ML Wed Nov 04 11:40:59 2009 +0100 @@ -69,7 +69,6 @@ fun wr1 s = sep (wr s) fun wrn f n = (fn [] => wr1 n | ts => par (wr n #> fold f ts)) -fun ins s f = (fn [] => I | x::xs => f x #> fold (fn x => wr1 s #> f x) xs) val term_marker = "__term" val formula_marker = "__form" @@ -88,7 +87,7 @@ SOME (loc', t') => wr_expr loc' env t' | NONE => wrn (wr_expr loc env) n ts) | T.SLet ((v, _), t1, t2) => - if loc then raise TERM ("SMTLIB: let expression in term", []) + if loc then raise SMT_Solver.SMT "SMTLIB: let expression in term" else let val (loc', t1') = the (dest_marker t1) @@ -104,7 +103,7 @@ val wre = wr_expr loc (map (tvar o fst) (rev vs) @ env) - fun wrp s ts = wr1 (":" ^ s ^ "{") #> ins "," wre ts #> wr1 "}" + fun wrp s ts = wr1 (":" ^ s ^ " {") #> fold wre ts #> wr1 "}" fun wr_pat (T.SPat ts) = wrp "pat" ts | wr_pat (T.SNoPat ts) = wrp "nopat" ts in par (wr_quant q #> fold wr_var vs #> wre b #> fold wr_pat ps) end) diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/SMT/Tools/yices_solver.ML --- a/src/HOL/SMT/Tools/yices_solver.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/HOL/SMT/Tools/yices_solver.ML Wed Nov 04 11:40:59 2009 +0100 @@ -15,16 +15,9 @@ val solver_name = "yices" val env_var = "YICES_SOLVER" -val options = ["--evidence", "--smtlib"] - -fun cex_kind true = "Counterexample" - | cex_kind false = "Possible counterexample" +val options = ["--smtlib"] -fun raise_cex real ctxt rtab ls = - let val p = Pretty.big_list (cex_kind real ^ " found:") (map Pretty.str ls) - in error (Pretty.string_of p) end - -structure S = SMT_Solver +fun raise_cex real = raise SMT_Solver.SMT_COUNTEREXAMPLE (real, []) fun core_oracle ({context, output, recon, ...} : SMT_Solver.proof_data) = let @@ -33,9 +26,9 @@ val (l, ls) = split_first (dropwhile empty_line output) in if String.isPrefix "unsat" l then @{cprop False} - else if String.isPrefix "sat" l then raise_cex true context recon ls - else if String.isPrefix "unknown" l then raise_cex false context recon ls - else error (solver_name ^ " failed") + else if String.isPrefix "sat" l then raise_cex true + else if String.isPrefix "unknown" l then raise_cex false + else raise SMT_Solver.SMT (solver_name ^ " failed") end fun smtlib_solver oracle _ = { diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/SMT/Tools/z3_model.ML --- a/src/HOL/SMT/Tools/z3_model.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/HOL/SMT/Tools/z3_model.ML Wed Nov 04 11:40:59 2009 +0100 @@ -12,15 +12,18 @@ structure Z3_Model: Z3_MODEL = struct -(* parsing primitives *) +(* counterexample expressions *) -fun lift f (x, y) = apsnd (pair x) (f y) -fun lift' f v (x, y) = apsnd (rpair y) (f v x) +datatype expr = True | False | Number of int * int option | Value of int | + Array of array +and array = Fresh of expr | Store of (array * expr) * expr -fun $$ s = lift (Scan.$$ s) -fun this s = lift (Scan.this_string s) + +(* parsing *) -fun par scan = $$ "(" |-- scan --| $$ ")" +val space = Scan.many Symbol.is_ascii_blank +fun in_parens p = Scan.$$ "(" |-- p --| Scan.$$ ")" +fun in_braces p = (space -- Scan.$$ "{") |-- p --| (space -- Scan.$$ "}") val digit = (fn "0" => SOME 0 | "1" => SOME 1 | "2" => SOME 2 | "3" => SOME 3 | @@ -29,131 +32,113 @@ val nat_num = Scan.repeat1 (Scan.some digit) >> (fn ds => fold (fn d => fn i => i * 10 + d) ds 0) -val int_num = Scan.optional (Scan.$$ "-" >> K (fn i => ~i)) I :|-- +val int_num = Scan.optional ($$ "-" >> K (fn i => ~i)) I :|-- (fn sign => nat_num >> sign) val is_char = Symbol.is_ascii_letter orf Symbol.is_ascii_digit orf member (op =) (explode "_+*-/%~=<>$&|?!.@^#") val name = Scan.many1 is_char >> implode - -(* parsing counterexamples *) - -datatype context = Context of { - ttab: term Symtab.table, - nctxt: Name.context, - vtab: term Inttab.table } +fun array_expr st = st |> + in_parens (space |-- ( + Scan.this_string "const" |-- expr >> Fresh || + Scan.this_string "store" -- space |-- array_expr -- expr -- expr >> Store)) -fun make_context (ttab, nctxt, vtab) = - Context {ttab=ttab, nctxt=nctxt, vtab=vtab} +and expr st = st |> (space |-- ( + Scan.this_string "true" >> K True || + Scan.this_string "false" >> K False || + int_num -- Scan.option (Scan.$$ "/" |-- int_num) >> Number || + Scan.this_string "val!" |-- nat_num >> Value || + array_expr >> Array)) -fun empty_context ({terms, ...} : SMT_Translate.recon) = - let - val ns = Symtab.fold (Term.add_free_names o snd) terms [] - val nctxt = Name.make_context ns - in make_context (terms, nctxt, Inttab.empty) end - -fun map_context f (Context {ttab, nctxt, vtab}) = - make_context (f (ttab, nctxt, vtab)) +val mapping = space -- Scan.this_string "->" +val value = mapping |-- expr -fun fresh_name (cx as Context {nctxt, ...}) = - let val (n, nctxt') = yield_singleton Name.variants "" nctxt - in (n, map_context (fn (ttab, _, vtab) => (ttab, nctxt', vtab)) cx) end +val args_case = Scan.repeat expr -- value +val else_case = space -- Scan.this_string "else" |-- value >> pair [] + +val func = + let fun cases st = (else_case >> single || args_case ::: cases) st + in in_braces cases end -fun ident name (cx as Context {ttab, ...}) = - (case Symtab.lookup ttab name of - SOME t => (t, cx) - | NONE => - let val (n, cx') = fresh_name cx - in (Free (n, Term.dummyT), cx) end) +val cex = space |-- Scan.repeat (space |-- name --| mapping -- + (func || expr >> (single o pair []))) -fun set_value t i = map_context (fn (ttab, nctxt, vtab) => - (ttab, nctxt, Inttab.update (i, t) vtab)) - -fun get_value T i (cx as Context {vtab, ...}) = - (case Inttab.lookup vtab i of - SOME t => (t, cx) - | _ => cx |> fresh_name |-> (fn n => - let val t = Free (n, T) - in set_value t i #> pair t end)) +fun read_cex ls = + explode (cat_lines ls) + |> try (fst o Scan.finite Symbol.stopper cex) + |> the_default [] -fun space s = lift (Scan.many Symbol.is_ascii_blank) s -fun spaced p = p --| space +(* translation into terms *) -val key = spaced (lift name) #-> lift' ident -fun mapping st = spaced (this "->") st -fun in_braces p = spaced ($$ "{") |-- p --| spaced ($$ "}") +fun lookup_term tab (name, e) = Option.map (rpair e) (Symtab.lookup tab name) -fun bool_expr st = - (this "true" >> K @{term True} || - this "false" >> K @{term False}) st - -fun number_expr T = +fun with_name_context tab f xs = let - val num = lift int_num >> HOLogic.mk_number T - fun frac n d = Const (@{const_name divide}, T --> T --> T) $ n $ d - in num :|-- (fn n => Scan.optional ($$ "/" |-- num >> frac n) n) end + val ns = Symtab.fold (Term.add_free_names o snd) tab [] + val nctxt = Name.make_context ns + in fst (fold_map f xs (Inttab.empty, nctxt)) end -fun value st = (this "val!" |-- lift nat_num) st -fun value_expr T = value #-> lift' (get_value T) +fun fresh_term T (tab, nctxt) = + let val (n, nctxt') = yield_singleton Name.variants "" nctxt + in (Free (n, T), (tab, nctxt')) end -val domT = Term.domain_type -val ranT = Term.range_type -fun const_array T t = Abs ("x", T, t) -fun upd_array T ((a, t), u) = - Const (@{const_name fun_upd}, [T, domT T, ranT T] ---> T) $ a $ t $ u -fun array_expr T st = if not (can domT T) then Scan.fail st else st |> ( - par (spaced (this "const") |-- expr (ranT T)) >> const_array (domT T) || - par (spaced (this "store") |-- spaced (array_expr T) -- - expr (Term.domain_type T) -- expr (Term.range_type T)) >> upd_array T) +fun term_of_value T i (cx as (tab, _)) = + (case Inttab.lookup tab i of + SOME t => (t, cx) + | NONE => + let val (t, (tab', nctxt')) = fresh_term T cx + in (t, (Inttab.update (i, t) tab', nctxt')) end) -and expr T st = - spaced (bool_expr || number_expr T || value_expr T || array_expr T) st +fun trans_expr _ True = pair @{term True} + | trans_expr _ False = pair @{term False} + | trans_expr T (Number (i, NONE)) = pair (HOLogic.mk_number T i) + | trans_expr T (Number (i, SOME j)) = + pair (Const (@{const_name divide}, [T, T] ---> T) $ + HOLogic.mk_number T i $ HOLogic.mk_number T j) + | trans_expr T (Value i) = term_of_value T i + | trans_expr T (Array a) = trans_array T a -fun const_val t = - let fun rep u = spaced value #-> apfst o set_value u #> pair [] +and trans_array T a = + let val dT = Term.domain_type T and rT = Term.range_type T in - if can HOLogic.dest_number t then rep t - else if t = @{term TT} then rep @{term True} - else expr (Term.fastype_of t) >> (fn u => [HOLogic.mk_eq (t, u)]) + (case a of + Fresh e => trans_expr rT e #>> (fn t => Abs ("x", dT, t)) + | Store ((a', e1), e2) => + trans_array T a' ##>> trans_expr dT e1 ##>> trans_expr rT e2 #>> + (fn ((m, k), v) => + Const (@{const_name fun_upd}, [T, dT, rT] ---> T) $ m $ k $ v)) end -fun func_value T = mapping |-- expr T +fun trans_pat i T f x = + f (Term.domain_type T) ##>> trans (i-1) (Term.range_type T) x #>> + (fn (u, (us, t)) => (u :: us, t)) + +and trans i T ([], v) = + if i > 0 then trans_pat i T fresh_term ([], v) + else trans_expr T v #>> pair [] + | trans i T (p :: ps, v) = trans_pat i T (fn U => trans_expr U p) (ps, v) -fun first_pat T = - let - fun args T = if not (can domT T) then Scan.succeed [] else - expr (domT T) ::: args (ranT T) - fun value ts = func_value (snd (SMT_Translate.dest_funT (length ts) T)) - in args T :-- value end +fun mk_eq' t us u = HOLogic.mk_eq (Term.list_comb (t, us), u) +fun mk_eq (Const (@{const_name apply}, _)) (u' :: us', u) = mk_eq' u' us' u + | mk_eq t (us, u) = mk_eq' t us u -fun func_pat (Ts, T) = fold_map expr Ts -- func_value T -fun else_pat (Ts, T) = - let fun else_arg T cx = cx |> fresh_name |>> (fn n => Free (n, T)) +fun translate (t, cs) = + let val T = Term.fastype_of t in - fold_map (lift' else_arg) Ts ##>> - spaced (this "else") |-- func_value T + (case (can HOLogic.dest_number t, cs) of + (true, [c]) => trans 0 T c #>> (fn (_, u) => [mk_eq u ([], t)]) + | (_, (es, _) :: _) => fold_map (trans (length es) T) cs #>> map (mk_eq t)) end -fun next_pats T (fts as (ts, _)) = - let val Tps = SMT_Translate.dest_funT (length ts) T - in Scan.repeat (func_pat Tps) @@@ (else_pat Tps >> single) >> cons fts end + + +(* overall procedure *) -fun mk_def' f (ts, t) = HOLogic.mk_eq (Term.list_comb (f, ts), t) -fun mk_def (Const (@{const_name apply}, _)) (u :: us, t) = mk_def' u (us, t) - | mk_def f (ts, t) = mk_def' f (ts, t) -fun func_pats t = - let val T = Term.fastype_of t - in first_pat T :|-- next_pats T >> map (mk_def t) end - -val assign = - key --| mapping :|-- (fn t => in_braces (func_pats t) || const_val t) - -val cex = space |-- Scan.repeat assign - -fun parse_counterex recon ls = - (empty_context recon, explode (cat_lines ls)) - |> Scan.catch (Scan.finite' Symbol.stopper (Scan.error cex)) - |> flat o fst +fun parse_counterex ({terms, ...} : SMT_Translate.recon) ls = + read_cex ls + |> map_filter (lookup_term terms) + |> with_name_context terms translate + |> flat end diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/SMT/Tools/z3_proof.ML --- a/src/HOL/SMT/Tools/z3_proof.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/HOL/SMT/Tools/z3_proof.ML Wed Nov 04 11:40:59 2009 +0100 @@ -16,7 +16,7 @@ structure T = Z3_Proof_Terms structure R = Z3_Proof_Rules -fun z3_exn msg = error ("Z3 proof reconstruction: " ^ msg) +fun z3_exn msg = raise SMT_Solver.SMT ("Z3 proof reconstruction: " ^ msg) fun lift f (x, y) = apsnd (pair x) (f y) @@ -223,10 +223,11 @@ def :|-- (fn k => expr thy k #> pair NONE || rule thy k #>> K NONE) || rule thy ~1 #>> SOME -fun parse_error line_no ((_, xs), _) = - "parse error at line " ^ string_of_int line_no ^ ": " ^ quote (implode xs) - -fun handle_errors ln scan = Scan.error (Scan.!! (parse_error ln) scan) +fun handle_errors line_no scan (st as (_, xs)) = + (case try scan st of + SOME y => y + | NONE => z3_exn ("parse error at line " ^ string_of_int line_no ^ ": " ^ + quote (implode xs))) fun parse_line thy l (st as (stop, line_no, cx)) = if is_some stop then st diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/SMT/Tools/z3_proof_rules.ML --- a/src/HOL/SMT/Tools/z3_proof_rules.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/HOL/SMT/Tools/z3_proof_rules.ML Wed Nov 04 11:40:59 2009 +0100 @@ -26,7 +26,7 @@ structure T = Z3_Proof_Terms -fun z3_exn msg = error ("Z3 proof reconstruction: " ^ msg) +fun z3_exn msg = raise SMT_Solver.SMT ("Z3 proof reconstruction: " ^ msg) (* proof rule names *) diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/SMT/Tools/z3_solver.ML --- a/src/HOL/SMT/Tools/z3_solver.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/HOL/SMT/Tools/z3_solver.ML Wed Nov 04 11:40:59 2009 +0100 @@ -52,7 +52,7 @@ if String.isPrefix "unsat" l then ls else if String.isPrefix "sat" l then raise_cex true recon ls else if String.isPrefix "unknown" l then raise_cex false recon ls - else error (solver_name ^ " failed") + else raise SMT_Solver.SMT (solver_name ^ " failed") end fun core_oracle ({output, recon, ...} : SMT_Solver.proof_data) = diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/SetInterval.thy --- a/src/HOL/SetInterval.thy Thu Oct 29 16:23:57 2009 +0100 +++ b/src/HOL/SetInterval.thy Wed Nov 04 11:40:59 2009 +0100 @@ -508,7 +508,7 @@ proof(induct A rule:finite_linorder_max_induct) case empty thus ?case by auto next - case (insert A b) + case (insert b A) moreover hence "b ~: A" by auto moreover have "A <= {k.. (case lookup_thread active thread of SOME (birth_time, _, description) => @@ -149,7 +149,7 @@ do (Synchronized.timed_access global_state (SOME o time_limit o #timeout_heap) action |> these - |> List.app (unregister (false, "Interrupted (reached timeout)")); + |> List.app (unregister "Interrupted (reached timeout)"); print_new_messages (); (*give threads some time to respond to interrupt*) OS.Process.sleep min_wait_time) @@ -263,14 +263,11 @@ let val _ = register birth_time death_time (Thread.self (), desc); val problem = ATP_Wrapper.problem_of_goal (! full_types) i (ctxt, (facts, goal)); - val result = - let val {success, message, ...} = prover (! timeout) problem; - in (success, message) end - handle ResHolClause.TOO_TRIVIAL => (* FIXME !? *) - (true, "Empty clause: Try this command: " ^ - Markup.markup Markup.sendback "apply metis") - | ERROR msg => (false, "Error: " ^ msg); - val _ = unregister result (Thread.self ()); + val message = #message (prover (! timeout) problem) + handle Res_HOL_Clause.TOO_TRIVIAL => (* FIXME !? *) + "Empty clause: Try this command: " ^ Markup.markup Markup.sendback "apply metis" + | ERROR msg => ("Error: " ^ msg); + val _ = unregister message (Thread.self ()); in () end) in () end); diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/Tools/ATP_Manager/atp_minimal.ML --- a/src/HOL/Tools/ATP_Manager/atp_minimal.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/HOL/Tools/ATP_Manager/atp_minimal.ML Wed Nov 04 11:40:59 2009 +0100 @@ -103,7 +103,7 @@ let val _ = priority ("Testing " ^ string_of_int (length name_thms_pairs) ^ " theorems... ") val name_thm_pairs = maps (fn (n, ths) => map (pair n) ths) name_thms_pairs - val axclauses = ResAxioms.cnf_rules_pairs (Proof.theory_of state) name_thm_pairs + val axclauses = Res_Axioms.cnf_rules_pairs (Proof.theory_of state) name_thm_pairs val {context = ctxt, facts, goal} = Proof.raw_goal state (* FIXME ?? *) val problem = {with_full_types = ! ATP_Manager.full_types, @@ -138,9 +138,9 @@ val to_use = if length ordered_used < length name_thms_pairs then filter (fn (name1, _) => List.exists (equal name1) ordered_used) name_thms_pairs - else - name_thms_pairs - val (min_thms, n) = if null to_use then ([], 0) + else name_thms_pairs + val (min_thms, n) = + if null to_use then ([], 0) else minimal (test_thms (SOME filtered)) to_use val min_names = sort_distinct string_ord (map fst min_thms) val _ = priority (cat_lines @@ -157,7 +157,7 @@ (NONE, "Error in prover: " ^ msg) | (Failure, _) => (NONE, "Failure: No proof with the theorems supplied")) - handle ResHolClause.TOO_TRIVIAL => + handle Res_HOL_Clause.TOO_TRIVIAL => (SOME ([], 0), "Trivial: Try this command: " ^ Markup.markup Markup.sendback "apply metis") | ERROR msg => (NONE, "Error: " ^ msg) end diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/Tools/ATP_Manager/atp_wrapper.ML --- a/src/HOL/Tools/ATP_Manager/atp_wrapper.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/HOL/Tools/ATP_Manager/atp_wrapper.ML Wed Nov 04 11:40:59 2009 +0100 @@ -117,8 +117,8 @@ (* get clauses and prepare them for writing *) val (ctxt, (chain_ths, th)) = goal; val thy = ProofContext.theory_of ctxt; - val chain_ths = map (Thm.put_name_hint ResReconstruct.chained_hint) chain_ths; - val goal_cls = #1 (ResAxioms.neg_conjecture_clauses ctxt th subgoalno); + val chain_ths = map (Thm.put_name_hint Res_Reconstruct.chained_hint) chain_ths; + val goal_cls = #1 (Res_Axioms.neg_conjecture_clauses ctxt th subgoalno); val the_filtered_clauses = (case filtered_clauses of NONE => relevance_filter goal goal_cls @@ -204,14 +204,14 @@ val {with_full_types, subgoal, goal, axiom_clauses, filtered_clauses} = problem; in external_prover - (ResAtp.get_relevant max_new_clauses insert_theory_const) - (ResAtp.prepare_clauses false) - (ResHolClause.tptp_write_file with_full_types) + (Res_ATP.get_relevant max_new_clauses insert_theory_const) + (Res_ATP.prepare_clauses false) + (Res_HOL_Clause.tptp_write_file with_full_types) command (arguments timeout) - ResReconstruct.find_failure - (if emit_structured_proof then ResReconstruct.structured_proof - else ResReconstruct.lemma_list false) + Res_Reconstruct.find_failure + (if emit_structured_proof then Res_Reconstruct.structured_proof + else Res_Reconstruct.lemma_list false) axiom_clauses filtered_clauses name @@ -280,13 +280,13 @@ val {with_full_types, subgoal, goal, axiom_clauses, filtered_clauses} = problem in external_prover - (ResAtp.get_relevant max_new_clauses insert_theory_const) - (ResAtp.prepare_clauses true) - (ResHolClause.dfg_write_file with_full_types) + (Res_ATP.get_relevant max_new_clauses insert_theory_const) + (Res_ATP.prepare_clauses true) + (Res_HOL_Clause.dfg_write_file with_full_types) command (arguments timeout) - ResReconstruct.find_failure - (ResReconstruct.lemma_list true) + Res_Reconstruct.find_failure + (Res_Reconstruct.lemma_list true) axiom_clauses filtered_clauses name diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/Tools/Datatype/datatype.ML --- a/src/HOL/Tools/Datatype/datatype.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/HOL/Tools/Datatype/datatype.ML Wed Nov 04 11:40:59 2009 +0100 @@ -63,9 +63,11 @@ val get_all = #types o DatatypesData.get; val get_info = Symtab.lookup o get_all; -fun the_info thy name = (case get_info thy name of - SOME info => info - | NONE => error ("Unknown datatype " ^ quote name)); + +fun the_info thy name = + (case get_info thy name of + SOME info => info + | NONE => error ("Unknown datatype " ^ quote name)); fun info_of_constr thy (c, T) = let @@ -94,6 +96,7 @@ cases = cases |> fold Symtab.update (map (fn (_, info as {case_name, ...}) => (case_name, info)) dt_infos)}); + (* complex queries *) fun the_spec thy dtco = @@ -155,6 +158,7 @@ | NONE => NONE; + (** various auxiliary **) (* prepare datatype specifications *) @@ -199,14 +203,15 @@ in -fun mk_case_names_induct descr = RuleCases.case_names (induct_cases descr); +fun mk_case_names_induct descr = Rule_Cases.case_names (induct_cases descr); fun mk_case_names_exhausts descr new = - map (RuleCases.case_names o exhaust_cases descr o #1) + map (Rule_Cases.case_names o exhaust_cases descr o #1) (filter (fn ((_, (name, _, _))) => member (op =) new name) descr); end; + (* translation rules for case *) fun make_case ctxt = DatatypeCase.make_case @@ -228,6 +233,7 @@ [], []); + (** document antiquotation **) val _ = ThyOutput.antiquotation "datatype" Args.tyname @@ -254,15 +260,17 @@ in ThyOutput.output (ThyOutput.maybe_pretty_source (K pretty_datatype) src [()]) end); + (** abstract theory extensions relative to a datatype characterisation **) -structure DatatypeInterpretation = InterpretationFun +structure Datatype_Interpretation = Interpretation (type T = config * string list val eq: T * T -> bool = eq_snd op =); -fun interpretation f = DatatypeInterpretation.interpretation (uncurry f); +fun interpretation f = Datatype_Interpretation.interpretation (uncurry f); fun make_dt_info alt_names descr sorts induct inducts rec_names rec_rewrites (index, (((((((((((_, (tname, _, _))), inject), distinct), - exhaust), nchotomy), case_name), case_rewrites), case_cong), weak_case_cong), (split, split_asm))) = + exhaust), nchotomy), case_name), case_rewrites), case_cong), weak_case_cong), + (split, split_asm))) = (tname, {index = index, alt_names = alt_names, @@ -309,7 +317,8 @@ config new_type_names descr sorts inject distinct exhaust case_rewrites thy8; val inducts = Project_Rule.projections (ProofContext.init thy2) induct; - val dt_infos = map_index (make_dt_info alt_names flat_descr sorts induct inducts rec_names rec_rewrites) + val dt_infos = map_index + (make_dt_info alt_names flat_descr sorts induct inducts rec_names rec_rewrites) (hd descr ~~ inject ~~ distinct ~~ exhaust ~~ nchotomys ~~ case_names ~~ case_rewrites ~~ case_congs ~~ weak_case_congs ~~ splits); val dt_names = map fst dt_infos; @@ -319,7 +328,7 @@ [((Binding.empty, [nth inducts index]), [Induct.induct_type tname]), ((Binding.empty, [nth exhaust index]), [Induct.cases_type tname])]) dt_names); val unnamed_rules = map (fn induct => - ((Binding.empty, [induct]), [Thm.kind_internal, Induct.induct_type ""])) + ((Binding.empty, [induct]), [Rule_Cases.inner_rule, Induct.induct_type ""])) (Library.drop (length dt_names, inducts)); in thy9 @@ -337,14 +346,16 @@ |> snd |> add_case_tr' case_names |> register dt_infos - |> DatatypeInterpretation.data (config, dt_names) + |> Datatype_Interpretation.data (config, dt_names) |> pair dt_names end; + (** declare existing type as datatype **) -fun prove_rep_datatype config dt_names alt_names descr sorts raw_inject half_distinct raw_induct thy1 = +fun prove_rep_datatype config dt_names alt_names descr sorts + raw_inject half_distinct raw_induct thy1 = let val raw_distinct = (map o maps) (fn thm => [thm, thm RS not_sym]) half_distinct; val new_type_names = map Long_Name.base_name (the_default dt_names alt_names); @@ -417,7 +428,8 @@ (*FIXME somehow dubious*) in ProofContext.theory_result - (prove_rep_datatype config dt_names alt_names descr vs raw_inject half_distinct raw_induct) + (prove_rep_datatype config dt_names alt_names descr vs + raw_inject half_distinct raw_induct) #-> after_qed end; in @@ -430,6 +442,7 @@ val rep_datatype_cmd = gen_rep_datatype Syntax.read_term_global default_config (K I); + (** definitional introduction of datatypes **) fun gen_add_datatype prep_typ config new_type_names dts thy = @@ -445,16 +458,20 @@ val (tyvars, _, _, _)::_ = dts; val (new_dts, types_syntax) = ListPair.unzip (map (fn (tvs, tname, mx, _) => let val full_tname = Sign.full_name tmp_thy (Binding.map_name (Syntax.type_name mx) tname) - in (case duplicates (op =) tvs of - [] => if eq_set (op =) (tyvars, tvs) then ((full_tname, tvs), (tname, mx)) - else error ("Mutually recursive datatypes must have same type parameters") - | dups => error ("Duplicate parameter(s) for datatype " ^ quote (Binding.str_of tname) ^ - " : " ^ commas dups)) + in + (case duplicates (op =) tvs of + [] => + if eq_set (op =) (tyvars, tvs) then ((full_tname, tvs), (tname, mx)) + else error ("Mutually recursive datatypes must have same type parameters") + | dups => error ("Duplicate parameter(s) for datatype " ^ quote (Binding.str_of tname) ^ + " : " ^ commas dups)) end) dts); val dt_names = map fst new_dts; - val _ = (case duplicates (op =) (map fst new_dts) @ duplicates (op =) new_type_names of - [] => () | dups => error ("Duplicate datatypes: " ^ commas dups)); + val _ = + (case duplicates (op =) (map fst new_dts) @ duplicates (op =) new_type_names of + [] => () + | dups => error ("Duplicate datatypes: " ^ commas dups)); fun prep_dt_spec ((tvs, tname, mx, constrs), tname') (dts', constr_syntax, sorts, i) = let @@ -508,13 +525,15 @@ val datatype_cmd = snd ooo gen_add_datatype read_typ default_config; + (** package setup **) (* setup theory *) val setup = trfun_setup #> - DatatypeInterpretation.init; + Datatype_Interpretation.init; + (* outer syntax *) diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/Tools/Datatype/datatype_abs_proofs.ML --- a/src/HOL/Tools/Datatype/datatype_abs_proofs.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/HOL/Tools/Datatype/datatype_abs_proofs.ML Wed Nov 04 11:40:59 2009 +0100 @@ -125,7 +125,7 @@ fun make_rec_intr T rec_set (cname, cargs) (rec_intr_ts, l) = let - fun mk_prem ((dt, U), (j, k, prems, t1s, t2s)) = + fun mk_prem (dt, U) (j, k, prems, t1s, t2s) = let val free1 = mk_Free "x" U j in (case (strip_dtyp dt, strip_type U) of ((_, DtRec m), (Us, _)) => @@ -141,7 +141,7 @@ end; val Ts = map (typ_of_dtyp descr' sorts) cargs; - val (_, _, prems, t1s, t2s) = List.foldr mk_prem (1, 1, [], [], []) (cargs ~~ Ts) + val (_, _, prems, t1s, t2s) = fold_rev mk_prem (cargs ~~ Ts) (1, 1, [], [], []) in (rec_intr_ts @ [Logic.list_implies (prems, HOLogic.mk_Trueprop (rec_set $ list_comb (Const (cname, Ts ---> T), t1s) $ @@ -292,7 +292,7 @@ val case_dummy_fns = map (fn (_, (_, _, constrs)) => map (fn (_, cargs) => let val Ts = map (typ_of_dtyp descr' sorts) cargs; - val Ts' = map mk_dummyT (List.filter is_rec_type cargs) + val Ts' = map mk_dummyT (filter is_rec_type cargs) in Const (@{const_name undefined}, Ts @ Ts' ---> T') end) constrs) descr'; @@ -305,7 +305,7 @@ val (fns1, fns2) = split_list (map (fn ((_, cargs), j) => let val Ts = map (typ_of_dtyp descr' sorts) cargs; - val Ts' = Ts @ map mk_dummyT (List.filter is_rec_type cargs); + val Ts' = Ts @ map mk_dummyT (filter is_rec_type cargs); val frees' = map (uncurry (mk_Free "x")) (Ts' ~~ (1 upto length Ts')); val frees = Library.take (length cargs, frees'); val free = mk_Free "f" (Ts ---> T') j diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/Tools/Datatype/datatype_aux.ML --- a/src/HOL/Tools/Datatype/datatype_aux.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/HOL/Tools/Datatype/datatype_aux.ML Wed Nov 04 11:40:59 2009 +0100 @@ -162,8 +162,7 @@ val prem' = hd (prems_of exhaustion); val _ $ (_ $ lhs $ _) = hd (rev (Logic.strip_assums_hyp prem')); val exhaustion' = cterm_instantiate [(cterm_of thy (head_of lhs), - cterm_of thy (List.foldr (fn ((_, T), t) => Abs ("z", T, t)) - (Bound 0) params))] exhaustion + cterm_of thy (fold_rev (fn (_, T) => fn t => Abs ("z", T, t)) params (Bound 0)))] exhaustion in compose_tac (false, exhaustion', nprems_of exhaustion) i state end; @@ -231,7 +230,7 @@ fun name_of_typ (Type (s, Ts)) = let val s' = Long_Name.base_name s - in space_implode "_" (List.filter (not o equal "") (map name_of_typ Ts) @ + in space_implode "_" (filter_out (equal "") (map name_of_typ Ts) @ [if Syntax.is_identifier s' then s' else "x"]) end | name_of_typ _ = ""; @@ -272,7 +271,7 @@ fun get_arities descr = fold (fn (_, (_, _, constrs)) => fold (fn (_, cargs) => fold (insert op =) (map (length o fst o strip_dtyp) - (List.filter is_rec_type cargs))) constrs) descr []; + (filter is_rec_type cargs))) constrs) descr []; (* interpret construction of datatype *) diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/Tools/Datatype/datatype_codegen.ML --- a/src/HOL/Tools/Datatype/datatype_codegen.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/HOL/Tools/Datatype/datatype_codegen.ML Wed Nov 04 11:40:59 2009 +0100 @@ -42,8 +42,8 @@ fun add_dt_defs thy defs dep module (descr: Datatype.descr) sorts gr = let - val descr' = List.filter (can (map DatatypeAux.dest_DtTFree o #2 o snd)) descr; - val rtnames = map (#1 o snd) (List.filter (fn (_, (_, _, cs)) => + val descr' = filter (can (map DatatypeAux.dest_DtTFree o #2 o snd)) descr; + val rtnames = map (#1 o snd) (filter (fn (_, (_, _, cs)) => exists (exists DatatypeAux.is_rec_type o snd) cs) descr'); val (_, (tname, _, _)) :: _ = descr'; diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/Tools/Datatype/datatype_prop.ML --- a/src/HOL/Tools/Datatype/datatype_prop.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/HOL/Tools/Datatype/datatype_prop.ML Wed Nov 04 11:40:59 2009 +0100 @@ -130,11 +130,11 @@ (make_pred (body_index dt) U $ app_bnds (Free (s, T)) (length Us))) end; - val recs = List.filter is_rec_type cargs; + val recs = filter is_rec_type cargs; val Ts = map (typ_of_dtyp descr' sorts) cargs; val recTs' = map (typ_of_dtyp descr' sorts) recs; val tnames = Name.variant_list pnames (make_tnames Ts); - val rec_tnames = map fst (List.filter (is_rec_type o snd) (tnames ~~ cargs)); + val rec_tnames = map fst (filter (is_rec_type o snd) (tnames ~~ cargs)); val frees = tnames ~~ Ts; val prems = map mk_prem (recs ~~ rec_tnames ~~ recTs'); @@ -190,7 +190,7 @@ map (fn (_, cargs) => let val Ts = map (typ_of_dtyp descr' sorts) cargs; - val recs = List.filter (is_rec_type o fst) (cargs ~~ Ts); + val recs = filter (is_rec_type o fst) (cargs ~~ Ts); fun mk_argT (dt, T) = binder_types T ---> List.nth (rec_result_Ts, body_index dt); @@ -223,11 +223,11 @@ fun make_primrec T comb_t (cname, cargs) (ts, f::fs) = let - val recs = List.filter is_rec_type cargs; + val recs = filter is_rec_type cargs; val Ts = map (typ_of_dtyp descr' sorts) cargs; val recTs' = map (typ_of_dtyp descr' sorts) recs; val tnames = make_tnames Ts; - val rec_tnames = map fst (List.filter (is_rec_type o snd) (tnames ~~ cargs)); + val rec_tnames = map fst (filter (is_rec_type o snd) (tnames ~~ cargs)); val frees = map Free (tnames ~~ Ts); val frees' = map Free (rec_tnames ~~ recTs'); @@ -313,20 +313,20 @@ val (_, fs) = strip_comb comb_t; val used = ["P", "x"] @ (map (fst o dest_Free) fs); - fun process_constr (((cname, cargs), f), (t1s, t2s)) = + fun process_constr ((cname, cargs), f) (t1s, t2s) = let val Ts = map (typ_of_dtyp descr' sorts) cargs; val frees = map Free (Name.variant_list used (make_tnames Ts) ~~ Ts); val eqn = HOLogic.mk_eq (Free ("x", T), list_comb (Const (cname, Ts ---> T), frees)); val P' = P $ list_comb (f, frees) - in ((List.foldr (fn (Free (s, T), t) => HOLogic.mk_all (s, T, t)) - (HOLogic.imp $ eqn $ P') frees)::t1s, - (List.foldr (fn (Free (s, T), t) => HOLogic.mk_exists (s, T, t)) - (HOLogic.conj $ eqn $ (HOLogic.Not $ P')) frees)::t2s) + in (fold_rev (fn Free (s, T) => fn t => HOLogic.mk_all (s, T, t)) frees + (HOLogic.imp $ eqn $ P') :: t1s, + fold_rev (fn Free (s, T) => fn t => HOLogic.mk_exists (s, T, t)) frees + (HOLogic.conj $ eqn $ (HOLogic.Not $ P')) :: t2s) end; - val (t1s, t2s) = List.foldr process_constr ([], []) (constrs ~~ fs); + val (t1s, t2s) = fold_rev process_constr (constrs ~~ fs) ([], []); val lhs = P $ (comb_t $ Free ("x", T)) in (HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, mk_conj t1s)), @@ -423,9 +423,9 @@ val tnames = Name.variant_list ["v"] (make_tnames Ts); val frees = tnames ~~ Ts in - List.foldr (fn ((s, T'), t) => HOLogic.mk_exists (s, T', t)) + fold_rev (fn (s, T') => fn t => HOLogic.mk_exists (s, T', t)) frees (HOLogic.mk_eq (Free ("v", T), - list_comb (Const (cname, Ts ---> T), map Free frees))) frees + list_comb (Const (cname, Ts ---> T), map Free frees))) end in map (fn ((_, (_, _, constrs)), T) => diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/Tools/Datatype/datatype_realizer.ML --- a/src/HOL/Tools/Datatype/datatype_realizer.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/HOL/Tools/Datatype/datatype_realizer.ML Wed Nov 04 11:40:59 2009 +0100 @@ -138,21 +138,24 @@ tname_of (body_type T) mem ["set", "bool"]) ivs); val ivs2 = map (fn (ixn, _) => Var (ixn, the (AList.lookup (op =) rvs ixn))) ivs; - val prf = List.foldr forall_intr_prf - (List.foldr (fn ((f, p), prf) => - (case head_of (strip_abs_body f) of - Free (s, T) => - let val T' = Logic.varifyT T - in Abst (s, SOME T', Proofterm.prf_abstract_over - (Var ((s, 0), T')) (AbsP ("H", SOME p, prf))) - end - | _ => AbsP ("H", SOME p, prf))) - (Proofterm.proof_combP - (prf_of thm', map PBound (length prems - 1 downto 0))) (rec_fns ~~ prems_of thm)) ivs2; + val prf = + List.foldr forall_intr_prf + (fold_rev (fn (f, p) => fn prf => + (case head_of (strip_abs_body f) of + Free (s, T) => + let val T' = Logic.varifyT T + in Abst (s, SOME T', Proofterm.prf_abstract_over + (Var ((s, 0), T')) (AbsP ("H", SOME p, prf))) + end + | _ => AbsP ("H", SOME p, prf))) + (rec_fns ~~ prems_of thm) + (Proofterm.proof_combP (prf_of thm', map PBound (length prems - 1 downto 0)))) ivs2; - val r' = if null is then r else Logic.varify (List.foldr (uncurry lambda) - r (map Logic.unvarify ivs1 @ filter_out is_unit - (map (head_of o strip_abs_body) rec_fns))); + val r' = + if null is then r + else Logic.varify (fold_rev lambda + (map Logic.unvarify ivs1 @ filter_out is_unit + (map (head_of o strip_abs_body) rec_fns)) r); in Extraction.add_realizers_i [(ind_name, (vs, r', prf))] thy' end; @@ -200,10 +203,10 @@ val P = Var (("P", 0), rT' --> HOLogic.boolT); val prf = forall_intr_prf (y, forall_intr_prf (P, - List.foldr (fn ((p, r), prf) => - forall_intr_prf (Logic.varify r, AbsP ("H", SOME (Logic.varify p), - prf))) (Proofterm.proof_combP (prf_of thm', - map PBound (length prems - 1 downto 0))) (prems ~~ rs))); + fold_rev (fn (p, r) => fn prf => + forall_intr_prf (Logic.varify r, AbsP ("H", SOME (Logic.varify p), prf))) + (prems ~~ rs) + (Proofterm.proof_combP (prf_of thm', map PBound (length prems - 1 downto 0))))); val r' = Logic.varify (Abs ("y", T, list_abs (map dest_Free rs, list_comb (r, map Bound ((length rs - 1 downto 0) @ [length rs]))))); diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/Tools/Datatype/datatype_rep_proofs.ML --- a/src/HOL/Tools/Datatype/datatype_rep_proofs.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/HOL/Tools/Datatype/datatype_rep_proofs.ML Wed Nov 04 11:40:59 2009 +0100 @@ -73,8 +73,9 @@ val branchT = if null branchTs then HOLogic.unitT else Balanced_Tree.make (fn (T, U) => Type ("+", [T, U])) branchTs; val arities = remove (op =) 0 (get_arities descr'); - val unneeded_vars = subtract (op =) (List.foldr OldTerm.add_typ_tfree_names [] (leafTs' @ branchTs)) (hd tyvars); - val leafTs = leafTs' @ (map (fn n => TFree (n, (the o AList.lookup (op =) sorts) n)) unneeded_vars); + val unneeded_vars = + subtract (op =) (List.foldr OldTerm.add_typ_tfree_names [] (leafTs' @ branchTs)) (hd tyvars); + val leafTs = leafTs' @ map (fn n => TFree (n, (the o AList.lookup (op =) sorts) n)) unneeded_vars; val recTs = get_rec_types descr' sorts; val newTs = Library.take (length (hd descr), recTs); val oldTs = Library.drop (length (hd descr), recTs); @@ -133,7 +134,7 @@ in mk_inj branchT (length branchTs) (1 + find_index (fn T'' => T'' = T') branchTs) end; - val mk_lim = List.foldr (fn (T, t) => Lim $ mk_fun_inj T (Abs ("x", T, t))); + fun mk_lim t Ts = fold_rev (fn T => fn t => Lim $ mk_fun_inj T (Abs ("x", T, t))) Ts t; (************** generate introduction rules for representing set **********) @@ -143,7 +144,8 @@ fun make_intr s n (i, (_, cargs)) = let - fun mk_prem (dt, (j, prems, ts)) = (case strip_dtyp dt of + fun mk_prem dt (j, prems, ts) = + (case strip_dtyp dt of (dts, DtRec k) => let val Ts = map (typ_of_dtyp descr' sorts) dts; @@ -159,7 +161,7 @@ in (j + 1, prems, (Leaf $ mk_inj T (mk_Free "x" T j))::ts) end); - val (_, prems, ts) = List.foldr mk_prem (1, [], []) cargs; + val (_, prems, ts) = fold_rev mk_prem cargs (1, [], []); val concl = HOLogic.mk_Trueprop (Free (s, UnivT') $ mk_univ_inj ts n i) in Logic.list_implies (prems, concl) @@ -213,7 +215,7 @@ fun make_constr_def tname T n ((cname, cargs), (cname', mx)) (thy, defs, eqns, i) = let - fun constr_arg (dt, (j, l_args, r_args)) = + fun constr_arg dt (j, l_args, r_args) = let val T = typ_of_dtyp descr' sorts dt; val free_t = mk_Free "x" T j in (case (strip_dtyp dt, strip_type T) of @@ -223,7 +225,7 @@ | _ => (j + 1, free_t::l_args, (Leaf $ mk_inj T free_t)::r_args)) end; - val (_, l_args, r_args) = List.foldr constr_arg (1, [], []) cargs; + val (_, l_args, r_args) = fold_rev constr_arg cargs (1, [], []); val constrT = (map (typ_of_dtyp descr' sorts) cargs) ---> T; val abs_name = Sign.intern_const thy ("Abs_" ^ tname); val rep_name = Sign.intern_const thy ("Rep_" ^ tname); @@ -387,7 +389,7 @@ val fun_congs = map (fn T => make_elim (Drule.instantiate' [SOME (ctyp_of thy5 T)] [] fun_cong)) branchTs; - fun prove_iso_thms (ds, (inj_thms, elem_thms)) = + fun prove_iso_thms ds (inj_thms, elem_thms) = let val (_, (tname, _, _)) = hd ds; val induct = (#induct o the o Symtab.lookup dt_info) tname; @@ -445,8 +447,8 @@ in (inj_thms'' @ inj_thms, elem_thms @ (split_conj_thm elem_thm)) end; - val (iso_inj_thms_unfolded, iso_elem_thms) = List.foldr prove_iso_thms - ([], map #3 newT_iso_axms) (tl descr); + val (iso_inj_thms_unfolded, iso_elem_thms) = + fold_rev prove_iso_thms (tl descr) ([], map #3 newT_iso_axms); val iso_inj_thms = map snd newT_iso_inj_thms @ map (fn r => r RS @{thm injD}) iso_inj_thms_unfolded; diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/Tools/Function/function.ML --- a/src/HOL/Tools/Function/function.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/HOL/Tools/Function/function.ML Wed Nov 04 11:40:59 2009 +0100 @@ -44,12 +44,12 @@ [Simplifier.simp_add, Nitpick_Psimps.add] -fun note_theorem ((name, atts), ths) = - LocalTheory.note Thm.generatedK ((Binding.qualified_name name, atts), ths) +fun note_theorem ((binding, atts), ths) = + LocalTheory.note Thm.generatedK ((binding, atts), ths) fun mk_defname fixes = fixes |> map (fst o fst) |> space_implode "_" -fun add_simps fnames post sort extra_qualify label moreatts simps lthy = +fun add_simps fnames post sort extra_qualify label mod_binding moreatts simps lthy = let val spec = post simps |> map (apfst (apsnd (fn ats => moreatts @ ats))) @@ -62,7 +62,8 @@ val simps_by_f = sort saved_simps fun add_for_f fname simps = - note_theorem ((Long_Name.qualify fname label, []), simps) #> snd + note_theorem ((mod_binding (Binding.qualify true fname (Binding.name label)), []), simps) + #> snd in (saved_simps, fold2 add_for_f fnames simps_by_f lthy) @@ -89,21 +90,24 @@ cont (Thm.close_derivation proof) val fnames = map (fst o fst) fixes - val qualify = Long_Name.qualify defname + fun qualify n = Binding.name n + |> Binding.qualify true defname + val conceal_partial = if partials then I else Binding.conceal + val addsmps = add_simps fnames post sort_cont val (((psimps', pinducts'), (_, [termination'])), lthy) = lthy - |> addsmps (Binding.qualify false "partial") "psimps" - psimp_attribs psimps - ||> fold_option (snd oo addsmps I "simps" simp_attribs) trsimps - ||>> note_theorem ((qualify "pinduct", - [Attrib.internal (K (RuleCases.case_names cnames)), - Attrib.internal (K (RuleCases.consumes 1)), + |> addsmps (conceal_partial o Binding.qualify false "partial") + "psimps" conceal_partial psimp_attribs psimps + ||> fold_option (snd oo addsmps I "simps" I simp_attribs) trsimps + ||>> note_theorem ((conceal_partial (qualify "pinduct"), + [Attrib.internal (K (Rule_Cases.case_names cnames)), + Attrib.internal (K (Rule_Cases.consumes 1)), Attrib.internal (K (Induct.induct_pred ""))]), simple_pinducts) - ||>> note_theorem ((qualify "termination", []), [termination]) + ||>> note_theorem ((Binding.conceal (qualify "termination"), []), [termination]) ||> (snd o note_theorem ((qualify "cases", - [Attrib.internal (K (RuleCases.case_names cnames))]), [cases])) + [Attrib.internal (K (Rule_Cases.case_names cnames))]), [cases])) ||> fold_option (snd oo curry note_theorem (qualify "domintros", [])) domintros val cdata = FunctionCtxData { add_simps=addsmps, case_names=cnames, psimps=psimps', @@ -147,23 +151,24 @@ full_simplify (HOL_basic_ss addsimps [totality, True_implies_equals]) val tsimps = map remove_domain_condition psimps val tinduct = map remove_domain_condition pinducts - val qualify = Long_Name.qualify defname; + fun qualify n = Binding.name n + |> Binding.qualify true defname in lthy - |> add_simps I "simps" simp_attribs tsimps |> snd + |> add_simps I "simps" I simp_attribs tsimps |> snd |> note_theorem ((qualify "induct", - [Attrib.internal (K (RuleCases.case_names case_names))]), + [Attrib.internal (K (Rule_Cases.case_names case_names))]), tinduct) |> snd end in lthy |> ProofContext.note_thmss "" - [((Binding.empty, [ContextRules.rule_del]), [([allI], [])])] |> snd + [((Binding.empty, [Context_Rules.rule_del]), [([allI], [])])] |> snd |> ProofContext.note_thmss "" - [((Binding.empty, [ContextRules.intro_bang (SOME 1)]), [([allI], [])])] |> snd + [((Binding.empty, [Context_Rules.intro_bang (SOME 1)]), [([allI], [])])] |> snd |> ProofContext.note_thmss "" - [((Binding.name "termination", [ContextRules.intro_bang (SOME 0)]), + [((Binding.name "termination", [Context_Rules.intro_bang (SOME 0)]), [([Goal.norm_result termination], [])])] |> snd |> Proof.theorem_i NONE afterqed [[(goal, [])]] end diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/Tools/Function/function_common.ML --- a/src/HOL/Tools/Function/function_common.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/HOL/Tools/Function/function_common.ML Wed Nov 04 11:40:59 2009 +0100 @@ -65,8 +65,9 @@ defname : string, (* contains no logical entities: invariant under morphisms *) - add_simps : (binding -> binding) -> string -> Attrib.src list -> thm list - -> local_theory -> thm list * local_theory, + add_simps : (binding -> binding) -> string -> (binding -> binding) -> + Attrib.src list -> thm list -> local_theory -> thm list * local_theory, + case_names : string list, fs : term list, @@ -92,9 +93,7 @@ structure FunctionData = GenericDataFun ( type T = (term * function_context_data) Item_Net.T; - val empty = Item_Net.init - (op aconv o pairself fst : (term * function_context_data) * (term * function_context_data) -> bool) - fst; + val empty : T = Item_Net.init (op aconv o pairself fst) (single o fst); val copy = I; val extend = I; fun merge _ (tab1, tab2) = Item_Net.merge (tab1, tab2) @@ -138,7 +137,7 @@ val all_function_data = Item_Net.content o get_function fun add_function_data (data as FunctionCtxData {fs, termination, ...}) = - FunctionData.map (fold (fn f => Item_Net.insert (f, data)) fs) + FunctionData.map (fold (fn f => Item_Net.update (f, data)) fs) #> store_termination_rule termination diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/Tools/Function/function_core.ML --- a/src/HOL/Tools/Function/function_core.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/HOL/Tools/Function/function_core.ML Wed Nov 04 11:40:59 2009 +0100 @@ -451,72 +451,99 @@ (goalstate, values) end +(* wrapper -- restores quantifiers in rule specifications *) +fun inductive_def (binding as ((R, T), _)) intrs lthy = + let + val ({intrs = intrs_gen, elims = [elim_gen], preds = [ Rdef ], induct, ...}, lthy) = + lthy + |> LocalTheory.conceal + |> Inductive.add_inductive_i + {quiet_mode = true, + verbose = ! trace, + kind = Thm.internalK, + alt_name = Binding.empty, + coind = false, + no_elim = false, + no_ind = false, + skip_mono = true, + fork_mono = false} + [binding] (* relation *) + [] (* no parameters *) + (map (fn t => (Attrib.empty_binding, t)) intrs) (* intro rules *) + [] (* no special monos *) + ||> LocalTheory.restore_naming lthy + + val cert = cterm_of (ProofContext.theory_of lthy) + fun requantify orig_intro thm = + let + val (qs, t) = dest_all_all orig_intro + val frees = frees_in_term lthy t |> remove (op =) (Binding.name_of R, T) + val vars = Term.add_vars (prop_of thm) [] |> rev + val varmap = AList.lookup (op =) (frees ~~ map fst vars) + #> the_default ("",0) + in + fold_rev (fn Free (n, T) => + forall_intr_rename (n, cert (Var (varmap (n, T), T)))) qs thm + end + in + ((Rdef, map2 requantify intrs intrs_gen, forall_intr_vars elim_gen, induct), lthy) + end fun define_graph Gname fvar domT ranT clauses RCss lthy = - let - val GT = domT --> ranT --> boolT - val Gvar = Free (the_single (Variable.variant_frees lthy [] [(Gname, GT)])) + let + val GT = domT --> ranT --> boolT + val (Gvar as (n, T)) = singleton (Variable.variant_frees lthy []) (Gname, GT) - fun mk_GIntro (ClauseContext {qs, gs, lhs, rhs, ...}) RCs = - let - fun mk_h_assm (rcfix, rcassm, rcarg) = - HOLogic.mk_Trueprop (Gvar $ rcarg $ (fvar $ rcarg)) - |> fold_rev (curry Logic.mk_implies o prop_of) rcassm - |> fold_rev (Logic.all o Free) rcfix - in - HOLogic.mk_Trueprop (Gvar $ lhs $ rhs) - |> fold_rev (curry Logic.mk_implies o mk_h_assm) RCs - |> fold_rev (curry Logic.mk_implies) gs - |> fold_rev Logic.all (fvar :: qs) - end + fun mk_GIntro (ClauseContext {qs, gs, lhs, rhs, ...}) RCs = + let + fun mk_h_assm (rcfix, rcassm, rcarg) = + HOLogic.mk_Trueprop (Free Gvar $ rcarg $ (fvar $ rcarg)) + |> fold_rev (curry Logic.mk_implies o prop_of) rcassm + |> fold_rev (Logic.all o Free) rcfix + in + HOLogic.mk_Trueprop (Free Gvar $ lhs $ rhs) + |> fold_rev (curry Logic.mk_implies o mk_h_assm) RCs + |> fold_rev (curry Logic.mk_implies) gs + |> fold_rev Logic.all (fvar :: qs) + end - val G_intros = map2 mk_GIntro clauses RCss - - val (GIntro_thms, (G, G_elim, G_induct, lthy)) = - Function_Inductive_Wrap.inductive_def G_intros ((dest_Free Gvar, NoSyn), lthy) - in - ((G, GIntro_thms, G_elim, G_induct), lthy) - end - - + val G_intros = map2 mk_GIntro clauses RCss + in + inductive_def ((Binding.name n, T), NoSyn) G_intros lthy + end fun define_function fdefname (fname, mixfix) domT ranT G default lthy = - let - val f_def = - Abs ("x", domT, Const (@{const_name FunDef.THE_default}, ranT --> (ranT --> boolT) --> ranT) $ (default $ Bound 0) $ - Abs ("y", ranT, G $ Bound 1 $ Bound 0)) - |> Syntax.check_term lthy - - val ((f, (_, f_defthm)), lthy) = - LocalTheory.define Thm.internalK - ((Binding.name (function_name fname), mixfix), - ((Binding.conceal (Binding.name fdefname), []), f_def)) lthy - in - ((f, f_defthm), lthy) - end - + let + val f_def = + Abs ("x", domT, Const (@{const_name FunDef.THE_default}, ranT --> (ranT --> boolT) --> ranT) + $ (default $ Bound 0) $ Abs ("y", ranT, G $ Bound 1 $ Bound 0)) + |> Syntax.check_term lthy + in + LocalTheory.define Thm.internalK + ((Binding.name (function_name fname), mixfix), + ((Binding.conceal (Binding.name fdefname), []), f_def)) lthy + end fun define_recursion_relation Rname domT ranT fvar f qglrs clauses RCss lthy = - let - - val RT = domT --> domT --> boolT - val Rvar = Free (the_single (Variable.variant_frees lthy [] [(Rname, RT)])) + let + val RT = domT --> domT --> boolT + val (Rvar as (n, T)) = singleton (Variable.variant_frees lthy []) (Rname, RT) - fun mk_RIntro (ClauseContext {qs, gs, lhs, ...}, (oqs, _, _, _)) (rcfix, rcassm, rcarg) = - HOLogic.mk_Trueprop (Rvar $ rcarg $ lhs) - |> fold_rev (curry Logic.mk_implies o prop_of) rcassm - |> fold_rev (curry Logic.mk_implies) gs - |> fold_rev (Logic.all o Free) rcfix - |> fold_rev mk_forall_rename (map fst oqs ~~ qs) - (* "!!qs xs. CS ==> G => (r, lhs) : R" *) + fun mk_RIntro (ClauseContext {qs, gs, lhs, ...}, (oqs, _, _, _)) (rcfix, rcassm, rcarg) = + HOLogic.mk_Trueprop (Free Rvar $ rcarg $ lhs) + |> fold_rev (curry Logic.mk_implies o prop_of) rcassm + |> fold_rev (curry Logic.mk_implies) gs + |> fold_rev (Logic.all o Free) rcfix + |> fold_rev mk_forall_rename (map fst oqs ~~ qs) + (* "!!qs xs. CS ==> G => (r, lhs) : R" *) - val R_intross = map2 (map o mk_RIntro) (clauses ~~ qglrs) RCss + val R_intross = map2 (map o mk_RIntro) (clauses ~~ qglrs) RCss - val (RIntro_thmss, (R, R_elim, _, lthy)) = - fold_burrow Function_Inductive_Wrap.inductive_def R_intross ((dest_Free Rvar, NoSyn), lthy) - in - ((R, RIntro_thmss, R_elim), lthy) - end + val ((R, RIntro_thms, R_elim, _), lthy) = + inductive_def ((Binding.name n, T), NoSyn) (flat R_intross) lthy + in + ((R, Library.unflat R_intross RIntro_thms, R_elim), lthy) + end fun fix_globals domT ranT fvar ctxt = @@ -892,7 +919,7 @@ val ((G, GIntro_thms, G_elim, G_induct), lthy) = PROFILE "def_graph" (define_graph (graph_name defname) fvar domT ranT clauses RCss) lthy - val ((f, f_defthm), lthy) = + val ((f, (_, f_defthm)), lthy) = PROFILE "def_fun" (define_function (defname ^ "_sumC_def") (fname, mixfix) domT ranT G default) lthy val RCss = map (map (inst_RC (ProofContext.theory_of lthy) fvar f)) RCss diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/Tools/Function/inductive_wrap.ML --- a/src/HOL/Tools/Function/inductive_wrap.ML Thu Oct 29 16:23:57 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,69 +0,0 @@ -(* Title: HOL/Tools/Function/inductive_wrap.ML - Author: Alexander Krauss, TU Muenchen - - -A wrapper around the inductive package, restoring the quantifiers in -the introduction and elimination rules. -*) - -signature FUNCTION_INDUCTIVE_WRAP = -sig - val inductive_def : term list - -> ((bstring * typ) * mixfix) * local_theory - -> thm list * (term * thm * thm * local_theory) -end - -structure Function_Inductive_Wrap: FUNCTION_INDUCTIVE_WRAP = -struct - -open Function_Lib - -fun requantify ctxt lfix orig_def thm = - let - val (qs, t) = dest_all_all orig_def - val thy = theory_of_thm thm - val frees = frees_in_term ctxt t - |> remove (op =) lfix - val vars = Term.add_vars (prop_of thm) [] |> rev - - val varmap = frees ~~ vars - in - fold_rev (fn Free (n, T) => - forall_intr_rename (n, cterm_of thy (Var (the_default (("",0), T) (AList.lookup (op =) varmap (n, T)))))) - qs - thm - end - - - -fun inductive_def defs (((R, T), mixfix), lthy) = - let - val ({intrs = intrs_gen, elims = [elim_gen], preds = [ Rdef ], induct, ...}, lthy) = - lthy - |> LocalTheory.conceal - |> Inductive.add_inductive_i - {quiet_mode = false, - verbose = ! Toplevel.debug, - kind = Thm.internalK, - alt_name = Binding.empty, - coind = false, - no_elim = false, - no_ind = false, - skip_mono = true, - fork_mono = false} - [((Binding.name R, T), NoSyn)] (* the relation *) - [] (* no parameters *) - (map (fn t => (Attrib.empty_binding, t)) defs) (* the intros *) - [] (* no special monos *) - ||> LocalTheory.restore_naming lthy - - val intrs = map2 (requantify lthy (R, T)) defs intrs_gen - - val elim = elim_gen - |> forall_intr_vars (* FIXME... *) - - in - (intrs, (Rdef, elim, induct, lthy)) - end - -end diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/Tools/Function/lexicographic_order.ML --- a/src/HOL/Tools/Function/lexicographic_order.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/HOL/Tools/Function/lexicographic_order.ML Wed Nov 04 11:40:59 2009 +0100 @@ -6,8 +6,8 @@ signature LEXICOGRAPHIC_ORDER = sig - val lex_order_tac : Proof.context -> tactic -> tactic - val lexicographic_order_tac : Proof.context -> tactic + val lex_order_tac : bool -> Proof.context -> tactic -> tactic + val lexicographic_order_tac : bool -> Proof.context -> tactic val lexicographic_order : Proof.context -> Proof.method val setup: theory -> theory @@ -187,7 +187,7 @@ (** The Main Function **) -fun lex_order_tac ctxt solve_tac (st: thm) = +fun lex_order_tac quiet ctxt solve_tac (st: thm) = let val thy = ProofContext.theory_of ctxt val ((trueprop $ (wf $ rel)) :: tl) = prems_of st @@ -197,29 +197,32 @@ val measure_funs = MeasureFunctions.get_measure_functions ctxt domT (* 1: generate measures *) (* 2: create table *) - val table = map (fn t => map (mk_cell thy solve_tac (dest_term t)) measure_funs) tl - - val order = the (search_table table) (* 3: search table *) - handle Option => error (no_order_msg ctxt table tl measure_funs) - - val clean_table = map (fn x => map (nth x) order) table + val table = Par_List.map (fn t => Par_List.map (mk_cell thy solve_tac (dest_term t)) measure_funs) tl + in + case search_table table of + NONE => if quiet then no_tac st else error (no_order_msg ctxt table tl measure_funs) + | SOME order => + let + val clean_table = map (fn x => map (nth x) order) table + val relation = mk_measures domT (map (nth measure_funs) order) + val _ = if not quiet + then writeln ("Found termination order: " ^ quote (Syntax.string_of_term ctxt relation)) + else () - val relation = mk_measures domT (map (nth measure_funs) order) - val _ = writeln ("Found termination order: " ^ quote (Syntax.string_of_term ctxt relation)) - - in (* 4: proof reconstruction *) - st |> (PRIMITIVE (cterm_instantiate [(cterm_of thy rel, cterm_of thy relation)]) - THEN (REPEAT (rtac @{thm "wf_mlex"} 1)) - THEN (rtac @{thm "wf_empty"} 1) - THEN EVERY (map prove_row clean_table)) + in (* 4: proof reconstruction *) + st |> (PRIMITIVE (cterm_instantiate [(cterm_of thy rel, cterm_of thy relation)]) + THEN (REPEAT (rtac @{thm "wf_mlex"} 1)) + THEN (rtac @{thm "wf_empty"} 1) + THEN EVERY (map prove_row clean_table)) + end end -fun lexicographic_order_tac ctxt = +fun lexicographic_order_tac quiet ctxt = TRY (Function_Common.apply_termination_rule ctxt 1) - THEN lex_order_tac ctxt + THEN lex_order_tac quiet ctxt (auto_tac (clasimpset_of ctxt addsimps2 Function_Common.Termination_Simps.get ctxt)) -val lexicographic_order = SIMPLE_METHOD o lexicographic_order_tac +val lexicographic_order = SIMPLE_METHOD o lexicographic_order_tac false val setup = Method.setup @{binding lexicographic_order} diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/Tools/Function/mutual.ML --- a/src/HOL/Tools/Function/mutual.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/HOL/Tools/Function/mutual.ML Wed Nov 04 11:40:59 2009 +0100 @@ -148,7 +148,8 @@ val ((f, (_, f_defthm)), lthy') = LocalTheory.define Thm.internalK ((Binding.name fname, mixfix), - (apfst Binding.conceal Attrib.empty_binding, Term.subst_bound (fsum, f_def))) lthy + ((Binding.conceal (Binding.name (fname ^ "_def")), []), + Term.subst_bound (fsum, f_def))) lthy in (MutualPart {i=i, i'=i', fvar=(fname, fT), cargTs=cargTs, f_def=f_def, f=SOME f, f_defthm=SOME f_defthm }, diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/Tools/Function/scnp_reconstruct.ML --- a/src/HOL/Tools/Function/scnp_reconstruct.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/HOL/Tools/Function/scnp_reconstruct.ML Wed Nov 04 11:40:59 2009 +0100 @@ -39,7 +39,6 @@ struct val PROFILE = Function_Common.PROFILE -fun TRACE x = if ! Function_Common.profile then tracing x else () open ScnpSolve @@ -197,7 +196,7 @@ else if b <= a then @{thm pair_leqI2} else @{thm pair_leqI1} in rtac rule 1 THEN PRIMITIVE (Thm.elim_implies stored_thm) - THEN (if tag_flag then Arith_Data.verbose_arith_tac ctxt 1 else all_tac) + THEN (if tag_flag then Arith_Data.arith_tac ctxt 1 else all_tac) end fun steps_tac MAX strict lq lp = @@ -347,17 +346,13 @@ end) end - fun single_scnp_tac use_tags orders ctxt cont err_cont D = Termination.CALLS (fn (cs, i) => let val ms_configured = is_some (Multiset_Setup.get (ProofContext.theory_of ctxt)) val orders' = if ms_configured then orders else filter_out (curry op = MS) orders val gp = gen_probl D cs -(* val _ = TRACE ("SCNP instance: " ^ makestring gp)*) val certificate = generate_certificate use_tags orders' gp -(* val _ = TRACE ("Certificate: " ^ makestring certificate)*) - in case certificate of NONE => err_cont D i diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/Tools/Function/size.ML --- a/src/HOL/Tools/Function/size.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/HOL/Tools/Function/size.ML Wed Nov 04 11:40:59 2009 +0100 @@ -115,7 +115,7 @@ then HOLogic.zero else foldl1 plus (ts @ [HOLogic.Suc_zero]) in - List.foldr (fn (T, t') => Abs ("x", T, t')) t (Ts @ replicate k HOLogic.natT) + fold_rev (fn T => fn t' => Abs ("x", T, t')) (Ts @ replicate k HOLogic.natT) t end; val fs = maps (fn (_, (name, _, constrs)) => diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/Tools/Nitpick/nitpick_hol.ML --- a/src/HOL/Tools/Nitpick/nitpick_hol.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/HOL/Tools/Nitpick/nitpick_hol.ML Wed Nov 04 11:40:59 2009 +0100 @@ -1634,7 +1634,7 @@ val cached_wf_props : (term * bool) list Unsynchronized.ref = Unsynchronized.ref [] -val termination_tacs = [Lexicographic_Order.lex_order_tac, +val termination_tacs = [Lexicographic_Order.lex_order_tac true, ScnpReconstruct.sizechange_tac] (* extended_context -> const_table -> styp -> bool *) diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/Tools/Predicate_Compile/pred_compile_fun.ML --- a/src/HOL/Tools/Predicate_Compile/pred_compile_fun.ML Thu Oct 29 16:23:57 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,437 +0,0 @@ -(* Author: Lukas Bulwahn, TU Muenchen - -Preprocessing functions to predicates -*) - -signature PREDICATE_COMPILE_FUN = -sig -val define_predicates : (string * thm list) list -> theory -> (string * thm list) list * theory - val rewrite_intro : theory -> thm -> thm list - val setup_oracle : theory -> theory - val pred_of_function : theory -> string -> string option -end; - -structure Predicate_Compile_Fun : PREDICATE_COMPILE_FUN = -struct - - -(* Oracle for preprocessing *) - -val (oracle : (string * (cterm -> thm)) option Unsynchronized.ref) = Unsynchronized.ref NONE; - -fun the_oracle () = - case !oracle of - NONE => error "Oracle is not setup" - | SOME (_, oracle) => oracle - -val setup_oracle = Thm.add_oracle (Binding.name "pred_compile_preprocessing", I) #-> - (fn ora => fn thy => let val _ = (oracle := SOME ora) in thy end) - - -fun is_funtype (Type ("fun", [_, _])) = true - | is_funtype _ = false; - -fun is_Type (Type _) = true - | is_Type _ = false - -(* returns true if t is an application of an datatype constructor *) -(* which then consequently would be splitted *) -(* else false *) -(* -fun is_constructor thy t = - if (is_Type (fastype_of t)) then - (case DatatypePackage.get_datatype thy ((fst o dest_Type o fastype_of) t) of - NONE => false - | SOME info => (let - val constr_consts = maps (fn (_, (_, _, constrs)) => map fst constrs) (#descr info) - val (c, _) = strip_comb t - in (case c of - Const (name, _) => name mem_string constr_consts - | _ => false) end)) - else false -*) - -(* must be exported in code.ML *) -fun is_constr thy = is_some o Code.get_datatype_of_constr thy; - -(* Table from constant name (string) to term of inductive predicate *) -structure Pred_Compile_Preproc = TheoryDataFun -( - type T = string Symtab.table; - val empty = Symtab.empty; - val copy = I; - val extend = I; - fun merge _ = Symtab.merge (op =); -) - -fun pred_of_function thy name = Symtab.lookup (Pred_Compile_Preproc.get thy) name - -fun defined thy = Symtab.defined (Pred_Compile_Preproc.get thy) - - -fun transform_ho_typ (T as Type ("fun", _)) = - let - val (Ts, T') = strip_type T - in if T' = @{typ "bool"} then T else (Ts @ [T']) ---> HOLogic.boolT end -| transform_ho_typ t = t - -fun transform_ho_arg arg = - case (fastype_of arg) of - (T as Type ("fun", _)) => - (case arg of - Free (name, _) => Free (name, transform_ho_typ T) - | _ => error "I am surprised") -| _ => arg - -fun pred_type T = - let - val (Ts, T') = strip_type T - val Ts' = map transform_ho_typ Ts - in - (Ts' @ [T']) ---> HOLogic.boolT - end; - -(* FIXME: create new predicate name -- does not avoid nameclashing *) -fun pred_of f = - let - val (name, T) = dest_Const f - in - if (body_type T = @{typ bool}) then - (Free (Long_Name.base_name name ^ "P", T)) - else - (Free (Long_Name.base_name name ^ "P", pred_type T)) - end - -fun mk_param thy lookup_pred (t as Free (v, _)) = lookup_pred t - | mk_param thy lookup_pred t = - let - val _ = tracing ("called param with " ^ (Syntax.string_of_term_global thy t)) - in if Predicate_Compile_Aux.is_predT (fastype_of t) then - t - else - let - val (vs, body) = strip_abs t - val names = Term.add_free_names body [] - val vs_names = Name.variant_list names (map fst vs) - val vs' = map2 (curry Free) vs_names (map snd vs) - val body' = subst_bounds (rev vs', body) - val (f, args) = strip_comb body' - val resname = Name.variant (vs_names @ names) "res" - val resvar = Free (resname, body_type (fastype_of body')) - (*val P = case try lookup_pred f of SOME P => P | NONE => error "mk_param" - val pred_body = list_comb (P, args @ [resvar]) - *) - val pred_body = HOLogic.mk_eq (body', resvar) - val param = fold_rev lambda (vs' @ [resvar]) pred_body - in param end - end -(* creates the list of premises for every intro rule *) -(* theory -> term -> (string list, term list list) *) - -fun dest_code_eqn eqn = let - val (lhs, rhs) = Logic.dest_equals (Logic.unvarify (Thm.prop_of eqn)) - val (func, args) = strip_comb lhs -in ((func, args), rhs) end; - -fun string_of_typ T = Syntax.string_of_typ_global @{theory} T - -fun string_of_term t = - case t of - Const (c, T) => "Const (" ^ c ^ ", " ^ string_of_typ T ^ ")" - | Free (c, T) => "Free (" ^ c ^ ", " ^ string_of_typ T ^ ")" - | Var ((c, i), T) => "Var ((" ^ c ^ ", " ^ string_of_int i ^ "), " ^ string_of_typ T ^ ")" - | Bound i => "Bound " ^ string_of_int i - | Abs (x, T, t) => "Abs (" ^ x ^ ", " ^ string_of_typ T ^ ", " ^ string_of_term t ^ ")" - | t1 $ t2 => "(" ^ string_of_term t1 ^ ") $ (" ^ string_of_term t2 ^ ")" - -fun ind_package_get_nparams thy name = - case try (Inductive.the_inductive (ProofContext.init thy)) name of - SOME (_, result) => length (Inductive.params_of (#raw_induct result)) - | NONE => error ("No such predicate: " ^ quote name) - -(* TODO: does not work with higher order functions yet *) -fun mk_rewr_eq (func, pred) = - let - val (argTs, resT) = (strip_type (fastype_of func)) - val nctxt = - Name.make_context (Term.fold_aterms (fn Free (x, _) => insert (op =) x | _ => I) (func $ pred) []) - val (argnames, nctxt') = Name.variants (replicate (length argTs) "a") nctxt - val ([resname], nctxt'') = Name.variants ["r"] nctxt' - val args = map Free (argnames ~~ argTs) - val res = Free (resname, resT) - in Logic.mk_equals - (HOLogic.mk_eq (res, list_comb (func, args)), list_comb (pred, args @ [res])) - end; - -fun has_split_rule_cname @{const_name "nat_case"} = true - | has_split_rule_cname @{const_name "list_case"} = true - | has_split_rule_cname _ = false - -fun has_split_rule_term thy (Const (@{const_name "nat_case"}, _)) = true - | has_split_rule_term thy (Const (@{const_name "list_case"}, _)) = true - | has_split_rule_term thy _ = false - -fun has_split_rule_term' thy (Const (@{const_name "If"}, _)) = true - | has_split_rule_term' thy (Const (@{const_name "Let"}, _)) = true - | has_split_rule_term' thy c = has_split_rule_term thy c - -fun prepare_split_thm ctxt split_thm = - (split_thm RS @{thm iffD2}) - |> LocalDefs.unfold ctxt [@{thm atomize_conjL[symmetric]}, - @{thm atomize_all[symmetric]}, @{thm atomize_imp[symmetric]}] - -fun find_split_thm thy (Const (name, typ)) = - let - fun split_name str = - case first_field "." str - of (SOME (field, rest)) => field :: split_name rest - | NONE => [str] - val splitted_name = split_name name - in - if length splitted_name > 0 andalso - String.isSuffix "_case" (List.last splitted_name) - then - (List.take (splitted_name, length splitted_name - 1)) @ ["split"] - |> space_implode "." - |> PureThy.get_thm thy - |> SOME - handle ERROR msg => NONE - else NONE - end - | find_split_thm _ _ = NONE - -fun find_split_thm' thy (Const (@{const_name "If"}, _)) = SOME @{thm split_if} - | find_split_thm' thy (Const (@{const_name "Let"}, _)) = SOME @{thm refl} (* TODO *) - | find_split_thm' thy c = find_split_thm thy c - -fun strip_all t = (Term.strip_all_vars t, Term.strip_all_body t) - -fun folds_map f xs y = - let - fun folds_map' acc [] y = [(rev acc, y)] - | folds_map' acc (x :: xs) y = - maps (fn (x, y) => folds_map' (x :: acc) xs y) (f x y) - in - folds_map' [] xs y - end; - -fun mk_prems thy (lookup_pred, get_nparams) t (names, prems) = - let - fun mk_prems' (t as Const (name, T)) (names, prems) = - if is_constr thy name orelse (is_none (try lookup_pred t)) then - [(t, (names, prems))] - else [(lookup_pred t, (names, prems))] - | mk_prems' (t as Free (f, T)) (names, prems) = - [(lookup_pred t, (names, prems))] - | mk_prems' (t as Abs _) (names, prems) = - if Predicate_Compile_Aux.is_predT (fastype_of t) then - [(t, (names, prems))] else error "mk_prems': Abs " - (* mk_param *) - | mk_prems' t (names, prems) = - if Predicate_Compile_Aux.is_constrt thy t then - [(t, (names, prems))] - else - if has_split_rule_term' thy (fst (strip_comb t)) then - let - val (f, args) = strip_comb t - val split_thm = prepare_split_thm (ProofContext.init thy) (the (find_split_thm' thy f)) - (* TODO: contextify things - this line is to unvarify the split_thm *) - (*val ((_, [isplit_thm]), _) = Variable.import true [split_thm] (ProofContext.init thy)*) - val (assms, concl) = Logic.strip_horn (Thm.prop_of split_thm) - val (P, [split_t]) = strip_comb (HOLogic.dest_Trueprop concl) - val subst = Pattern.match thy (split_t, t) (Vartab.empty, Vartab.empty) - val (_, split_args) = strip_comb split_t - val match = split_args ~~ args - fun mk_prems_of_assm assm = - let - val (vTs, assm') = strip_all (Envir.beta_norm (Envir.subst_term subst assm)) - val var_names = Name.variant_list names (map fst vTs) - val vars = map Free (var_names ~~ (map snd vTs)) - val (prems', pre_res) = Logic.strip_horn (subst_bounds (rev vars, assm')) - val (_, [inner_t]) = strip_comb (HOLogic.dest_Trueprop pre_res) - in - mk_prems' inner_t (var_names @ names, prems' @ prems) - end - in - maps mk_prems_of_assm assms - end - else - let - val (f, args) = strip_comb t - (* TODO: special procedure for higher-order functions: split arguments in - simple types and function types *) - val resname = Name.variant names "res" - val resvar = Free (resname, body_type (fastype_of t)) - val names' = resname :: names - fun mk_prems'' (t as Const (c, _)) = - if is_constr thy c orelse (is_none (try lookup_pred t)) then - folds_map mk_prems' args (names', prems) |> - map - (fn (argvs, (names'', prems')) => - let - val prem = HOLogic.mk_Trueprop (HOLogic.mk_eq (resvar, list_comb (f, argvs))) - in (names'', prem :: prems') end) - else - let - val pred = lookup_pred t - val nparams = get_nparams pred - val (params, args) = chop nparams args - val params' = map (mk_param thy lookup_pred) params - in - folds_map mk_prems' args (names', prems) - |> map (fn (argvs, (names'', prems')) => - let - val prem = HOLogic.mk_Trueprop (list_comb (pred, params' @ argvs @ [resvar])) - in (names'', prem :: prems') end) - end - | mk_prems'' (t as Free (_, _)) = - let - (* higher order argument call *) - val pred = lookup_pred t - in - folds_map mk_prems' args (resname :: names, prems) - |> map (fn (argvs, (names', prems')) => - let - val prem = HOLogic.mk_Trueprop (list_comb (pred, argvs @ [resvar])) - in (names', prem :: prems') end) - end - | mk_prems'' t = - error ("Invalid term: " ^ Syntax.string_of_term_global thy t) - in - map (pair resvar) (mk_prems'' f) - end - in - mk_prems' t (names, prems) - end; - -(* assumption: mutual recursive predicates all have the same parameters. *) -fun define_predicates specs thy = - if forall (fn (const, _) => member (op =) (Symtab.keys (Pred_Compile_Preproc.get thy)) const) specs then - ([], thy) - else - let - val consts = map fst specs - val eqns = maps snd specs - (*val eqns = maps (Predicate_Compile_Preproc_Data.get_specification thy) consts*) - (* create prednames *) - val ((funs, argss), rhss) = map_split dest_code_eqn eqns |>> split_list - val argss' = map (map transform_ho_arg) argss - val pnames = map dest_Free (distinct (op =) (maps (filter (is_funtype o fastype_of)) argss')) - val preds = map pred_of funs - val prednames = map (fst o dest_Free) preds - val funnames = map (fst o dest_Const) funs - val fun_pred_names = (funnames ~~ prednames) - (* mapping from term (Free or Const) to term *) - fun lookup_pred (Const (name, T)) = - (case (Symtab.lookup (Pred_Compile_Preproc.get thy) name) of - SOME c => Const (c, pred_type T) - | NONE => - (case AList.lookup op = fun_pred_names name of - SOME f => Free (f, pred_type T) - | NONE => Const (name, T))) - | lookup_pred (Free (name, T)) = - if member op = (map fst pnames) name then - Free (name, transform_ho_typ T) - else - Free (name, T) - | lookup_pred t = - error ("lookup function is not defined for " ^ Syntax.string_of_term_global thy t) - - (* mapping from term (predicate term, not function term!) to int *) - fun get_nparams (Const (name, _)) = - the_default 0 (try (ind_package_get_nparams thy) name) - | get_nparams (Free (name, _)) = - (if member op = prednames name then - length pnames - else 0) - | get_nparams t = error ("No parameters for " ^ (Syntax.string_of_term_global thy t)) - - (* create intro rules *) - - fun mk_intros ((func, pred), (args, rhs)) = - if (body_type (fastype_of func) = @{typ bool}) then - (*TODO: preprocess predicate definition of rhs *) - [Logic.list_implies ([HOLogic.mk_Trueprop rhs], HOLogic.mk_Trueprop (list_comb (pred, args)))] - else - let - val names = Term.add_free_names rhs [] - in mk_prems thy (lookup_pred, get_nparams) rhs (names, []) - |> map (fn (resultt, (names', prems)) => - Logic.list_implies (prems, HOLogic.mk_Trueprop (list_comb (pred, args @ [resultt])))) - end - fun mk_rewr_thm (func, pred) = @{thm refl} - in - case try (maps mk_intros) ((funs ~~ preds) ~~ (argss' ~~ rhss)) of - NONE => ([], thy) - | SOME intr_ts => - if is_some (try (map (cterm_of thy)) intr_ts) then - let - val (ind_result, thy') = - Inductive.add_inductive_global (serial ()) - {quiet_mode = false, verbose = false, kind = Thm.internalK, - alt_name = Binding.empty, coind = false, no_elim = false, - no_ind = false, skip_mono = false, fork_mono = false} - (map (fn (s, T) => ((Binding.name s, T), NoSyn)) (distinct (op =) (map dest_Free preds))) - pnames - (map (fn x => (Attrib.empty_binding, x)) intr_ts) - [] thy - val prednames = map (fst o dest_Const) (#preds ind_result) - (* val rewr_thms = map mk_rewr_eq ((distinct (op =) funs) ~~ (#preds ind_result)) *) - (* add constants to my table *) - val specs = map (fn predname => (predname, filter (Predicate_Compile_Aux.is_intro predname) (#intrs ind_result))) prednames - val thy'' = Pred_Compile_Preproc.map (fold Symtab.update_new (consts ~~ prednames)) thy' - in - (specs, thy'') - end - else - let - val _ = tracing "Introduction rules of function_predicate are not welltyped" - in ([], thy) end - end - -(* preprocessing intro rules - uses oracle *) - -(* theory -> thm -> thm *) -fun rewrite_intro thy intro = - let - fun lookup_pred (Const (name, T)) = - (case (Symtab.lookup (Pred_Compile_Preproc.get thy) name) of - SOME c => Const (c, pred_type T) - | NONE => error ("Function " ^ name ^ " is not inductified")) - | lookup_pred (Free (name, T)) = Free (name, T) - | lookup_pred _ = error "lookup function is not defined!" - - fun get_nparams (Const (name, _)) = - the_default 0 (try (ind_package_get_nparams thy) name) - | get_nparams (Free _) = 0 - | get_nparams t = error ("No parameters for " ^ (Syntax.string_of_term_global thy t)) - - val intro_t = (Logic.unvarify o prop_of) intro - val (prems, concl) = Logic.strip_horn intro_t - val frees = map fst (Term.add_frees intro_t []) - fun rewrite prem names = - let - val t = (HOLogic.dest_Trueprop prem) - val (lit, mk_lit) = case try HOLogic.dest_not t of - SOME t => (t, HOLogic.mk_not) - | NONE => (t, I) - val (P, args) = (strip_comb lit) - in - folds_map ( - fn t => if (is_funtype (fastype_of t)) then (fn x => [(t, x)]) - else mk_prems thy (lookup_pred, get_nparams) t) args (names, []) - |> map (fn (resargs, (names', prems')) => - let - val prem' = HOLogic.mk_Trueprop (mk_lit (list_comb (P, resargs))) - in (prem'::prems', names') end) - end - val intro_ts' = folds_map rewrite prems frees - |> maps (fn (prems', frees') => - rewrite concl frees' - |> map (fn (concl'::conclprems, _) => - Logic.list_implies ((flat prems') @ conclprems, concl'))) - in - map (Drule.standard o the_oracle () o cterm_of thy) intro_ts' - end; - -end; diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/Tools/Predicate_Compile/predicate_compile.ML --- a/src/HOL/Tools/Predicate_Compile/predicate_compile.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/HOL/Tools/Predicate_Compile/predicate_compile.ML Wed Nov 04 11:40:59 2009 +0100 @@ -18,8 +18,6 @@ open Predicate_Compile_Aux; -val priority = tracing; - (* Some last processing *) fun remove_pointless_clauses intro = @@ -27,20 +25,21 @@ [] else [intro] -fun tracing s = () - fun print_intross options thy msg intross = if show_intermediate_results options then - Output.tracing (msg ^ - (space_implode "\n" (map - (fn (c, intros) => "Introduction rule(s) of " ^ c ^ ":\n" ^ - commas (map (Display.string_of_thm_global thy) intros)) intross))) + tracing (msg ^ + (space_implode "\n" (map + (fn (c, intros) => "Introduction rule(s) of " ^ c ^ ":\n" ^ + commas (map (Display.string_of_thm_global thy) intros)) intross))) else () fun print_specs thy specs = map (fn (c, thms) => "Constant " ^ c ^ " has specification:\n" ^ (space_implode "\n" (map (Display.string_of_thm_global thy) thms)) ^ "\n") specs +(* TODO: *) +fun overload_const thy s = s + fun map_specs f specs = map (fn (s, ths) => (s, f ths)) specs @@ -63,15 +62,14 @@ val (intross3, (new_defs, thy''')) = Predicate_Compile_Pred.flat_higher_order_arguments (intross2, thy'') val (new_intross, thy'''') = if not (null new_defs) then - let - val _ = print_step options "Recursively obtaining introduction rules for new definitions..." - in process_specification options new_defs thy''' end - else ([], thy''') + let + val _ = print_step options "Recursively obtaining introduction rules for new definitions..." + in process_specification options new_defs thy''' end + else ([], thy''') in (intross3 @ new_intross, thy'''') end - fun preprocess_strong_conn_constnames options gr constnames thy = let val get_specs = map (fn k => (k, Graph.get_node gr k)) @@ -89,7 +87,7 @@ val _ = print_intross options thy''' "Introduction rules with new constants: " intross3 val intross4 = map_specs (maps remove_pointless_clauses) intross3 val _ = print_intross options thy''' "After removing pointless clauses: " intross4 - (*val intross5 = map (fn s, ths) => ( s, map (AxClass.overload thy''') ths)) intross4*) + val intross5 = (*map (fn (s, ths) => (overload_const s, map (AxClass.overload thy''') ths))*) intross4 val intross6 = map_specs (map (expand_tuples thy''')) intross4 val _ = print_intross options thy''' "introduction rules before registering: " intross6 val _ = print_step options "Registering introduction rules..." @@ -101,7 +99,7 @@ fun preprocess options const thy = let val _ = print_step options "Fetching definitions from theory..." - val table = Predicate_Compile_Data.make_const_spec_table thy + val table = Predicate_Compile_Data.make_const_spec_table options thy val gr = Predicate_Compile_Data.obtain_specification_graph options thy table const val gr = Graph.subgraph (member (op =) (Graph.all_succs gr [const])) gr in fold_rev (preprocess_strong_conn_constnames options gr) @@ -122,7 +120,7 @@ show_compilation = chk "show_compilation", skip_proof = chk "skip_proof", inductify = chk "inductify", - rpred = chk "rpred", + random = chk "random", depth_limited = chk "depth_limited" } end @@ -151,15 +149,40 @@ val setup = Predicate_Compile_Fun.setup_oracle #> Predicate_Compile_Core.setup val bool_options = ["show_steps", "show_intermediate_results", "show_proof_trace", "show_modes", - "show_mode_inference", "show_compilation", "skip_proof", "inductify", "rpred", "depth_limited"] + "show_mode_inference", "show_compilation", "skip_proof", "inductify", "random", "depth_limited"] local structure P = OuterParse in +(*val parse_argmode' = P.nat >> rpair NONE || P.$$$ "(" |-- P.enum1 "," --| P.$$$ ")"*) +datatype raw_argmode = Argmode of string | Argmode_Tuple of string list + +val parse_argmode' = + ((Args.$$$ "i" || Args.$$$ "o") >> Argmode) || + (Args.$$$ "(" |-- P.enum1 "," (Args.$$$ "i" || Args.$$$ "o") --| Args.$$$ ")" >> Argmode_Tuple) + +fun mk_numeral_mode ss = flat (map_index (fn (i, s) => if s = "i" then [i + 1] else []) ss) + +val parse_smode' = P.$$$ "[" |-- P.enum1 "," parse_argmode' --| P.$$$ "]" + >> (fn m => flat (map_index + (fn (i, Argmode s) => if s = "i" then [(i + 1, NONE)] else [] + | (i, Argmode_Tuple ss) => [(i + 1, SOME (mk_numeral_mode ss))]) m)) + +val parse_smode = (P.$$$ "[" |-- P.enum "," P.nat --| P.$$$ "]") + >> map (rpair (NONE : int list option)) + +fun gen_parse_mode smode_parser = + (Scan.optional + ((P.enum "=>" ((Args.$$$ "X" >> K NONE) || (smode_parser >> SOME))) --| Args.$$$ "==>") []) + -- smode_parser + +val parse_mode = gen_parse_mode parse_smode + +val parse_mode' = gen_parse_mode parse_smode' + val opt_modes = Scan.optional (P.$$$ "(" |-- Args.$$$ "mode" |-- P.$$$ ":" |-- - P.enum1 "," (P.$$$ "[" |-- P.enum "," P.nat --| P.$$$ "]") - --| P.$$$ ")" >> SOME) NONE + P.enum1 "," (parse_mode || parse_mode') --| P.$$$ ")" >> SOME) NONE val scan_params = let @@ -170,8 +193,7 @@ val _ = OuterSyntax.local_theory_to_proof "code_pred" "prove equations for predicate specified by intro/elim rules" - OuterKeyword.thy_goal (opt_modes -- scan_params -- P.term_group >> - code_pred_cmd) + OuterKeyword.thy_goal (opt_modes -- scan_params -- P.term_group >> code_pred_cmd) end diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/Tools/Predicate_Compile/predicate_compile_aux.ML --- a/src/HOL/Tools/Predicate_Compile/predicate_compile_aux.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_aux.ML Wed Nov 04 11:40:59 2009 +0100 @@ -9,6 +9,29 @@ structure Predicate_Compile_Aux = struct + +(* mode *) + +type smode = (int * int list option) list +type mode = smode option list * smode +datatype tmode = Mode of mode * smode * tmode option list; + +fun string_of_smode js = + commas (map + (fn (i, is) => + string_of_int i ^ (case is of NONE => "" + | SOME is => "p" ^ enclose "[" "]" (commas (map string_of_int is)))) js) + +fun string_of_mode (iss, is) = space_implode " -> " (map + (fn NONE => "X" + | SOME js => enclose "[" "]" (string_of_smode js)) + (iss @ [SOME is])); + +fun string_of_tmode (Mode (predmode, termmode, param_modes)) = + "predmode: " ^ (string_of_mode predmode) ^ + (if null param_modes then "" else + "; " ^ "params: " ^ commas (map (the_default "NONE" o Option.map string_of_tmode) param_modes)) + (* general syntactic functions *) (*Like dest_conj, but flattens conjunctions however nested*) @@ -136,7 +159,7 @@ (* Different options for compiler *) datatype options = Options of { - expected_modes : (string * int list list) option, + expected_modes : (string * mode list) option, show_steps : bool, show_proof_trace : bool, show_intermediate_results : bool, @@ -146,7 +169,7 @@ skip_proof : bool, inductify : bool, - rpred : bool, + random : bool, depth_limited : bool }; @@ -160,7 +183,7 @@ fun skip_proof (Options opt) = #skip_proof opt fun is_inductify (Options opt) = #inductify opt -fun is_rpred (Options opt) = #rpred opt +fun is_random (Options opt) = #random opt fun is_depth_limited (Options opt) = #depth_limited opt val default_options = Options { @@ -174,7 +197,7 @@ skip_proof = false, inductify = false, - rpred = false, + random = false, depth_limited = false } diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/Tools/Predicate_Compile/predicate_compile_core.ML --- a/src/HOL/Tools/Predicate_Compile/predicate_compile_core.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_core.ML Wed Nov 04 11:40:59 2009 +0100 @@ -9,24 +9,20 @@ val setup: theory -> theory val code_pred: Predicate_Compile_Aux.options -> string -> Proof.context -> Proof.state val code_pred_cmd: Predicate_Compile_Aux.options -> string -> Proof.context -> Proof.state - type smode = (int * int list option) list - type mode = smode option list * smode - datatype tmode = Mode of mode * smode * tmode option list; val register_predicate : (string * thm list * thm * int) -> theory -> theory val register_intros : string * thm list -> theory -> theory val is_registered : theory -> string -> bool - val predfun_intro_of: theory -> string -> mode -> thm - val predfun_elim_of: theory -> string -> mode -> thm - val predfun_name_of: theory -> string -> mode -> string + val predfun_intro_of: theory -> string -> Predicate_Compile_Aux.mode -> thm + val predfun_elim_of: theory -> string -> Predicate_Compile_Aux.mode -> thm + val predfun_name_of: theory -> string -> Predicate_Compile_Aux.mode -> string val all_preds_of : theory -> string list - val modes_of: theory -> string -> mode list - val depth_limited_modes_of: theory -> string -> mode list - val depth_limited_function_name_of : theory -> string -> mode -> string - val generator_modes_of: theory -> string -> mode list - val generator_name_of : theory -> string -> mode -> string - val all_modes_of : theory -> (string * mode list) list - val all_generator_modes_of : theory -> (string * mode list) list - val string_of_mode : mode -> string + val modes_of: theory -> string -> Predicate_Compile_Aux.mode list + val depth_limited_modes_of: theory -> string -> Predicate_Compile_Aux.mode list + val depth_limited_function_name_of : theory -> string -> Predicate_Compile_Aux.mode -> string + val generator_modes_of: theory -> string -> Predicate_Compile_Aux.mode list + val generator_name_of : theory -> string -> Predicate_Compile_Aux.mode -> string + val all_modes_of : theory -> (string * Predicate_Compile_Aux.mode list) list + val all_generator_modes_of : theory -> (string * Predicate_Compile_Aux.mode list) list val intros_of: theory -> string -> thm list val nparams_of: theory -> string -> int val add_intro: thm -> theory -> theory @@ -67,8 +63,6 @@ (* debug stuff *) -fun tracing s = (if ! Toplevel.debug then Output.tracing s else ()); - fun print_tac s = Seq.single; fun print_tac' options s = @@ -140,9 +134,6 @@ type mode = arg_mode list type tmode = Mode of mode * *) -type smode = (int * int list option) list -type mode = smode option list * smode; -datatype tmode = Mode of mode * smode * tmode option list; fun gen_split_smode (mk_tuple, strip_tuple) smode ts = let @@ -165,32 +156,16 @@ (split_smode' smode (i+1) ts) in split_smode' smode 1 ts end -val split_smode = gen_split_smode (HOLogic.mk_tuple, HOLogic.strip_tuple) -val split_smodeT = gen_split_smode (HOLogic.mk_tupleT, HOLogic.strip_tupleT) +fun split_smode smode ts = gen_split_smode (HOLogic.mk_tuple, HOLogic.strip_tuple) smode ts +fun split_smodeT smode ts = gen_split_smode (HOLogic.mk_tupleT, HOLogic.strip_tupleT) smode ts fun gen_split_mode split_smode (iss, is) ts = let val (t1, t2) = chop (length iss) ts in (t1, split_smode is t2) end -val split_mode = gen_split_mode split_smode -val split_modeT = gen_split_mode split_smodeT - -fun string_of_smode js = - commas (map - (fn (i, is) => - string_of_int i ^ (case is of NONE => "" - | SOME is => "p" ^ enclose "[" "]" (commas (map string_of_int is)))) js) - -fun string_of_mode (iss, is) = space_implode " -> " (map - (fn NONE => "X" - | SOME js => enclose "[" "]" (string_of_smode js)) - (iss @ [SOME is])); - -fun string_of_tmode (Mode (predmode, termmode, param_modes)) = - "predmode: " ^ (string_of_mode predmode) ^ - (if null param_modes then "" else - "; " ^ "params: " ^ commas (map (the_default "NONE" o Option.map string_of_tmode) param_modes)) +fun split_mode (iss, is) ts = gen_split_mode split_smode (iss, is) ts +fun split_modeT (iss, is) ts = gen_split_mode split_smodeT (iss, is) ts datatype indprem = Prem of term list * term | Negprem of term list * term | Sidecond of term | Generator of (string * typ); @@ -279,7 +254,7 @@ val depth_limited_modes_of = (map fst) o #depth_limited_functions oo the_pred_data -val rpred_modes_of = (map fst) o #generators oo the_pred_data +val random_modes_of = (map fst) o #generators oo the_pred_data fun all_modes_of thy = map (fn name => (name, modes_of thy name)) (all_preds_of thy) @@ -327,13 +302,11 @@ val depth_limited_function_name_of = #name ooo the_depth_limited_function_data -(*val generator_modes_of = (map fst) o #generators oo the_pred_data*) - (* diagnostic display functions *) fun print_modes options modes = if show_modes options then - Output.tracing ("Inferred modes:\n" ^ + tracing ("Inferred modes:\n" ^ cat_lines (map (fn (s, ms) => s ^ ": " ^ commas (map string_of_mode ms)) modes)) else () @@ -344,7 +317,7 @@ ^ (string_of_entry pred mode entry) fun print_pred (pred, modes) = "predicate " ^ pred ^ ": " ^ cat_lines (map (print_mode pred) modes) - val _ = Output.tracing (cat_lines (map print_pred pred_mode_table)) + val _ = tracing (cat_lines (map print_pred pred_mode_table)) in () end; fun string_of_prem thy (Prem (ts, p)) = @@ -423,10 +396,10 @@ case expected_modes options of SOME (s, ms) => (case AList.lookup (op =) modes s of SOME modes => - if not (eq_set (op =) (map (map (rpair NONE)) ms, map snd modes)) then + if not (eq_set (op =) (ms, modes)) then error ("expected modes were not inferred:\n" - ^ "inferred modes for " ^ s ^ ": " - ^ commas (map ((enclose "[" "]") o string_of_smode o snd) modes)) + ^ "inferred modes for " ^ s ^ ": " ^ commas (map string_of_mode modes) + ^ "\n expected modes for " ^ s ^ ": " ^ commas (map string_of_mode ms)) else () | NONE => ()) | NONE => () @@ -444,7 +417,7 @@ val rec_consts = fold add_term_consts_2 cs' []; val intr_consts = fold add_term_consts_2 intr_ts' []; fun unify (cname, cT) = - let val consts = map snd (List.filter (fn c => fst c = cname) intr_consts) + let val consts = map snd (filter (fn c => fst c = cname) intr_consts) in fold (Sign.typ_unify thy) ((replicate (length consts) cT) ~~ consts) end; val (env, _) = fold unify rec_consts (Vartab.empty, i'); val subst = map_types (Envir.norm_type env) @@ -661,7 +634,7 @@ fun cons_intro gr = case try (Graph.get_node gr) name of SOME pred_data => Graph.map_node name (map_pred_data - (apfst (fn (intros, elim, nparams) => (thm::intros, elim, nparams)))) gr + (apfst (fn (intros, elim, nparams) => (intros @ [thm], elim, nparams)))) gr | NONE => let val nparams = the_default (guess_nparams T) (try (#nparams o rep_pred_data o (fetch_pred_data thy)) name) @@ -1052,16 +1025,16 @@ fun print_failed_mode options thy modes p m rs i = if show_mode_inference options then let - val _ = Output.tracing ("Clause " ^ string_of_int (i + 1) ^ " of " ^ + val _ = tracing ("Clause " ^ string_of_int (i + 1) ^ " of " ^ p ^ " violates mode " ^ string_of_mode m) - val _ = Output.tracing (string_of_clause thy p (nth rs i)) + val _ = tracing (string_of_clause thy p (nth rs i)) in () end else () fun check_modes_pred options with_generator thy param_vs clauses modes gen_modes (p, ms) = let val rs = case AList.lookup (op =) clauses p of SOME rs => rs | NONE => [] - in (p, List.filter (fn m => case find_index + in (p, filter (fn m => case find_index (is_none o check_mode_clause with_generator thy param_vs modes gen_modes m) rs of ~1 => true | i => (print_failed_mode options thy modes p m rs i; false)) ms) @@ -1191,6 +1164,28 @@ (t, names) end; +structure Comp_Mod = +struct + +datatype comp_modifiers = Comp_Modifiers of +{ + const_name_of : theory -> string -> Predicate_Compile_Aux.mode -> string, + funT_of : compilation_funs -> mode -> typ -> typ, + additional_arguments : string list -> term list, + wrap_compilation : compilation_funs -> string -> typ -> mode -> term list -> term -> term, + transform_additional_arguments : indprem -> term list -> term list +} + +fun dest_comp_modifiers (Comp_Modifiers c) = c + +val const_name_of = #const_name_of o dest_comp_modifiers +val funT_of = #funT_of o dest_comp_modifiers +val additional_arguments = #additional_arguments o dest_comp_modifiers +val wrap_compilation = #wrap_compilation o dest_comp_modifiers +val transform_additional_arguments = #transform_additional_arguments o dest_comp_modifiers + +end; + fun compile_arg compilation_modifiers compfuns additional_arguments thy param_vs iss arg = let fun map_params (t as Free (f, T)) = @@ -1198,7 +1193,7 @@ case (the (AList.lookup (op =) (param_vs ~~ iss) f)) of SOME is => let - val T' = #funT_of compilation_modifiers compfuns ([], is) T + val T' = Comp_Mod.funT_of compilation_modifiers compfuns ([], is) T in fst (mk_Eval_of additional_arguments ((Free (f, T'), T), SOME is) []) end | NONE => t else t @@ -1248,9 +1243,9 @@ val params' = map (compile_param compilation_modifiers compfuns thy) (ms ~~ params) val f' = case f of - Const (name, T) => Const (#const_name_of compilation_modifiers thy name mode, - #funT_of compilation_modifiers compfuns mode T) - | Free (name, T) => Free (name, #funT_of compilation_modifiers compfuns mode T) + Const (name, T) => Const (Comp_Mod.const_name_of compilation_modifiers thy name mode, + Comp_Mod.funT_of compilation_modifiers compfuns mode T) + | Free (name, T) => Free (name, Comp_Mod.funT_of compilation_modifiers compfuns mode T) | _ => error ("PredicateCompiler: illegal parameter term") in list_comb (f', params' @ args') @@ -1262,13 +1257,13 @@ let val params' = map (compile_param compilation_modifiers compfuns thy) (ms ~~ params) (*val mk_fun_of = if depth_limited then mk_depth_limited_fun_of else mk_fun_of*) - val name' = #const_name_of compilation_modifiers thy name mode - val T' = #funT_of compilation_modifiers compfuns mode T + val name' = Comp_Mod.const_name_of compilation_modifiers thy name mode + val T' = Comp_Mod.funT_of compilation_modifiers compfuns mode T in (list_comb (Const (name', T'), params' @ inargs @ additional_arguments)) end | (Free (name, T), params) => - list_comb (Free (name, #funT_of compilation_modifiers compfuns mode T), params @ inargs @ additional_arguments) + list_comb (Free (name, Comp_Mod.funT_of compilation_modifiers compfuns mode T), params @ inargs @ additional_arguments) fun compile_clause compilation_modifiers compfuns thy all_vs param_vs additional_arguments (iss, is) inp (ts, moded_ps) = let @@ -1302,7 +1297,7 @@ val (out_ts'', (names'', constr_vs')) = fold_map distinct_v out_ts' ((names', map (rpair []) vs)) val additional_arguments' = - #transform_additional_arguments compilation_modifiers p additional_arguments + Comp_Mod.transform_additional_arguments compilation_modifiers p additional_arguments val (compiled_clause, rest) = case p of Prem (us, t) => let @@ -1356,7 +1351,7 @@ val (Ts1, Ts2) = chop (length (fst mode)) (binder_types T) val (Us1, Us2) = split_smodeT (snd mode) Ts2 val Ts1' = - map2 (fn NONE => I | SOME is => #funT_of compilation_modifiers compfuns ([], is)) (fst mode) Ts1 + map2 (fn NONE => I | SOME is => Comp_Mod.funT_of compilation_modifiers compfuns ([], is)) (fst mode) Ts1 fun mk_input_term (i, NONE) = [Free (Name.variant (all_vs @ param_vs) ("x" ^ string_of_int i), nth Ts2 (i - 1))] | mk_input_term (i, SOME pis) = case HOLogic.strip_tupleT (nth Ts2 (i - 1)) of @@ -1370,17 +1365,17 @@ else [HOLogic.mk_tuple (map Free (vnames ~~ map (fn j => nth Ts (j - 1)) pis))] end val in_ts = maps mk_input_term (snd mode) val params = map2 (fn s => fn T => Free (s, T)) param_vs Ts1' - val additional_arguments = #additional_arguments compilation_modifiers (all_vs @ param_vs) + val additional_arguments = Comp_Mod.additional_arguments compilation_modifiers (all_vs @ param_vs) val cl_ts = map (compile_clause compilation_modifiers compfuns thy all_vs param_vs additional_arguments mode (HOLogic.mk_tuple in_ts)) moded_cls; - val compilation = #wrap_compilation compilation_modifiers compfuns s T mode additional_arguments + val compilation = Comp_Mod.wrap_compilation compilation_modifiers compfuns s T mode additional_arguments (if null cl_ts then mk_bot compfuns (HOLogic.mk_tupleT Us2) else foldr1 (mk_sup compfuns) cl_ts) val fun_const = - Const (#const_name_of compilation_modifiers thy s mode, - #funT_of compilation_modifiers compfuns mode T) + Const (Comp_Mod.const_name_of compilation_modifiers thy s mode, + Comp_Mod.funT_of compilation_modifiers compfuns mode T) in HOLogic.mk_Trueprop (HOLogic.mk_eq (list_comb (fun_const, params @ in_ts @ additional_arguments), compilation)) @@ -1589,7 +1584,7 @@ (mk_predT RandomPredCompFuns.compfuns (HOLogic.mk_tupleT outargTs)) end -fun rpred_create_definitions preds (name, modes) thy = +fun random_create_definitions preds (name, modes) thy = let val T = AList.lookup (op =) preds name |> the fun create_definition mode thy = @@ -2023,6 +2018,8 @@ map_preds_modes (fn pred => fn mode => fn t => Drule.standard (Skip_Proof.make_thm thy t)) compiled_terms +(* preparation of introduction rules into special datastructures *) + fun dest_prem thy params t = (case strip_comb t of (v as Free _, ts) => if v mem params then Prem (ts, v) else Sidecond t @@ -2096,6 +2093,8 @@ val all_modes = map (fn (s, T) => (s, modes_of_typ T)) preds in (preds, nparams, all_vs, param_vs, extra_modes, clauses, all_modes) end; +(* sanity check of introduction rules *) + fun check_format_of_intro_rule thy intro = let val concl = Logic.strip_imp_concl (prop_of intro) @@ -2137,40 +2136,93 @@ in forall check prednames end *) +(* create code equation *) + +fun add_code_equations thy nparams preds result_thmss = + let + fun add_code_equation ((predname, T), (pred, result_thms)) = + let + val full_mode = (replicate nparams NONE, + map (rpair NONE) (1 upto (length (binder_types T) - nparams))) + in + if member (op =) (modes_of thy predname) full_mode then + let + val Ts = binder_types T + val arg_names = Name.variant_list [] + (map (fn i => "x" ^ string_of_int i) (1 upto length Ts)) + val args = map Free (arg_names ~~ Ts) + val predfun = Const (predfun_name_of thy predname full_mode, + Ts ---> PredicateCompFuns.mk_predT @{typ unit}) + val rhs = PredicateCompFuns.mk_Eval (list_comb (predfun, args), @{term "Unity"}) + val eq_term = HOLogic.mk_Trueprop + (HOLogic.mk_eq (list_comb (Const (predname, T), args), rhs)) + val def = predfun_definition_of thy predname full_mode + val tac = fn {...} => Simplifier.simp_tac + (HOL_basic_ss addsimps [def, @{thm eval_pred}]) 1 + val eq = Goal.prove (ProofContext.init thy) arg_names [] eq_term tac + in + (pred, result_thms @ [eq]) + end + else + (pred, result_thms) + end + in + map add_code_equation (preds ~~ result_thmss) + end + (** main function of predicate compiler **) +datatype steps = Steps of + { + compile_preds : theory -> string list -> string list -> (string * typ) list + -> (moded_clause list) pred_mode_table -> term pred_mode_table, + create_definitions: (string * typ) list -> string * mode list -> theory -> theory, + infer_modes : options -> theory -> (string * mode list) list -> (string * mode list) list + -> string list -> (string * (term list * indprem list) list) list + -> moded_clause list pred_mode_table, + prove : options -> theory -> (string * (term list * indprem list) list) list + -> (string * typ) list -> (string * mode list) list + -> moded_clause list pred_mode_table -> term pred_mode_table -> thm pred_mode_table, + add_code_equations : theory -> int -> (string * typ) list + -> (string * thm list) list -> (string * thm list) list, + are_not_defined : theory -> string list -> bool, + qname : bstring + } + + fun add_equations_of steps options prednames thy = let + fun dest_steps (Steps s) = s val _ = print_step options ("Starting predicate compiler for predicates " ^ commas prednames ^ "...") - val _ = tracing (commas (map (Display.string_of_thm_global thy) (maps (intros_of thy) prednames))) (*val _ = check_intros_elim_match thy prednames*) (*val _ = map (check_format_of_intro_rule thy) (maps (intros_of thy) prednames)*) val (preds, nparams, all_vs, param_vs, extra_modes, clauses, all_modes) = prepare_intrs thy prednames (maps (intros_of thy) prednames) val _ = print_step options "Infering modes..." - val moded_clauses = #infer_modes steps options thy extra_modes all_modes param_vs clauses + val moded_clauses = #infer_modes (dest_steps steps) options thy extra_modes all_modes param_vs clauses val modes = map (fn (p, mps) => (p, map fst mps)) moded_clauses val _ = check_expected_modes options modes val _ = print_modes options modes (*val _ = print_moded_clauses thy moded_clauses*) val _ = print_step options "Defining executable functions..." - val thy' = fold (#create_definitions steps preds) modes thy + val thy' = fold (#create_definitions (dest_steps steps) preds) modes thy |> Theory.checkpoint val _ = print_step options "Compiling equations..." val compiled_terms = - (#compile_preds steps) thy' all_vs param_vs preds moded_clauses + #compile_preds (dest_steps steps) thy' all_vs param_vs preds moded_clauses val _ = print_compiled_terms options thy' compiled_terms val _ = print_step options "Proving equations..." - val result_thms = #prove steps options thy' clauses preds (extra_modes @ modes) + val result_thms = #prove (dest_steps steps) options thy' clauses preds (extra_modes @ modes) moded_clauses compiled_terms - val qname = #qname steps + val result_thms' = #add_code_equations (dest_steps steps) thy' nparams preds + (maps_modes result_thms) + val qname = #qname (dest_steps steps) val attrib = fn thy => Attrib.attribute_i thy (Attrib.internal (K (Thm.declaration_attribute (fn thm => Context.mapping (Code.add_eqn thm) I)))) val thy'' = fold (fn (name, result_thms) => fn thy => snd (PureThy.add_thmss [((Binding.qualify true (Long_Name.base_name name) (Binding.name qname), result_thms), [attrib thy ])] thy)) - (maps_modes result_thms) thy' - |> Theory.checkpoint + result_thms' thy' |> Theory.checkpoint in thy'' end @@ -2181,7 +2233,7 @@ SOME v => (G, v) | NONE => (Graph.new_node (key, value_of key) G, value_of key) val (G'', visited') = fold (extend' value_of edges_of) (subtract (op =) visited (edges_of (key, v))) - (G', key :: visited) + (G', key :: visited) in (fold (Graph.add_edge o (pair key)) (edges_of (key, v)) G'', visited') end; @@ -2190,6 +2242,7 @@ fun gen_add_equations steps options names thy = let + fun dest_steps (Steps s) = s val thy' = PredData.map (fold (extend (fetch_pred_data thy) (depending_preds_of thy)) names) thy |> Theory.checkpoint; fun strong_conn_of gr keys = @@ -2197,24 +2250,25 @@ val scc = strong_conn_of (PredData.get thy') names val thy'' = fold_rev (fn preds => fn thy => - if #are_not_defined steps thy preds then + if #are_not_defined (dest_steps steps) thy preds then add_equations_of steps options preds thy else thy) scc thy' |> Theory.checkpoint in thy'' end (* different instantiantions of the predicate compiler *) -val predicate_comp_modifiers = - {const_name_of = predfun_name_of, - funT_of = funT_of, +val predicate_comp_modifiers = Comp_Mod.Comp_Modifiers + {const_name_of = predfun_name_of : (theory -> string -> mode -> string), + funT_of = funT_of : (compilation_funs -> mode -> typ -> typ), additional_arguments = K [], - wrap_compilation = K (K (K (K (K I)))), - transform_additional_arguments = K I + wrap_compilation = K (K (K (K (K I)))) + : (compilation_funs -> string -> typ -> mode -> term list -> term -> term), + transform_additional_arguments = K I : (indprem -> term list -> term list) } -val depth_limited_comp_modifiers = +val depth_limited_comp_modifiers = Comp_Mod.Comp_Modifiers {const_name_of = depth_limited_function_name_of, - funT_of = depth_limited_funT_of, + funT_of = depth_limited_funT_of : (compilation_funs -> mode -> typ -> typ), additional_arguments = fn names => let val [depth_name, polarity_name] = Name.variant_list names ["depth", "polarity"] @@ -2245,38 +2299,41 @@ in [polarity', depth'] end } -val rpred_comp_modifiers = +val random_comp_modifiers = Comp_Mod.Comp_Modifiers {const_name_of = generator_name_of, - funT_of = K generator_funT_of, + funT_of = K generator_funT_of : (compilation_funs -> mode -> typ -> typ), additional_arguments = fn names => [Free (Name.variant names "size", @{typ code_numeral})], - wrap_compilation = K (K (K (K (K I)))), - transform_additional_arguments = K I + wrap_compilation = K (K (K (K (K I)))) + : (compilation_funs -> string -> typ -> mode -> term list -> term -> term), + transform_additional_arguments = K I : (indprem -> term list -> term list) } - val add_equations = gen_add_equations - {infer_modes = infer_modes, + (Steps {infer_modes = infer_modes, create_definitions = create_definitions, compile_preds = compile_preds predicate_comp_modifiers PredicateCompFuns.compfuns, prove = prove, + add_code_equations = add_code_equations, are_not_defined = fn thy => forall (null o modes_of thy), - qname = "equation"} + qname = "equation"}) val add_depth_limited_equations = gen_add_equations - {infer_modes = infer_modes, + (Steps {infer_modes = infer_modes, create_definitions = create_definitions_of_depth_limited_functions, compile_preds = compile_preds depth_limited_comp_modifiers PredicateCompFuns.compfuns, prove = prove_by_skip, + add_code_equations = K (K (K I)), are_not_defined = fn thy => forall (null o depth_limited_modes_of thy), - qname = "depth_limited_equation"} + qname = "depth_limited_equation"}) val add_quickcheck_equations = gen_add_equations - {infer_modes = infer_modes_with_generator, - create_definitions = rpred_create_definitions, - compile_preds = compile_preds rpred_comp_modifiers RandomPredCompFuns.compfuns, + (Steps {infer_modes = infer_modes_with_generator, + create_definitions = random_create_definitions, + compile_preds = compile_preds random_comp_modifiers RandomPredCompFuns.compfuns, prove = prove_by_skip, - are_not_defined = fn thy => forall (null o rpred_modes_of thy), - qname = "rpred_equation"} + add_code_equations = K (K (K I)), + are_not_defined = fn thy => forall (null o random_modes_of thy), + qname = "random_equation"}) (** user interface **) @@ -2307,7 +2364,7 @@ (extend (fetch_pred_data thy) (depending_preds_of thy) const)) lthy |> LocalTheory.checkpoint val thy' = ProofContext.theory_of lthy' - val preds = Graph.all_preds (PredData.get thy') [const] |> filter_out (has_elim thy') + val preds = Graph.all_succs (PredData.get thy') [const] |> filter_out (has_elim thy') fun mk_cases const = let val T = Sign.the_const_type thy const @@ -2317,7 +2374,7 @@ in mk_casesrule lthy' pred nparams intros end val cases_rules = map mk_cases preds val cases = - map (fn case_rule => RuleCases.Case {fixes = [], + map (fn case_rule => Rule_Cases.Case {fixes = [], assumes = [("", Logic.strip_imp_prems case_rule)], binds = [], cases = []}) cases_rules val case_env = map2 (fn p => fn c => (Long_Name.base_name p, SOME c)) preds cases @@ -2330,7 +2387,7 @@ (ProofContext.init (ProofContext.theory_of goal_ctxt)) (map the_single thms) in goal_ctxt |> LocalTheory.theory (fold set_elim global_thms #> - (if is_rpred options then + (if is_random options then (add_equations options [const] #> add_quickcheck_equations options [const]) else if is_depth_limited options then @@ -2380,7 +2437,7 @@ | SOME d => [@{term "True"}, HOLogic.mk_number @{typ "code_numeral"} d] val comp_modifiers = case depth_limit of NONE => - (if random then rpred_comp_modifiers else predicate_comp_modifiers) | SOME _ => depth_limited_comp_modifiers + (if random then random_comp_modifiers else predicate_comp_modifiers) | SOME _ => depth_limited_comp_modifiers val mk_fun_of = if random then mk_generator_of else if (is_some depth_limit) then mk_depth_limited_fun_of else mk_fun_of val t_pred = compile_expr comp_modifiers compfuns thy diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/Tools/Predicate_Compile/predicate_compile_data.ML --- a/src/HOL/Tools/Predicate_Compile/predicate_compile_data.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_data.ML Wed Nov 04 11:40:59 2009 +0100 @@ -7,7 +7,7 @@ signature PREDICATE_COMPILE_DATA = sig type specification_table; - val make_const_spec_table : theory -> specification_table + val make_const_spec_table : Predicate_Compile_Aux.options -> theory -> specification_table val get_specification : specification_table -> string -> thm list val obtain_specification_graph : Predicate_Compile_Aux.options -> theory -> specification_table -> string -> thm list Graph.T @@ -126,13 +126,22 @@ |> split_all_pairs thy |> tap check_equation_format -fun inline_equations thy th = Conv.fconv_rule (Simplifier.rewrite -((Simplifier.theory_context thy Simplifier.empty_ss) addsimps (Predicate_Compile_Inline_Defs.get (ProofContext.init thy)))) th +fun inline_equations options thy th = + let + val inline_defs = Predicate_Compile_Inline_Defs.get (ProofContext.init thy) + val th' = (Simplifier.full_simplify (HOL_basic_ss addsimps inline_defs)) th + (*val _ = print_step options + ("Inlining " ^ (Syntax.string_of_term_global thy (prop_of th)) + ^ "with " ^ (commas (map ((Syntax.string_of_term_global thy) o prop_of) inline_defs)) + ^" to " ^ (Syntax.string_of_term_global thy (prop_of th')))*) + in + th' + end -fun store_thm_in_table ignore_consts thy th= +fun store_thm_in_table options ignore_consts thy th= let val th = th - |> inline_equations thy + |> inline_equations options thy |> Predicate_Compile_Set.unfold_set_notation |> AxClass.unoverload thy val (const, th) = @@ -150,9 +159,9 @@ else I end -fun make_const_spec_table thy = +fun make_const_spec_table options thy = let - fun store ignore_const f = fold (store_thm_in_table ignore_const thy) (map (Thm.transfer thy) (f (ProofContext.init thy))) + fun store ignore_const f = fold (store_thm_in_table options ignore_const thy) (map (Thm.transfer thy) (f (ProofContext.init thy))) val table = Symtab.empty |> store [] Predicate_Compile_Alternative_Defs.get val ignore_consts = Symtab.keys table diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/Tools/Predicate_Compile/predicate_compile_pred.ML --- a/src/HOL/Tools/Predicate_Compile/predicate_compile_pred.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_pred.ML Wed Nov 04 11:40:59 2009 +0100 @@ -141,11 +141,8 @@ fun process constname atom (new_defs, thy) = let val (pred, args) = strip_comb atom - val abs_args = filter is_Abs args fun replace_abs_arg (abs_arg as Abs _ ) (new_defs, thy) = let - val _ = tracing ("Introduce new constant for " ^ - Syntax.string_of_term_global thy abs_arg) val vars = map Var (Term.add_vars abs_arg []) val abs_arg' = Logic.unvarify abs_arg val frees = map Free (Term.add_frees abs_arg' []) @@ -156,14 +153,18 @@ val const = Const (full_constname, constT) val lhs = list_comb (const, frees) val def = Logic.mk_equals (lhs, abs_arg') - val _ = tracing (Syntax.string_of_term_global thy def) val ([definition], thy') = thy |> Sign.add_consts_i [(Binding.name constname, constT, NoSyn)] |> PureThy.add_defs false [((Binding.name (constname ^ "_def"), def), [])] in (list_comb (Logic.varify const, vars), ((full_constname, [definition])::new_defs, thy')) end - | replace_abs_arg arg (new_defs, thy) = (arg, (new_defs, thy)) + | replace_abs_arg arg (new_defs, thy) = + if (is_predT (fastype_of arg)) then + process constname arg (new_defs, thy) + else + (arg, (new_defs, thy)) + val (args', (new_defs', thy')) = fold_map replace_abs_arg args (new_defs, thy) in (list_comb (pred, args'), (new_defs', thy')) diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/Tools/Predicate_Compile/predicate_compile_quickcheck.ML --- a/src/HOL/Tools/Predicate_Compile/predicate_compile_quickcheck.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_quickcheck.ML Wed Nov 04 11:40:59 2009 +0100 @@ -32,7 +32,7 @@ skip_proof = false, inductify = false, - rpred = false, + random = false, depth_limited = false } diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/Tools/TFL/post.ML --- a/src/HOL/Tools/TFL/post.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/HOL/Tools/TFL/post.ML Wed Nov 04 11:40:59 2009 +0100 @@ -28,7 +28,7 @@ *--------------------------------------------------------------------------*) fun termination_goals rules = map (Type.freeze o HOLogic.dest_Trueprop) - (List.foldr (fn (th,A) => uncurry (union (op aconv)) (prems_of th, A)) [] rules); + (fold_rev (union (op aconv) o prems_of) rules []); (*--------------------------------------------------------------------------- * Three postprocessors are applied to the definition. It @@ -71,7 +71,7 @@ | _ => r RS P_imp_P_eq_True (*Is this the best way to invoke the simplifier??*) -fun rewrite L = rewrite_rule (map mk_meta_eq (List.filter(not o id_thm) L)) +fun rewrite L = rewrite_rule (map mk_meta_eq (filter_out id_thm L)) fun join_assums th = let val thy = Thm.theory_of_thm th diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/Tools/TFL/rules.ML --- a/src/HOL/Tools/TFL/rules.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/HOL/Tools/TFL/rules.ML Wed Nov 04 11:40:59 2009 +0100 @@ -131,8 +131,7 @@ fun FILTER_DISCH_ALL P thm = let fun check tm = P (#t (Thm.rep_cterm tm)) - in List.foldr (fn (tm,th) => if check tm then DISCH tm th else th) - thm (chyps thm) + in fold_rev (fn tm => fn th => if check tm then DISCH tm th else th) (chyps thm) thm end; (* freezeT expensive! *) @@ -807,7 +806,7 @@ (prover names) (ss0 addsimps [cut_lemma'] addeqcongs congs) ctm val th2 = equal_elim th1 th in - (th2, List.filter (not o restricted) (!tc_list)) + (th2, filter_out restricted (!tc_list)) end; diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/Tools/TFL/tfl.ML --- a/src/HOL/Tools/TFL/tfl.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/HOL/Tools/TFL/tfl.ML Wed Nov 04 11:40:59 2009 +0100 @@ -423,13 +423,13 @@ end; -fun givens pats = map pat_of (List.filter given pats); +fun givens pats = map pat_of (filter given pats); fun post_definition meta_tflCongs (theory, (def, pats)) = let val tych = Thry.typecheck theory val f = #lhs(S.dest_eq(concl def)) val corollary = R.MATCH_MP Thms.WFREC_COROLLARY def - val pats' = List.filter given pats + val pats' = filter given pats val given_pats = map pat_of pats' val rows = map row_of_pat pats' val WFR = #ant(S.dest_imp(concl corollary)) @@ -529,7 +529,7 @@ then writeln (cat_lines ("Extractants =" :: map (Display.string_of_thm_global thy) extractants)) else () - val TCs = List.foldr (uncurry (union (op aconv))) [] TCl + val TCs = fold_rev (union (op aconv)) TCl [] val full_rqt = WFR::TCs val R' = S.mk_select{Bvar=R1, Body=S.list_mk_conj full_rqt} val R'abs = S.rand R' @@ -821,7 +821,7 @@ let val ex_tm = S.mk_exists{Bvar=v,Body=tm} in (ex_tm, R.CHOOSE(tych v, R.ASSUME (tych ex_tm)) thm) end - val [veq] = List.filter (can S.dest_eq) (#1 (R.dest_thm thm)) + val [veq] = filter (can S.dest_eq) (#1 (R.dest_thm thm)) val {lhs,rhs} = S.dest_eq veq val L = S.free_vars_lr rhs in #2 (fold_rev CHOOSER L (veq,thm)) end; diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/Tools/arith_data.ML --- a/src/HOL/Tools/arith_data.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/HOL/Tools/arith_data.ML Wed Nov 04 11:40:59 2009 +0100 @@ -1,7 +1,7 @@ (* Title: HOL/Tools/arith_data.ML Author: Markus Wenzel, Stefan Berghofer, and Tobias Nipkow -Common arithmetic proof auxiliary. +Common arithmetic proof auxiliary and legacy. *) signature ARITH_DATA = @@ -11,6 +11,11 @@ val add_tactic: string -> (bool -> Proof.context -> int -> tactic) -> theory -> theory val get_arith_facts: Proof.context -> thm list + val mk_number: typ -> int -> term + val mk_sum: typ -> term list -> term + val long_mk_sum: typ -> term list -> term + val dest_sum: term -> term list + val prove_conv_nohyps: tactic list -> Proof.context -> term * term -> thm option val prove_conv: tactic list -> Proof.context -> thm list -> term * term -> thm option val prove_conv2: tactic -> (simpset -> tactic) -> simpset -> term * term -> thm @@ -67,6 +72,36 @@ "various arithmetic decision procedures"; +(* some specialized syntactic operations *) + +fun mk_number T n = HOLogic.number_of_const T $ HOLogic.mk_numeral n; + +val mk_plus = HOLogic.mk_binop @{const_name HOL.plus}; + +fun mk_minus t = + let val T = Term.fastype_of t + in Const (@{const_name HOL.uminus}, T --> T) $ t end; + +(*Thus mk_sum[t] yields t+0; longer sums don't have a trailing zero*) +fun mk_sum T [] = mk_number T 0 + | mk_sum T [t,u] = mk_plus (t, u) + | mk_sum T (t :: ts) = mk_plus (t, mk_sum T ts); + +(*this version ALWAYS includes a trailing zero*) +fun long_mk_sum T [] = mk_number T 0 + | long_mk_sum T (t :: ts) = mk_plus (t, mk_sum T ts); + +(*decompose additions AND subtractions as a sum*) +fun dest_summing (pos, Const (@{const_name HOL.plus}, _) $ t $ u, ts) = + dest_summing (pos, t, dest_summing (pos, u, ts)) + | dest_summing (pos, Const (@{const_name HOL.minus}, _) $ t $ u, ts) = + dest_summing (pos, t, dest_summing (not pos, u, ts)) + | dest_summing (pos, t, ts) = + if pos then t::ts else mk_minus t :: ts; + +fun dest_sum t = dest_summing (true, t, []); + + (* various auxiliary and legacy *) fun prove_conv_nohyps tacs ctxt (t, u) = diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/Tools/choice_specification.ML --- a/src/HOL/Tools/choice_specification.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/HOL/Tools/choice_specification.ML Wed Nov 04 11:40:59 2009 +0100 @@ -120,7 +120,8 @@ val frees = OldTerm.term_frees prop val _ = forall (fn v => Sign.of_sort thy (type_of v,HOLogic.typeS)) frees orelse error "Specificaton: Only free variables of sort 'type' allowed" - val prop_closed = List.foldr (fn ((vname,T),prop) => HOLogic.mk_all (vname,T,prop)) prop (map dest_Free frees) + val prop_closed = fold_rev (fn (vname, T) => fn prop => + HOLogic.mk_all (vname, T, prop)) (map dest_Free frees) prop in (prop_closed,frees) end @@ -143,7 +144,7 @@ val (_, c') = Type.varify [] c val (cname,ctyp) = dest_Const c' in - case List.filter (fn t => let val (name,typ) = dest_Const t + case filter (fn t => let val (name,typ) = dest_Const t in name = cname andalso Sign.typ_equiv thy (typ, ctyp) end) thawed_prop_consts of [] => error ("Specification: No suitable instances of constant \"" ^ Syntax.string_of_term_global thy c ^ "\" found") @@ -151,7 +152,7 @@ | _ => error ("Specification: Several variations of \"" ^ Syntax.string_of_term_global thy c ^ "\" found (try applying explicit type constraints)") end val proc_consts = map proc_const consts - fun mk_exist (c,prop) = + fun mk_exist c prop = let val T = type_of c val cname = Long_Name.base_name (fst (dest_Const c)) @@ -161,7 +162,7 @@ in HOLogic.exists_const T $ Abs(vname,T,Term.abstract_over (c,prop)) end - val ex_prop = List.foldr mk_exist prop proc_consts + val ex_prop = fold_rev mk_exist proc_consts prop val cnames = map (fst o dest_Const) proc_consts fun post_process (arg as (thy,thm)) = let diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/Tools/inductive.ML --- a/src/HOL/Tools/inductive.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/HOL/Tools/inductive.ML Wed Nov 04 11:40:59 2009 +0100 @@ -392,7 +392,7 @@ list_all (params', Logic.list_implies (map (HOLogic.mk_Trueprop o HOLogic.mk_eq) (frees ~~ us) @ ts, P)); - val c_intrs = (List.filter (equal c o #1 o #1 o #1) intrs); + val c_intrs = filter (equal c o #1 o #1 o #1) intrs; val prems = HOLogic.mk_Trueprop (list_comb (c, params @ frees)) :: map mk_elim_prem (map #1 c_intrs) in @@ -517,16 +517,17 @@ | (Abs (a, T, t)) => (Abs (a, T, fst (subst t)), NONE) | _ => (s, NONE))); - fun mk_prem (s, prems) = (case subst s of - (_, SOME (t, u)) => t :: u :: prems - | (t, _) => t :: prems); + fun mk_prem s prems = + (case subst s of + (_, SOME (t, u)) => t :: u :: prems + | (t, _) => t :: prems); val SOME (_, i, ys, _) = dest_predicate cs params (HOLogic.dest_Trueprop (Logic.strip_assums_concl r)) in list_all_free (Logic.strip_params r, - Logic.list_implies (map HOLogic.mk_Trueprop (List.foldr mk_prem - [] (map HOLogic.dest_Trueprop (Logic.strip_assums_hyp r))), + Logic.list_implies (map HOLogic.mk_Trueprop (fold_rev mk_prem + (map HOLogic.dest_Trueprop (Logic.strip_assums_hyp r)) []), HOLogic.mk_Trueprop (list_comb (List.nth (preds, i), ys)))) end; @@ -549,9 +550,9 @@ (* make predicate for instantiation of abstract induction rule *) val ind_pred = fold_rev lambda (bs @ xs) (foldr1 HOLogic.mk_conj - (map_index (fn (i, P) => List.foldr HOLogic.mk_imp - (list_comb (P, make_args' argTs xs (binder_types (fastype_of P)))) - (make_bool_args HOLogic.mk_not I bs i)) preds)); + (map_index (fn (i, P) => fold_rev (curry HOLogic.mk_imp) + (make_bool_args HOLogic.mk_not I bs i) + (list_comb (P, make_args' argTs xs (binder_types (fastype_of P))))) preds)); val ind_concl = HOLogic.mk_Trueprop (HOLogic.mk_binrel "HOL.ord_class.less_eq" (rec_const, ind_pred)); @@ -631,9 +632,10 @@ map HOLogic.mk_eq (make_args' argTs xs Ts ~~ ts) @ map (subst o HOLogic.dest_Trueprop) (Logic.strip_assums_hyp r) - in List.foldr (fn ((x, T), P) => HOLogic.exists_const T $ (Abs (x, T, P))) - (if null ps then HOLogic.true_const else foldr1 HOLogic.mk_conj ps) - (Logic.strip_params r) + in + fold_rev (fn (x, T) => fn P => HOLogic.exists_const T $ Abs (x, T, P)) + (Logic.strip_params r) + (if null ps then HOLogic.true_const else foldr1 HOLogic.mk_conj ps) end (* make a disjunction of all introduction rules *) @@ -691,22 +693,22 @@ val rec_name = Binding.name_of rec_binding; fun rec_qualified qualified = Binding.qualify qualified rec_name; val intr_names = map Binding.name_of intr_bindings; - val ind_case_names = RuleCases.case_names intr_names; + val ind_case_names = Rule_Cases.case_names intr_names; val induct = if coind then - (raw_induct, [RuleCases.case_names [rec_name], - RuleCases.case_conclusion (rec_name, intr_names), - RuleCases.consumes 1, Induct.coinduct_pred (hd cnames)]) + (raw_induct, [Rule_Cases.case_names [rec_name], + Rule_Cases.case_conclusion (rec_name, intr_names), + Rule_Cases.consumes 1, Induct.coinduct_pred (hd cnames)]) else if no_ind orelse length cnames > 1 then - (raw_induct, [ind_case_names, RuleCases.consumes 0]) - else (raw_induct RSN (2, rev_mp), [ind_case_names, RuleCases.consumes 1]); + (raw_induct, [ind_case_names, Rule_Cases.consumes 0]) + else (raw_induct RSN (2, rev_mp), [ind_case_names, Rule_Cases.consumes 1]); val (intrs', ctxt1) = ctxt |> LocalTheory.notes kind (map (rec_qualified false) intr_bindings ~~ intr_atts ~~ map (fn th => [([th], - [Attrib.internal (K (ContextRules.intro_query NONE)), + [Attrib.internal (K (Context_Rules.intro_query NONE)), Attrib.internal (K Nitpick_Intros.add)])]) intrs) |>> map (hd o snd); val (((_, elims'), (_, [induct'])), ctxt2) = @@ -714,10 +716,10 @@ LocalTheory.note kind ((rec_qualified true (Binding.name "intros"), []), intrs') ||>> fold_map (fn (name, (elim, cases)) => LocalTheory.note kind ((Binding.qualify true (Long_Name.base_name name) (Binding.name "cases"), - [Attrib.internal (K (RuleCases.case_names cases)), - Attrib.internal (K (RuleCases.consumes 1)), + [Attrib.internal (K (Rule_Cases.case_names cases)), + Attrib.internal (K (Rule_Cases.consumes 1)), Attrib.internal (K (Induct.cases_pred name)), - Attrib.internal (K (ContextRules.elim_query NONE))]), [elim]) #> + Attrib.internal (K (Context_Rules.elim_query NONE))]), [elim]) #> apfst (hd o snd)) (if null elims then [] else cnames ~~ elims) ||>> LocalTheory.note kind ((rec_qualified true (Binding.name (coind_prefix coind ^ "induct")), @@ -730,7 +732,7 @@ LocalTheory.notes kind [((rec_qualified true (Binding.name "inducts"), []), inducts |> map (fn (name, th) => ([th], [Attrib.internal (K ind_case_names), - Attrib.internal (K (RuleCases.consumes 1)), + Attrib.internal (K (Rule_Cases.consumes 1)), Attrib.internal (K (Induct.induct_pred name))])))] |> snd end in (intrs', elims', induct', ctxt3) end; diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/Tools/inductive_codegen.ML --- a/src/HOL/Tools/inductive_codegen.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/HOL/Tools/inductive_codegen.ML Wed Nov 04 11:40:59 2009 +0100 @@ -71,8 +71,7 @@ {intros = intros |> Symtab.update (s, (AList.update Thm.eq_thm_prop (thm, (thyname_of s, nparms)) rules)), - graph = List.foldr (uncurry (Graph.add_edge o pair s)) - (fold add_node (s :: cs) graph) cs, + graph = fold_rev (Graph.add_edge o pair s) cs (fold add_node (s :: cs) graph), eqns = eqns} thy end | _ => (warn thm; thy)) @@ -174,7 +173,7 @@ let val is' = map (fn i => i - k) (List.drop (is, k)) in map (fn x => Mode (m, is', x)) (cprods (map (fn (NONE, _) => [NONE] - | (SOME js, arg) => map SOME (List.filter + | (SOME js, arg) => map SOME (filter (fn Mode (_, js', _) => js=js') (modes_of modes arg))) (iss ~~ args1))) end @@ -237,7 +236,7 @@ fun check_modes_pred thy arg_vs preds modes (p, ms) = let val SOME rs = AList.lookup (op =) preds p - in (p, List.filter (fn m => case find_index + in (p, filter (fn m => case find_index (not o check_mode_clause thy arg_vs modes m) rs of ~1 => true | i => (message ("Clause " ^ string_of_int (i+1) ^ " of " ^ diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/Tools/inductive_realizer.ML --- a/src/HOL/Tools/inductive_realizer.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/HOL/Tools/inductive_realizer.ML Wed Nov 04 11:40:59 2009 +0100 @@ -263,8 +263,7 @@ val rlz' = fold_rev Logic.all (vs2 @ rs) (prop_of rrule); val rlz'' = fold_rev Logic.all vs2 rlz in (name, (vs, - if rt = Extraction.nullt then rt else - List.foldr (uncurry lambda) rt vs1, + if rt = Extraction.nullt then rt else fold_rev lambda vs1 rt, ProofRewriteRules.un_hhf_proof rlz' rlz'' (fold_rev forall_intr_prf (vs2 @ rs) (prf_of rrule)))) end; diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/Tools/inductive_set.ML --- a/src/HOL/Tools/inductive_set.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/HOL/Tools/inductive_set.ML Wed Nov 04 11:40:59 2009 +0100 @@ -343,7 +343,7 @@ Simplifier.full_simplify (HOL_basic_ss addsimprocs [to_pred_simproc (mem_Collect_eq :: split_conv :: to_pred_simps)]) |> eta_contract_thm (is_pred pred_arities) |> - RuleCases.save thm + Rule_Cases.save thm end; val to_pred_att = Thm.rule_attribute o to_pred; @@ -374,7 +374,7 @@ Thm.instantiate ([], insts) |> Simplifier.full_simplify (HOL_basic_ss addsimps to_set_simps addsimprocs [strong_ind_simproc pred_arities, collect_mem_simproc]) |> - RuleCases.save thm + Rule_Cases.save thm end; val to_set_att = Thm.rule_attribute o to_set; @@ -522,7 +522,7 @@ Inductive.declare_rules kind rec_name coind no_ind cnames (map (to_set [] (Context.Proof ctxt3)) intrs) intr_names intr_atts (map (fn th => (to_set [] (Context.Proof ctxt3) th, - map fst (fst (RuleCases.get th)))) elims) + map fst (fst (Rule_Cases.get th)))) elims) raw_induct' ctxt3 in ({intrs = intrs', elims = elims', induct = induct, diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/Tools/lin_arith.ML --- a/src/HOL/Tools/lin_arith.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/HOL/Tools/lin_arith.ML Wed Nov 04 11:40:59 2009 +0100 @@ -645,9 +645,9 @@ fun filter_prems_tac_items (p : term -> bool) (terms : term list) : term list = let - fun filter_prems (t, (left, right)) = - if p t then (left, right @ [t]) else (left @ right, []) - val (left, right) = List.foldl filter_prems ([], []) terms + fun filter_prems t (left, right) = + if p t then (left, right @ [t]) else (left @ right, []) + val (left, right) = fold filter_prems terms ([], []) in right @ left end; @@ -681,7 +681,7 @@ val beta_eta_norm = map (apsnd (map (Envir.eta_contract o Envir.beta_norm))) split_goals (* TRY (etac notE) THEN eq_assume_tac: *) - val result = List.filter (not o negated_term_occurs_positively o snd) + val result = filter_out (negated_term_occurs_positively o snd) beta_eta_norm in result diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/Tools/meson.ML --- a/src/HOL/Tools/meson.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/HOL/Tools/meson.ML Wed Nov 04 11:40:59 2009 +0100 @@ -432,7 +432,7 @@ (*Generate Horn clauses for all contrapositives of a clause. The input, th, is a HOL disjunction.*) -fun add_contras crules (th,hcs) = +fun add_contras crules th hcs = let fun rots (0,th) = hcs | rots (k,th) = zero_var_indexes (make_horn crules th) :: rots(k-1, assoc_right (th RS disj_comm)) @@ -443,14 +443,14 @@ (*Use "theorem naming" to label the clauses*) fun name_thms label = - let fun name1 (th, (k,ths)) = + let fun name1 th (k, ths) = (k-1, Thm.put_name_hint (label ^ string_of_int k) th :: ths) - in fn ths => #2 (List.foldr name1 (length ths, []) ths) end; + in fn ths => #2 (fold_rev name1 ths (length ths, [])) end; (*Is the given disjunction an all-negative support clause?*) fun is_negative th = forall (not o #1) (literals (prop_of th)); -val neg_clauses = List.filter is_negative; +val neg_clauses = filter is_negative; (***** MESON PROOF PROCEDURE *****) @@ -491,9 +491,9 @@ TRYALL_eq_assume_tac; (*Sums the sizes of the subgoals, ignoring hypotheses (ancestors)*) -fun addconcl(prem,sz) = size_of_term(Logic.strip_assums_concl prem) + sz +fun addconcl prem sz = size_of_term (Logic.strip_assums_concl prem) + sz; -fun size_of_subgoals st = List.foldr addconcl 0 (prems_of st); +fun size_of_subgoals st = fold_rev addconcl (prems_of st) 0; (*Negation Normal Form*) @@ -553,19 +553,19 @@ (trace_msg (fn () => "Failed to Skolemize " ^ Display.string_of_thm ctxt th); skolemize_nnf_list ctxt ths); -fun add_clauses (th,cls) = +fun add_clauses th cls = let val ctxt0 = Variable.thm_context th - val (cnfs,ctxt) = make_cnf [] th ctxt0 + val (cnfs, ctxt) = make_cnf [] th ctxt0 in Variable.export ctxt ctxt0 cnfs @ cls end; (*Make clauses from a list of theorems, previously Skolemized and put into nnf. The resulting clauses are HOL disjunctions.*) -fun make_clauses ths = sort_clauses (List.foldr add_clauses [] ths); +fun make_clauses ths = sort_clauses (fold_rev add_clauses ths []); (*Convert a list of clauses (disjunctions) to Horn clauses (contrapositives)*) fun make_horns ths = name_thms "Horn#" - (distinct Thm.eq_thm_prop (List.foldr (add_contras clause_rules) [] ths)); + (distinct Thm.eq_thm_prop (fold_rev (add_contras clause_rules) ths [])); (*Could simply use nprems_of, which would count remaining subgoals -- no discrimination as to their size! With BEST_FIRST, fails for problem 41.*) diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/Tools/metis_tools.ML --- a/src/HOL/Tools/metis_tools.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/HOL/Tools/metis_tools.ML Wed Nov 04 11:40:59 2009 +0100 @@ -63,62 +63,62 @@ fun metis_lit b c args = (b, (c, args)); -fun hol_type_to_fol (ResClause.AtomV x) = Metis.Term.Var x - | hol_type_to_fol (ResClause.AtomF x) = Metis.Term.Fn(x,[]) - | hol_type_to_fol (ResClause.Comp(tc,tps)) = Metis.Term.Fn(tc, map hol_type_to_fol tps); +fun hol_type_to_fol (Res_Clause.AtomV x) = Metis.Term.Var x + | hol_type_to_fol (Res_Clause.AtomF x) = Metis.Term.Fn(x,[]) + | hol_type_to_fol (Res_Clause.Comp(tc,tps)) = Metis.Term.Fn(tc, map hol_type_to_fol tps); (*These two functions insert type literals before the real literals. That is the opposite order from TPTP linkup, but maybe OK.*) fun hol_term_to_fol_FO tm = - case ResHolClause.strip_comb tm of - (ResHolClause.CombConst(c,_,tys), tms) => + case Res_HOL_Clause.strip_comb tm of + (Res_HOL_Clause.CombConst(c,_,tys), tms) => let val tyargs = map hol_type_to_fol tys val args = map hol_term_to_fol_FO tms in Metis.Term.Fn (c, tyargs @ args) end - | (ResHolClause.CombVar(v,_), []) => Metis.Term.Var v + | (Res_HOL_Clause.CombVar(v,_), []) => Metis.Term.Var v | _ => error "hol_term_to_fol_FO"; -fun hol_term_to_fol_HO (ResHolClause.CombVar (a, _)) = Metis.Term.Var a - | hol_term_to_fol_HO (ResHolClause.CombConst (a, _, tylist)) = +fun hol_term_to_fol_HO (Res_HOL_Clause.CombVar (a, _)) = Metis.Term.Var a + | hol_term_to_fol_HO (Res_HOL_Clause.CombConst (a, _, tylist)) = Metis.Term.Fn (fn_isa_to_met a, map hol_type_to_fol tylist) - | hol_term_to_fol_HO (ResHolClause.CombApp (tm1, tm2)) = + | hol_term_to_fol_HO (Res_HOL_Clause.CombApp (tm1, tm2)) = Metis.Term.Fn (".", map hol_term_to_fol_HO [tm1, tm2]); (*The fully-typed translation, to avoid type errors*) fun wrap_type (tm, ty) = Metis.Term.Fn("ti", [tm, hol_type_to_fol ty]); -fun hol_term_to_fol_FT (ResHolClause.CombVar(a, ty)) = +fun hol_term_to_fol_FT (Res_HOL_Clause.CombVar(a, ty)) = wrap_type (Metis.Term.Var a, ty) - | hol_term_to_fol_FT (ResHolClause.CombConst(a, ty, _)) = + | hol_term_to_fol_FT (Res_HOL_Clause.CombConst(a, ty, _)) = wrap_type (Metis.Term.Fn(fn_isa_to_met a, []), ty) - | hol_term_to_fol_FT (tm as ResHolClause.CombApp(tm1,tm2)) = + | hol_term_to_fol_FT (tm as Res_HOL_Clause.CombApp(tm1,tm2)) = wrap_type (Metis.Term.Fn(".", map hol_term_to_fol_FT [tm1,tm2]), - ResHolClause.type_of_combterm tm); + Res_HOL_Clause.type_of_combterm tm); -fun hol_literal_to_fol FO (ResHolClause.Literal (pol, tm)) = - let val (ResHolClause.CombConst(p,_,tys), tms) = ResHolClause.strip_comb tm +fun hol_literal_to_fol FO (Res_HOL_Clause.Literal (pol, tm)) = + let val (Res_HOL_Clause.CombConst(p,_,tys), tms) = Res_HOL_Clause.strip_comb tm val tylits = if p = "equal" then [] else map hol_type_to_fol tys val lits = map hol_term_to_fol_FO tms in metis_lit pol (fn_isa_to_met p) (tylits @ lits) end - | hol_literal_to_fol HO (ResHolClause.Literal (pol, tm)) = - (case ResHolClause.strip_comb tm of - (ResHolClause.CombConst("equal",_,_), tms) => + | hol_literal_to_fol HO (Res_HOL_Clause.Literal (pol, tm)) = + (case Res_HOL_Clause.strip_comb tm of + (Res_HOL_Clause.CombConst("equal",_,_), tms) => metis_lit pol "=" (map hol_term_to_fol_HO tms) | _ => metis_lit pol "{}" [hol_term_to_fol_HO tm]) (*hBOOL*) - | hol_literal_to_fol FT (ResHolClause.Literal (pol, tm)) = - (case ResHolClause.strip_comb tm of - (ResHolClause.CombConst("equal",_,_), tms) => + | hol_literal_to_fol FT (Res_HOL_Clause.Literal (pol, tm)) = + (case Res_HOL_Clause.strip_comb tm of + (Res_HOL_Clause.CombConst("equal",_,_), tms) => metis_lit pol "=" (map hol_term_to_fol_FT tms) | _ => metis_lit pol "{}" [hol_term_to_fol_FT tm]) (*hBOOL*); fun literals_of_hol_thm thy mode t = - let val (lits, types_sorts) = ResHolClause.literals_of_term thy t + let val (lits, types_sorts) = Res_HOL_Clause.literals_of_term thy t in (map (hol_literal_to_fol mode) lits, types_sorts) end; (*Sign should be "true" for conjecture type constraints, "false" for type lits in clauses.*) -fun metis_of_typeLit pos (ResClause.LTVar (s,x)) = metis_lit pos s [Metis.Term.Var x] - | metis_of_typeLit pos (ResClause.LTFree (s,x)) = metis_lit pos s [Metis.Term.Fn(x,[])]; +fun metis_of_typeLit pos (Res_Clause.LTVar (s,x)) = metis_lit pos s [Metis.Term.Var x] + | metis_of_typeLit pos (Res_Clause.LTFree (s,x)) = metis_lit pos s [Metis.Term.Fn(x,[])]; fun default_sort _ (TVar _) = false | default_sort ctxt (TFree (x, s)) = (s = the_default [] (Variable.def_sort ctxt (x, ~1))); @@ -132,9 +132,9 @@ (literals_of_hol_thm thy mode o HOLogic.dest_Trueprop o prop_of) th in if is_conjecture then - (Metis.Thm.axiom (Metis.LiteralSet.fromList mlits), ResClause.add_typs types_sorts) + (Metis.Thm.axiom (Metis.LiteralSet.fromList mlits), Res_Clause.add_typs types_sorts) else - let val tylits = ResClause.add_typs + let val tylits = Res_Clause.add_typs (filter (not o default_sort ctxt) types_sorts) val mtylits = if Config.get ctxt type_lits then map (metis_of_typeLit false) tylits else [] @@ -145,13 +145,13 @@ (* ARITY CLAUSE *) -fun m_arity_cls (ResClause.TConsLit (c,t,args)) = - metis_lit true (ResClause.make_type_class c) [Metis.Term.Fn(t, map Metis.Term.Var args)] - | m_arity_cls (ResClause.TVarLit (c,str)) = - metis_lit false (ResClause.make_type_class c) [Metis.Term.Var str]; +fun m_arity_cls (Res_Clause.TConsLit (c,t,args)) = + metis_lit true (Res_Clause.make_type_class c) [Metis.Term.Fn(t, map Metis.Term.Var args)] + | m_arity_cls (Res_Clause.TVarLit (c,str)) = + metis_lit false (Res_Clause.make_type_class c) [Metis.Term.Var str]; (*TrueI is returned as the Isabelle counterpart because there isn't any.*) -fun arity_cls (ResClause.ArityClause{conclLit,premLits,...}) = +fun arity_cls (Res_Clause.ArityClause{conclLit,premLits,...}) = (TrueI, Metis.Thm.axiom (Metis.LiteralSet.fromList (map m_arity_cls (conclLit :: premLits)))); @@ -160,7 +160,7 @@ fun m_classrel_cls subclass superclass = [metis_lit false subclass [Metis.Term.Var "T"], metis_lit true superclass [Metis.Term.Var "T"]]; -fun classrel_cls (ResClause.ClassrelClause {subclass, superclass, ...}) = +fun classrel_cls (Res_Clause.ClassrelClause {subclass, superclass, ...}) = (TrueI, Metis.Thm.axiom (Metis.LiteralSet.fromList (m_classrel_cls subclass superclass))); (* ------------------------------------------------------------------------- *) @@ -209,14 +209,14 @@ | strip_happ args x = (x, args); fun fol_type_to_isa _ (Metis.Term.Var v) = - (case ResReconstruct.strip_prefix ResClause.tvar_prefix v of - SOME w => ResReconstruct.make_tvar w - | NONE => ResReconstruct.make_tvar v) + (case Res_Reconstruct.strip_prefix Res_Clause.tvar_prefix v of + SOME w => Res_Reconstruct.make_tvar w + | NONE => Res_Reconstruct.make_tvar v) | fol_type_to_isa ctxt (Metis.Term.Fn(x, tys)) = - (case ResReconstruct.strip_prefix ResClause.tconst_prefix x of - SOME tc => Term.Type (ResReconstruct.invert_type_const tc, map (fol_type_to_isa ctxt) tys) + (case Res_Reconstruct.strip_prefix Res_Clause.tconst_prefix x of + SOME tc => Term.Type (Res_Reconstruct.invert_type_const tc, map (fol_type_to_isa ctxt) tys) | NONE => - case ResReconstruct.strip_prefix ResClause.tfree_prefix x of + case Res_Reconstruct.strip_prefix Res_Clause.tfree_prefix x of SOME tf => mk_tfree ctxt tf | NONE => error ("fol_type_to_isa: " ^ x)); @@ -225,10 +225,10 @@ let val thy = ProofContext.theory_of ctxt val _ = trace_msg (fn () => "fol_term_to_hol: " ^ Metis.Term.toString fol_tm) fun tm_to_tt (Metis.Term.Var v) = - (case ResReconstruct.strip_prefix ResClause.tvar_prefix v of - SOME w => Type (ResReconstruct.make_tvar w) + (case Res_Reconstruct.strip_prefix Res_Clause.tvar_prefix v of + SOME w => Type (Res_Reconstruct.make_tvar w) | NONE => - case ResReconstruct.strip_prefix ResClause.schematic_var_prefix v of + case Res_Reconstruct.strip_prefix Res_Clause.schematic_var_prefix v of SOME w => Term (mk_var (w, HOLogic.typeT)) | NONE => Term (mk_var (v, HOLogic.typeT)) ) (*Var from Metis with a name like _nnn; possibly a type variable*) @@ -245,14 +245,14 @@ and applic_to_tt ("=",ts) = Term (list_comb(Const ("op =", HOLogic.typeT), terms_of (map tm_to_tt ts))) | applic_to_tt (a,ts) = - case ResReconstruct.strip_prefix ResClause.const_prefix a of + case Res_Reconstruct.strip_prefix Res_Clause.const_prefix a of SOME b => - let val c = ResReconstruct.invert_const b - val ntypes = ResReconstruct.num_typargs thy c + let val c = Res_Reconstruct.invert_const b + val ntypes = Res_Reconstruct.num_typargs thy c val nterms = length ts - ntypes val tts = map tm_to_tt ts val tys = types_of (List.take(tts,ntypes)) - val ntyargs = ResReconstruct.num_typargs thy c + val ntyargs = Res_Reconstruct.num_typargs thy c in if length tys = ntyargs then apply_list (Const (c, dummyT)) nterms (List.drop(tts,ntypes)) else error ("Constant " ^ c ^ " expects " ^ Int.toString ntyargs ^ @@ -263,14 +263,14 @@ cat_lines (map (Syntax.string_of_term ctxt) (terms_of tts))) end | NONE => (*Not a constant. Is it a type constructor?*) - case ResReconstruct.strip_prefix ResClause.tconst_prefix a of + case Res_Reconstruct.strip_prefix Res_Clause.tconst_prefix a of SOME b => - Type (Term.Type (ResReconstruct.invert_type_const b, types_of (map tm_to_tt ts))) + Type (Term.Type (Res_Reconstruct.invert_type_const b, types_of (map tm_to_tt ts))) | NONE => (*Maybe a TFree. Should then check that ts=[].*) - case ResReconstruct.strip_prefix ResClause.tfree_prefix a of + case Res_Reconstruct.strip_prefix Res_Clause.tfree_prefix a of SOME b => Type (mk_tfree ctxt b) | NONE => (*a fixed variable? They are Skolem functions.*) - case ResReconstruct.strip_prefix ResClause.fixed_var_prefix a of + case Res_Reconstruct.strip_prefix Res_Clause.fixed_var_prefix a of SOME b => let val opr = Term.Free(b, HOLogic.typeT) in apply_list opr (length ts) (map tm_to_tt ts) end @@ -281,16 +281,16 @@ fun fol_term_to_hol_FT ctxt fol_tm = let val _ = trace_msg (fn () => "fol_term_to_hol_FT: " ^ Metis.Term.toString fol_tm) fun cvt (Metis.Term.Fn ("ti", [Metis.Term.Var v, _])) = - (case ResReconstruct.strip_prefix ResClause.schematic_var_prefix v of + (case Res_Reconstruct.strip_prefix Res_Clause.schematic_var_prefix v of SOME w => mk_var(w, dummyT) | NONE => mk_var(v, dummyT)) | cvt (Metis.Term.Fn ("ti", [Metis.Term.Fn ("=",[]), _])) = Const ("op =", HOLogic.typeT) | cvt (Metis.Term.Fn ("ti", [Metis.Term.Fn (x,[]), ty])) = - (case ResReconstruct.strip_prefix ResClause.const_prefix x of - SOME c => Const (ResReconstruct.invert_const c, dummyT) + (case Res_Reconstruct.strip_prefix Res_Clause.const_prefix x of + SOME c => Const (Res_Reconstruct.invert_const c, dummyT) | NONE => (*Not a constant. Is it a fixed variable??*) - case ResReconstruct.strip_prefix ResClause.fixed_var_prefix x of + case Res_Reconstruct.strip_prefix Res_Clause.fixed_var_prefix x of SOME v => Free (v, fol_type_to_isa ctxt ty) | NONE => error ("fol_term_to_hol_FT bad constant: " ^ x)) | cvt (Metis.Term.Fn ("ti", [Metis.Term.Fn (".",[tm1,tm2]), _])) = @@ -301,10 +301,10 @@ | cvt (Metis.Term.Fn ("=", [tm1,tm2])) = list_comb(Const ("op =", HOLogic.typeT), map cvt [tm1,tm2]) | cvt (t as Metis.Term.Fn (x, [])) = - (case ResReconstruct.strip_prefix ResClause.const_prefix x of - SOME c => Const (ResReconstruct.invert_const c, dummyT) + (case Res_Reconstruct.strip_prefix Res_Clause.const_prefix x of + SOME c => Const (Res_Reconstruct.invert_const c, dummyT) | NONE => (*Not a constant. Is it a fixed variable??*) - case ResReconstruct.strip_prefix ResClause.fixed_var_prefix x of + case Res_Reconstruct.strip_prefix Res_Clause.fixed_var_prefix x of SOME v => Free (v, dummyT) | NONE => (trace_msg (fn () => "fol_term_to_hol_FT bad const: " ^ x); fol_term_to_hol_RAW ctxt t)) @@ -383,9 +383,9 @@ " in " ^ Display.string_of_thm ctxt i_th); NONE) fun remove_typeinst (a, t) = - case ResReconstruct.strip_prefix ResClause.schematic_var_prefix a of + case Res_Reconstruct.strip_prefix Res_Clause.schematic_var_prefix a of SOME b => SOME (b, t) - | NONE => case ResReconstruct.strip_prefix ResClause.tvar_prefix a of + | NONE => case Res_Reconstruct.strip_prefix Res_Clause.tvar_prefix a of SOME _ => NONE (*type instantiations are forbidden!*) | NONE => SOME (a,t) (*internal Metis var?*) val _ = trace_msg (fn () => " isa th: " ^ Display.string_of_thm ctxt i_th) @@ -452,7 +452,7 @@ in cterm_instantiate [(refl_x, c_t)] REFL_THM end; fun get_ty_arg_size _ (Const ("op =", _)) = 0 (*equality has no type arguments*) - | get_ty_arg_size thy (Const (c, _)) = (ResReconstruct.num_typargs thy c handle TYPE _ => 0) + | get_ty_arg_size thy (Const (c, _)) = (Res_Reconstruct.num_typargs thy c handle TYPE _ => 0) | get_ty_arg_size _ _ = 0; (* INFERENCE RULE: EQUALITY *) @@ -538,7 +538,7 @@ | step ctxt mode _ (_, Metis.Proof.Equality (f_lit, f_p, f_r)) = equality_inf ctxt mode f_lit f_p f_r; -fun real_literal (_, (c, _)) = not (String.isPrefix ResClause.class_prefix c); +fun real_literal (_, (c, _)) = not (String.isPrefix Res_Clause.class_prefix c); fun translate _ _ thpairs [] = thpairs | translate mode ctxt thpairs ((fol_th, inf) :: infpairs) = @@ -564,23 +564,23 @@ (* Translation of HO Clauses *) (* ------------------------------------------------------------------------- *) -fun cnf_th thy th = hd (ResAxioms.cnf_axiom thy th); +fun cnf_th thy th = hd (Res_Axioms.cnf_axiom thy th); val equal_imp_fequal' = cnf_th @{theory} @{thm equal_imp_fequal}; val fequal_imp_equal' = cnf_th @{theory} @{thm fequal_imp_equal}; -val comb_I = cnf_th @{theory} ResHolClause.comb_I; -val comb_K = cnf_th @{theory} ResHolClause.comb_K; -val comb_B = cnf_th @{theory} ResHolClause.comb_B; -val comb_C = cnf_th @{theory} ResHolClause.comb_C; -val comb_S = cnf_th @{theory} ResHolClause.comb_S; +val comb_I = cnf_th @{theory} Res_HOL_Clause.comb_I; +val comb_K = cnf_th @{theory} Res_HOL_Clause.comb_K; +val comb_B = cnf_th @{theory} Res_HOL_Clause.comb_B; +val comb_C = cnf_th @{theory} Res_HOL_Clause.comb_C; +val comb_S = cnf_th @{theory} Res_HOL_Clause.comb_S; fun type_ext thy tms = - let val subs = ResAtp.tfree_classes_of_terms tms - val supers = ResAtp.tvar_classes_of_terms tms - and tycons = ResAtp.type_consts_of_terms thy tms - val (supers', arity_clauses) = ResClause.make_arity_clauses thy tycons supers - val classrel_clauses = ResClause.make_classrel_clauses thy subs supers' + let val subs = Res_ATP.tfree_classes_of_terms tms + val supers = Res_ATP.tvar_classes_of_terms tms + and tycons = Res_ATP.type_consts_of_terms thy tms + val (supers', arity_clauses) = Res_Clause.make_arity_clauses thy tycons supers + val classrel_clauses = Res_Clause.make_classrel_clauses thy subs supers' in map classrel_cls classrel_clauses @ map arity_cls arity_clauses end; @@ -590,7 +590,7 @@ type logic_map = {axioms : (Metis.Thm.thm * thm) list, - tfrees : ResClause.type_literal list}; + tfrees : Res_Clause.type_literal list}; fun const_in_metis c (pred, tm_list) = let @@ -602,7 +602,7 @@ (*Extract TFree constraints from context to include as conjecture clauses*) fun init_tfrees ctxt = let fun add ((a,i),s) Ts = if i = ~1 then TFree(a,s) :: Ts else Ts - in ResClause.add_typs (Vartab.fold add (#2 (Variable.constraints_of ctxt)) []) end; + in Res_Clause.add_typs (Vartab.fold add (#2 (Variable.constraints_of ctxt)) []) end; (*transform isabelle type / arity clause to metis clause *) fun add_type_thm [] lmap = lmap @@ -628,15 +628,14 @@ | set_mode FT = FT val mode = set_mode mode0 (*transform isabelle clause to metis clause *) - fun add_thm is_conjecture (ith, {axioms, tfrees}) : logic_map = + fun add_thm is_conjecture ith {axioms, tfrees} : logic_map = let val (mth, tfree_lits) = hol_thm_to_fol is_conjecture ctxt mode ith in {axioms = (mth, Meson.make_meta_clause ith) :: axioms, tfrees = union (op =) tfree_lits tfrees} end; - val lmap0 = List.foldl (add_thm true) - {axioms = [], tfrees = init_tfrees ctxt} cls - val lmap = List.foldl (add_thm false) (add_tfrees lmap0) ths + val lmap0 = fold (add_thm true) cls {axioms = [], tfrees = init_tfrees ctxt} + val lmap = fold (add_thm false) ths (add_tfrees lmap0) val clause_lists = map (Metis.Thm.clause o #1) (#axioms lmap) fun used c = exists (Metis.LiteralSet.exists (const_in_metis c o #2)) clause_lists (*Now check for the existence of certain combinators*) @@ -647,7 +646,7 @@ val thS = if used "c_COMBS" then [comb_S] else [] val thEQ = if used "c_fequal" then [fequal_imp_equal', equal_imp_fequal'] else [] val lmap' = if mode=FO then lmap - else List.foldl (add_thm false) lmap (thEQ @ thS @ thC @ thB @ thK @ thI) + else fold (add_thm false) (thEQ @ thS @ thC @ thB @ thK @ thI) lmap in (mode, add_type_thm (type_ext thy (map prop_of (cls @ ths))) lmap') end; @@ -664,7 +663,7 @@ (* Main function to start metis prove and reconstruction *) fun FOL_SOLVE mode ctxt cls ths0 = let val thy = ProofContext.theory_of ctxt - val th_cls_pairs = map (fn th => (Thm.get_name_hint th, ResAxioms.cnf_axiom thy th)) ths0 + val th_cls_pairs = map (fn th => (Thm.get_name_hint th, Res_Axioms.cnf_axiom thy th)) ths0 val ths = maps #2 th_cls_pairs val _ = trace_msg (fn () => "FOL_SOLVE: CONJECTURE CLAUSES") val _ = app (fn th => trace_msg (fn () => Display.string_of_thm ctxt th)) cls @@ -673,14 +672,14 @@ val (mode, {axioms,tfrees}) = build_map mode ctxt cls ths val _ = if null tfrees then () else (trace_msg (fn () => "TFREE CLAUSES"); - app (fn tf => trace_msg (fn _ => ResClause.tptp_of_typeLit true tf)) tfrees) + app (fn tf => trace_msg (fn _ => Res_Clause.tptp_of_typeLit true tf)) tfrees) val _ = trace_msg (fn () => "CLAUSES GIVEN TO METIS") val thms = map #1 axioms val _ = app (fn th => trace_msg (fn () => Metis.Thm.toString th)) thms val _ = trace_msg (fn () => "mode = " ^ string_of_mode mode) val _ = trace_msg (fn () => "START METIS PROVE PROCESS") in - case List.filter (is_false o prop_of) cls of + case filter (is_false o prop_of) cls of false_th::_ => [false_th RS @{thm FalseE}] | [] => case refute thms of @@ -694,10 +693,11 @@ and used = map_filter (used_axioms axioms) proof val _ = trace_msg (fn () => "METIS COMPLETED...clauses actually used:") val _ = app (fn th => trace_msg (fn () => Display.string_of_thm ctxt th)) used - val unused = filter (fn (_, cls) => not (common_thm used cls)) th_cls_pairs + val unused = th_cls_pairs |> map_filter (fn (name, cls) => + if common_thm used cls then NONE else SOME name) in if null unused then () - else warning ("Metis: unused theorems " ^ commas_quote (map #1 unused)); + else warning ("Metis: unused theorems " ^ commas_quote unused); case result of (_,ith)::_ => (trace_msg (fn () => "success: " ^ Display.string_of_thm ctxt ith); @@ -714,12 +714,12 @@ let val _ = trace_msg (fn () => "Metis called with theorems " ^ cat_lines (map (Display.string_of_thm ctxt) ths)) in - if exists_type ResAxioms.type_has_empty_sort (prop_of st0) + if exists_type Res_Axioms.type_has_empty_sort (prop_of st0) then (warning "Proof state contains the empty sort"; Seq.empty) else - (Meson.MESON ResAxioms.neg_clausify + (Meson.MESON Res_Axioms.neg_clausify (fn cls => resolve_tac (FOL_SOLVE mode ctxt cls ths) 1) ctxt i - THEN ResAxioms.expand_defs_tac st0) st0 + THEN Res_Axioms.expand_defs_tac st0) st0 end handle METIS s => (warning ("Metis: " ^ s); Seq.empty); @@ -736,7 +736,7 @@ method @{binding metisF} FO "METIS for FOL problems" #> method @{binding metisFT} FT "METIS With-fully typed translation" #> Method.setup @{binding finish_clausify} - (Scan.succeed (K (SIMPLE_METHOD (ResAxioms.expand_defs_tac refl)))) + (Scan.succeed (K (SIMPLE_METHOD (Res_Axioms.expand_defs_tac refl)))) "cleanup after conversion to clauses"; end; diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/Tools/numeral_simprocs.ML --- a/src/HOL/Tools/numeral_simprocs.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/HOL/Tools/numeral_simprocs.ML Wed Nov 04 11:40:59 2009 +0100 @@ -16,9 +16,6 @@ signature NUMERAL_SIMPROCS = sig - val mk_sum: typ -> term list -> term - val dest_sum: term -> term list - val assoc_fold_simproc: simproc val combine_numerals: simproc val cancel_numerals: simproc list @@ -32,39 +29,10 @@ structure Numeral_Simprocs : NUMERAL_SIMPROCS = struct -fun mk_number T n = HOLogic.number_of_const T $ HOLogic.mk_numeral n; - -fun find_first_numeral past (t::terms) = - ((snd (HOLogic.dest_number t), rev past @ terms) - handle TERM _ => find_first_numeral (t::past) terms) - | find_first_numeral past [] = raise TERM("find_first_numeral", []); - -val mk_plus = HOLogic.mk_binop @{const_name HOL.plus}; - -fun mk_minus t = - let val T = Term.fastype_of t - in Const (@{const_name HOL.uminus}, T --> T) $ t end; - -(*Thus mk_sum[t] yields t+0; longer sums don't have a trailing zero*) -fun mk_sum T [] = mk_number T 0 - | mk_sum T [t,u] = mk_plus (t, u) - | mk_sum T (t :: ts) = mk_plus (t, mk_sum T ts); - -(*this version ALWAYS includes a trailing zero*) -fun long_mk_sum T [] = mk_number T 0 - | long_mk_sum T (t :: ts) = mk_plus (t, mk_sum T ts); - -val dest_plus = HOLogic.dest_bin @{const_name HOL.plus} Term.dummyT; - -(*decompose additions AND subtractions as a sum*) -fun dest_summing (pos, Const (@{const_name HOL.plus}, _) $ t $ u, ts) = - dest_summing (pos, t, dest_summing (pos, u, ts)) - | dest_summing (pos, Const (@{const_name HOL.minus}, _) $ t $ u, ts) = - dest_summing (pos, t, dest_summing (not pos, u, ts)) - | dest_summing (pos, t, ts) = - if pos then t::ts else mk_minus t :: ts; - -fun dest_sum t = dest_summing (true, t, []); +val mk_number = Arith_Data.mk_number; +val mk_sum = Arith_Data.mk_sum; +val long_mk_sum = Arith_Data.long_mk_sum; +val dest_sum = Arith_Data.dest_sum; val mk_diff = HOLogic.mk_binop @{const_name HOL.minus}; val dest_diff = HOLogic.dest_bin @{const_name HOL.minus} Term.dummyT; @@ -95,6 +63,11 @@ in dest_prod t @ dest_prod u end handle TERM _ => [t]; +fun find_first_numeral past (t::terms) = + ((snd (HOLogic.dest_number t), rev past @ terms) + handle TERM _ => find_first_numeral (t::past) terms) + | find_first_numeral past [] = raise TERM("find_first_numeral", []); + (*DON'T do the obvious simplifications; that would create special cases*) fun mk_coeff (k, t) = mk_times (mk_number (Term.fastype_of t) k, t); diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/Tools/old_primrec.ML --- a/src/HOL/Tools/old_primrec.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/HOL/Tools/old_primrec.ML Wed Nov 04 11:40:59 2009 +0100 @@ -34,11 +34,11 @@ same type in all introduction rules*) fun unify_consts thy cs intr_ts = (let - fun varify (t, (i, ts)) = + fun varify t (i, ts) = let val t' = map_types (Logic.incr_tvar (i + 1)) (snd (Type.varify [] t)) in (maxidx_of_term t', t'::ts) end; - val (i, cs') = List.foldr varify (~1, []) cs; - val (i', intr_ts') = List.foldr varify (i, []) intr_ts; + val (i, cs') = fold_rev varify cs (~1, []); + val (i', intr_ts') = fold_rev varify intr_ts (i, []); val rec_consts = fold Term.add_consts cs' []; val intr_consts = fold Term.add_consts intr_ts' []; fun unify (cname, cT) = @@ -195,7 +195,7 @@ NONE => let val dummy_fns = map (fn (_, cargs) => Const ("HOL.undefined", - replicate ((length cargs) + (length (List.filter is_rec_type cargs))) + replicate (length cargs + length (filter is_rec_type cargs)) dummyT ---> HOLogic.unitT)) constrs; val _ = warning ("No function definition for datatype " ^ quote tname) in @@ -237,8 +237,8 @@ val params_of = these o AList.lookup (op =) (maps constrs_of rec_eqns); in induct - |> RuleCases.rename_params (map params_of (maps (map #1 o #3 o #2) descr)) - |> RuleCases.save induct + |> Rule_Cases.rename_params (map params_of (maps (map #1 o #3 o #2) descr)) + |> Rule_Cases.save induct end; local diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/Tools/primrec.ML --- a/src/HOL/Tools/primrec.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/HOL/Tools/primrec.ML Wed Nov 04 11:40:59 2009 +0100 @@ -178,7 +178,7 @@ NONE => let val dummy_fns = map (fn (_, cargs) => Const ("HOL.undefined", - replicate ((length cargs) + (length (List.filter is_rec_type cargs))) + replicate (length cargs + length (filter is_rec_type cargs)) dummyT ---> HOLogic.unitT)) constrs; val _ = warning ("No function definition for datatype " ^ quote tname) in diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/Tools/record.ML --- a/src/HOL/Tools/record.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/HOL/Tools/record.ML Wed Nov 04 11:40:59 2009 +0100 @@ -262,7 +262,7 @@ infixr 0 ==>; val op $$ = Term.list_comb; -val op :== = PrimitiveDefs.mk_defpair; +val op :== = Primitive_Defs.mk_defpair; val op === = Trueprop o HOLogic.mk_eq; val op ==> = Logic.mk_implies; @@ -1553,7 +1553,7 @@ (* attributes *) -fun case_names_fields x = RuleCases.case_names ["fields"] x; +fun case_names_fields x = Rule_Cases.case_names ["fields"] x; fun induct_type_global name = [case_names_fields, Induct.induct_type name]; fun cases_type_global name = [case_names_fields, Induct.cases_type name]; diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/Tools/refute.ML --- a/src/HOL/Tools/refute.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/HOL/Tools/refute.ML Wed Nov 04 11:40:59 2009 +0100 @@ -394,7 +394,7 @@ (* (string * int) list *) val sizes = map_filter (fn (name, value) => Option.map (pair name) (Int.fromString value)) - (List.filter (fn (name, _) => name<>"minsize" andalso name<>"maxsize" + (filter (fn (name, _) => name<>"minsize" andalso name<>"maxsize" andalso name<>"maxvars" andalso name<>"maxtime" andalso name<>"satsolver") allparams) in @@ -954,7 +954,7 @@ (* required for mutually recursive datatypes; those need to *) (* be added even if they are an instance of an otherwise non- *) (* recursive datatype *) - fun collect_dtyp (d, acc) = + fun collect_dtyp d acc = let val dT = typ_of_dtyp descr typ_assoc d in @@ -962,7 +962,7 @@ DatatypeAux.DtTFree _ => collect_types dT acc | DatatypeAux.DtType (_, ds) => - collect_types dT (List.foldr collect_dtyp acc ds) + collect_types dT (fold_rev collect_dtyp ds acc) | DatatypeAux.DtRec i => if dT mem acc then acc (* prevent infinite recursion *) @@ -976,9 +976,9 @@ insert (op =) dT acc else acc (* collect argument types *) - val acc_dtyps = List.foldr collect_dtyp acc_dT dtyps + val acc_dtyps = fold_rev collect_dtyp dtyps acc_dT (* collect constructor types *) - val acc_dconstrs = List.foldr collect_dtyp acc_dtyps (maps snd dconstrs) + val acc_dconstrs = fold_rev collect_dtyp (maps snd dconstrs) acc_dtyps in acc_dconstrs end @@ -986,7 +986,7 @@ in (* argument types 'Ts' could be added here, but they are also *) (* added by 'collect_dtyp' automatically *) - collect_dtyp (DatatypeAux.DtRec index, acc) + collect_dtyp (DatatypeAux.DtRec index) acc end | NONE => (* not an inductive datatype, e.g. defined via "typedef" or *) @@ -1070,8 +1070,7 @@ (* continue search *) next max (len+1) (sum+x1) (x2::xs) (* only consider those types for which the size is not fixed *) - val mutables = List.filter - (not o (AList.defined (op =) sizes) o string_of_typ o fst) xs + val mutables = filter_out (AList.defined (op =) sizes o string_of_typ o fst) xs (* subtract 'minsize' from every size (will be added again at the end) *) val diffs = map (fn (_, n) => n-minsize) mutables in @@ -1597,8 +1596,9 @@ val Ts = Term.binder_types (Term.fastype_of t) val t' = Term.incr_boundvars i t in - List.foldr (fn (T, term) => Abs ("", T, term)) - (Term.list_comb (t', map Bound (i-1 downto 0))) (List.take (Ts, i)) + fold_rev (fn T => fn term => Abs ("", T, term)) + (List.take (Ts, i)) + (Term.list_comb (t', map Bound (i-1 downto 0))) end; (* ------------------------------------------------------------------------- *) @@ -2059,7 +2059,7 @@ Const (@{const_name insert}, HOLogic_prodT --> HOLogic_setT --> HOLogic_setT) in (* functions as graphs, i.e. as a (HOL) set of pairs "(x, y)" *) - map (List.foldr (fn (pair, acc) => HOLogic_insert $ pair $ acc) + map (fn ps => fold_rev (fn pair => fn acc => HOLogic_insert $ pair $ acc) ps HOLogic_empty_set) pairss end | Type (s, Ts) => @@ -2552,7 +2552,7 @@ (* arguments *) val (_, _, constrs) = the (AList.lookup (op =) descr idx) val (_, dtyps) = List.nth (constrs, c) - val rec_dtyps_args = List.filter + val rec_dtyps_args = filter (DatatypeAux.is_rec_type o fst) (dtyps ~~ args) (* map those indices to interpretations *) val rec_dtyps_intrs = map (fn (dtyp, arg) => @@ -2591,8 +2591,8 @@ (* interpretation list *) val arg_intrs = map (uncurry rec_intr) rec_dtyps_intrs (* apply 'intr' to all recursive arguments *) - val result = List.foldl (fn (arg_i, i) => - interpretation_apply (i, arg_i)) intr arg_intrs + val result = fold (fn arg_i => fn i => + interpretation_apply (i, arg_i)) arg_intrs intr (* update 'REC_OPERATORS' *) val _ = Array.update (arr, elem, (true, result)) in @@ -2971,11 +2971,11 @@ "intersection: interpretation for set is not a node") (* interpretation -> interpretaion *) fun lfp (Node resultsets) = - List.foldl (fn ((set, resultset), acc) => + fold (fn (set, resultset) => fn acc => if is_subset (resultset, set) then intersection (acc, set) else - acc) i_univ (i_sets ~~ resultsets) + acc) (i_sets ~~ resultsets) i_univ | lfp _ = raise REFUTE ("lfp_interpreter", "lfp: interpretation for function is not a node") @@ -3026,11 +3026,11 @@ "union: interpretation for set is not a node") (* interpretation -> interpretaion *) fun gfp (Node resultsets) = - List.foldl (fn ((set, resultset), acc) => + fold (fn (set, resultset) => fn acc => if is_subset (set, resultset) then union (acc, set) else - acc) i_univ (i_sets ~~ resultsets) + acc) (i_sets ~~ resultsets) i_univ | gfp _ = raise REFUTE ("gfp_interpreter", "gfp: interpretation for function is not a node") @@ -3128,8 +3128,7 @@ val HOLogic_insert = Const (@{const_name insert}, HOLogic_prodT --> HOLogic_setT --> HOLogic_setT) in - SOME (List.foldr (fn (pair, acc) => HOLogic_insert $ pair $ acc) - HOLogic_empty_set pairs) + SOME (fold_rev (fn pair => fn acc => HOLogic_insert $ pair $ acc) pairs HOLogic_empty_set) end | Type ("prop", []) => (case index_from_interpretation intr of diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/Tools/res_atp.ML --- a/src/HOL/Tools/res_atp.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/HOL/Tools/res_atp.ML Wed Nov 04 11:40:59 2009 +0100 @@ -1,4 +1,6 @@ -(* Author: Jia Meng, Cambridge University Computer Laboratory, NICTA *) +(* Title: HOL/Tools/res_atp.ML + Author: Jia Meng, Cambridge University Computer Laboratory, NICTA +*) signature RES_ATP = sig @@ -9,14 +11,14 @@ val get_relevant : int -> bool -> Proof.context * (thm list * 'a) -> thm list -> (thm * (string * int)) list val prepare_clauses : bool -> thm list -> thm list -> - (thm * (ResHolClause.axiom_name * ResHolClause.clause_id)) list -> - (thm * (ResHolClause.axiom_name * ResHolClause.clause_id)) list -> theory -> - ResHolClause.axiom_name vector * - (ResHolClause.clause list * ResHolClause.clause list * ResHolClause.clause list * - ResHolClause.clause list * ResClause.classrelClause list * ResClause.arityClause list) + (thm * (Res_HOL_Clause.axiom_name * Res_HOL_Clause.clause_id)) list -> + (thm * (Res_HOL_Clause.axiom_name * Res_HOL_Clause.clause_id)) list -> theory -> + Res_HOL_Clause.axiom_name vector * + (Res_HOL_Clause.clause list * Res_HOL_Clause.clause list * Res_HOL_Clause.clause list * + Res_HOL_Clause.clause list * Res_Clause.classrelClause list * Res_Clause.arityClause list) end; -structure ResAtp: RES_ATP = +structure Res_ATP: RES_ATP = struct @@ -236,10 +238,10 @@ let val cls = sort compare_pairs newpairs val accepted = List.take (cls, max_new) in - ResAxioms.trace_msg (fn () => ("Number of candidates, " ^ Int.toString nnew ^ + Res_Axioms.trace_msg (fn () => ("Number of candidates, " ^ Int.toString nnew ^ ", exceeds the limit of " ^ Int.toString (max_new))); - ResAxioms.trace_msg (fn () => ("Effective pass mark: " ^ Real.toString (#2 (List.last accepted)))); - ResAxioms.trace_msg (fn () => "Actually passed: " ^ + Res_Axioms.trace_msg (fn () => ("Effective pass mark: " ^ Real.toString (#2 (List.last accepted)))); + Res_Axioms.trace_msg (fn () => "Actually passed: " ^ space_implode ", " (map (fn (((_,(name,_)),_),_) => name) accepted)); (map #1 accepted, map #1 (List.drop (cls, max_new))) @@ -254,7 +256,7 @@ val rel_consts' = List.foldl add_const_typ_table rel_consts new_consts val newp = p + (1.0-p) / convergence in - ResAxioms.trace_msg (fn () => ("relevant this iteration: " ^ Int.toString (length newrels))); + Res_Axioms.trace_msg (fn () => "relevant this iteration: " ^ Int.toString (length newrels)); (map #1 newrels) @ (relevant_clauses max_new thy ctab newp rel_consts' (more_rejects@rejects)) end @@ -262,12 +264,12 @@ let val weight = clause_weight ctab rel_consts consts_typs in if p <= weight orelse (follow_defs andalso defines thy (#1 clsthm) rel_consts) - then (ResAxioms.trace_msg (fn () => (name ^ " clause " ^ Int.toString n ^ + then (Res_Axioms.trace_msg (fn () => (name ^ " clause " ^ Int.toString n ^ " passes: " ^ Real.toString weight)); relevant ((ax,weight)::newrels, rejects) axs) else relevant (newrels, ax::rejects) axs end - in ResAxioms.trace_msg (fn () => ("relevant_clauses, current pass mark = " ^ Real.toString p)); + in Res_Axioms.trace_msg (fn () => ("relevant_clauses, current pass mark = " ^ Real.toString p)); relevant ([],[]) end; @@ -276,12 +278,12 @@ then let val const_tab = List.foldl (count_axiom_consts theory_const thy) Symtab.empty axioms val goal_const_tab = get_goal_consts_typs thy goals - val _ = ResAxioms.trace_msg (fn () => ("Initial constants: " ^ + val _ = Res_Axioms.trace_msg (fn () => ("Initial constants: " ^ space_implode ", " (Symtab.keys goal_const_tab))); val rels = relevant_clauses max_new thy const_tab (pass_mark) goal_const_tab (map (pair_consts_typs_axiom theory_const thy) axioms) in - ResAxioms.trace_msg (fn () => ("Total relevant: " ^ Int.toString (length rels))); + Res_Axioms.trace_msg (fn () => ("Total relevant: " ^ Int.toString (length rels))); rels end else axioms; @@ -335,7 +337,7 @@ fun make_unique xs = let val ht = mk_clause_table 7000 in - ResAxioms.trace_msg (fn () => ("make_unique gets " ^ Int.toString (length xs) ^ " clauses")); + Res_Axioms.trace_msg (fn () => ("make_unique gets " ^ Int.toString (length xs) ^ " clauses")); app (ignore o Polyhash.peekInsert ht) xs; Polyhash.listItems ht end; @@ -352,11 +354,12 @@ fun valid_facts facts = Facts.fold_static (fn (name, ths) => - if run_blacklist_filter andalso is_package_def name then I + if Facts.is_concealed facts name orelse + (run_blacklist_filter andalso is_package_def name) then I else let val xname = Facts.extern facts name in if Name_Space.is_hidden xname then I - else cons (xname, filter_out ResAxioms.bad_for_atp ths) + else cons (xname, filter_out Res_Axioms.bad_for_atp ths) end) facts []; fun all_valid_thms ctxt = @@ -365,29 +368,29 @@ val local_facts = ProofContext.facts_of ctxt; in valid_facts global_facts @ valid_facts local_facts end; -fun multi_name a (th, (n,pairs)) = - (n+1, (a ^ "(" ^ Int.toString n ^ ")", th) :: pairs) +fun multi_name a th (n, pairs) = + (n + 1, (a ^ "(" ^ Int.toString n ^ ")", th) :: pairs); -fun add_single_names ((a, []), pairs) = pairs - | add_single_names ((a, [th]), pairs) = (a,th)::pairs - | add_single_names ((a, ths), pairs) = #2 (List.foldl (multi_name a) (1,pairs) ths); +fun add_single_names (a, []) pairs = pairs + | add_single_names (a, [th]) pairs = (a, th) :: pairs + | add_single_names (a, ths) pairs = #2 (fold (multi_name a) ths (1, pairs)); (*Ignore blacklisted basenames*) -fun add_multi_names ((a, ths), pairs) = - if (Long_Name.base_name a) mem_string ResAxioms.multi_base_blacklist then pairs - else add_single_names ((a, ths), pairs); +fun add_multi_names (a, ths) pairs = + if (Long_Name.base_name a) mem_string Res_Axioms.multi_base_blacklist then pairs + else add_single_names (a, ths) pairs; fun is_multi (a, ths) = length ths > 1 orelse String.isSuffix ".axioms" a; (*The single theorems go BEFORE the multiple ones. Blacklist is applied to all.*) fun name_thm_pairs ctxt = - let val (mults,singles) = List.partition is_multi (all_valid_thms ctxt) - val ht = mk_clause_table 900 (*ht for blacklisted theorems*) - fun blacklisted x = run_blacklist_filter andalso is_some (Polyhash.peek ht x) + let + val (mults, singles) = List.partition is_multi (all_valid_thms ctxt) + fun blacklisted (_, th) = + run_blacklist_filter andalso Res_Blacklist.blacklisted ctxt th in - app (fn th => ignore (Polyhash.peekInsert ht (th,()))) (ResBlacklist.get ctxt); - filter (not o blacklisted o #2) - (List.foldl add_single_names (List.foldl add_multi_names [] mults) singles) + filter_out blacklisted + (fold add_single_names singles (fold add_multi_names mults [])) end; fun check_named ("", th) = @@ -396,7 +399,7 @@ fun get_all_lemmas ctxt = let val included_thms = - tap (fn ths => ResAxioms.trace_msg + tap (fn ths => Res_Axioms.trace_msg (fn () => ("Including all " ^ Int.toString (length ths) ^ " theorems"))) (name_thm_pairs ctxt) in @@ -499,17 +502,14 @@ | Fol => true | Hol => false -fun ths_to_cls thy ths = - ResAxioms.cnf_rules_pairs thy (filter check_named (map ResAxioms.pairname ths)) - fun get_relevant max_new theory_const (ctxt, (chain_ths, th)) goal_cls = let val thy = ProofContext.theory_of ctxt val isFO = isFO thy goal_cls - val included_thms = get_all_lemmas ctxt - val included_cls = included_thms |> ResAxioms.cnf_rules_pairs thy |> make_unique - |> restrict_to_logic thy isFO - |> remove_unwanted_clauses + val included_cls = get_all_lemmas ctxt + |> Res_Axioms.cnf_rules_pairs thy |> make_unique + |> restrict_to_logic thy isFO + |> remove_unwanted_clauses in relevance_filter max_new theory_const thy included_cls (map prop_of goal_cls) end; @@ -519,24 +519,25 @@ fun prepare_clauses dfg goal_cls chain_ths axcls extra_cls thy = let (* add chain thms *) - val chain_cls = ths_to_cls thy chain_ths + val chain_cls = + Res_Axioms.cnf_rules_pairs thy (filter check_named (map Res_Axioms.pairname chain_ths)) val axcls = chain_cls @ axcls val extra_cls = chain_cls @ extra_cls val isFO = isFO thy goal_cls val ccls = subtract_cls goal_cls extra_cls - val _ = app (fn th => ResAxioms.trace_msg (fn _ => Display.string_of_thm_global thy th)) ccls + val _ = app (fn th => Res_Axioms.trace_msg (fn _ => Display.string_of_thm_global thy th)) ccls val ccltms = map prop_of ccls and axtms = map (prop_of o #1) extra_cls val subs = tfree_classes_of_terms ccltms and supers = tvar_classes_of_terms axtms and tycons = type_consts_of_terms thy (ccltms@axtms) (*TFrees in conjecture clauses; TVars in axiom clauses*) - val conjectures = ResHolClause.make_conjecture_clauses dfg thy ccls - val (_, extra_clauses) = ListPair.unzip (ResHolClause.make_axiom_clauses dfg thy extra_cls) - val (clnames,axiom_clauses) = ListPair.unzip (ResHolClause.make_axiom_clauses dfg thy axcls) - val helper_clauses = ResHolClause.get_helper_clauses dfg thy isFO (conjectures, extra_cls, []) - val (supers',arity_clauses) = ResClause.make_arity_clauses_dfg dfg thy tycons supers - val classrel_clauses = ResClause.make_classrel_clauses thy subs supers' + val conjectures = Res_HOL_Clause.make_conjecture_clauses dfg thy ccls + val (_, extra_clauses) = ListPair.unzip (Res_HOL_Clause.make_axiom_clauses dfg thy extra_cls) + val (clnames,axiom_clauses) = ListPair.unzip (Res_HOL_Clause.make_axiom_clauses dfg thy axcls) + val helper_clauses = Res_HOL_Clause.get_helper_clauses dfg thy isFO (conjectures, extra_cls, []) + val (supers',arity_clauses) = Res_Clause.make_arity_clauses_dfg dfg thy tycons supers + val classrel_clauses = Res_Clause.make_classrel_clauses thy subs supers' in (Vector.fromList clnames, (conjectures, axiom_clauses, extra_clauses, helper_clauses, classrel_clauses, arity_clauses)) diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/Tools/res_axioms.ML --- a/src/HOL/Tools/res_axioms.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/HOL/Tools/res_axioms.ML Wed Nov 04 11:40:59 2009 +0100 @@ -1,4 +1,5 @@ -(* Author: Jia Meng, Cambridge University Computer Laboratory +(* Title: HOL/Tools/res_axioms.ML + Author: Jia Meng, Cambridge University Computer Laboratory Transformation of axiom rules (elim/intro/etc) into CNF forms. *) @@ -22,7 +23,7 @@ val setup: theory -> theory end; -structure ResAxioms: RES_AXIOMS = +structure Res_Axioms: RES_AXIOMS = struct val trace = Unsynchronized.ref false; @@ -285,7 +286,7 @@ map (skolem_of_def o assume o (cterm_of (theory_of_thm th))) (assume_skofuns s th); -(*** Blacklisting (duplicated in ResAtp?) ***) +(*** Blacklisting (duplicated in Res_ATP?) ***) val max_lambda_nesting = 3; @@ -316,18 +317,17 @@ fun is_strange_thm th = case head_of (concl_of th) of - Const (a,_) => (a <> "Trueprop" andalso a <> "==") + Const (a, _) => (a <> "Trueprop" andalso a <> "==") | _ => false; fun bad_for_atp th = - Thm.is_internal th - orelse too_complex (prop_of th) + too_complex (prop_of th) orelse exists_type type_has_empty_sort (prop_of th) orelse is_strange_thm th; val multi_base_blacklist = ["defs","select_defs","update_defs","induct","inducts","split","splits","split_asm", - "cases","ext_cases"]; (*FIXME: put other record thms here, or use the "Internal" marker*) + "cases","ext_cases"]; (* FIXME put other record thms here, or declare as "noatp" *) (*Keep the full complexity of the original name*) fun flatten_name s = space_implode "_X" (Long_Name.explode s); @@ -387,14 +387,14 @@ fun cnf_rules_pairs_aux _ pairs [] = pairs | cnf_rules_pairs_aux thy pairs ((name,th)::ths) = let val pairs' = (pair_name_cls 0 (name, cnf_axiom thy th)) @ pairs - handle THM _ => pairs | ResClause.CLAUSE _ => pairs + handle THM _ => pairs | Res_Clause.CLAUSE _ => pairs in cnf_rules_pairs_aux thy pairs' ths end; (*The combination of rev and tail recursion preserves the original order*) fun cnf_rules_pairs thy l = cnf_rules_pairs_aux thy [] (rev l); -(**** Convert all facts of the theory into clauses (ResClause.clause, or ResHolClause.clause) ****) +(**** Convert all facts of the theory into clauses (Res_Clause.clause, or Res_HOL_Clause.clause) ****) local @@ -421,8 +421,10 @@ fun saturate_skolem_cache thy = let - val new_facts = (PureThy.facts_of thy, []) |-> Facts.fold_static (fn (name, ths) => - if already_seen thy name then I else cons (name, ths)); + val facts = PureThy.facts_of thy; + val new_facts = (facts, []) |-> Facts.fold_static (fn (name, ths) => + if Facts.is_concealed facts name orelse already_seen thy name then I + else cons (name, ths)); val new_thms = (new_facts, []) |-> fold (fn (name, ths) => if member (op =) multi_base_blacklist (Long_Name.base_name name) then I else fold_index (fn (i, th) => @@ -474,7 +476,7 @@ val remaining_hyps = filter_out (member (op aconv) (map Thm.term_of defs)) (map Thm.term_of hyps) val fixed = OldTerm.term_frees (concl_of st) @ - List.foldl (uncurry (union (op aconv))) [] (map OldTerm.term_frees remaining_hyps) + fold (union (op aconv)) (map OldTerm.term_frees remaining_hyps) [] in Seq.of_list [LocalDefs.expand (filter (is_okdef fixed o Thm.term_of) defs) st] end; diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/Tools/res_blacklist.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Tools/res_blacklist.ML Wed Nov 04 11:40:59 2009 +0100 @@ -0,0 +1,43 @@ +(* Title: HOL/Tools/res_blacklist.ML + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Author: Makarius + +Theorems blacklisted to sledgehammer. These theorems typically +produce clauses that are prolific (match too many equality or +membership literals) and relate to seldom-used facts. Some duplicate +other rules. +*) + +signature RES_BLACKLIST = +sig + val setup: theory -> theory + val blacklisted: Proof.context -> thm -> bool + val add: attribute + val del: attribute +end; + +structure Res_Blacklist: RES_BLACKLIST = +struct + +structure Data = GenericDataFun +( + type T = thm Termtab.table; + val empty = Termtab.empty; + val extend = I; + fun merge _ tabs = Termtab.merge (K true) tabs; +); + +fun blacklisted ctxt th = + Termtab.defined (Data.get (Context.Proof ctxt)) (Thm.full_prop_of th); + +fun add_thm th = Data.map (Termtab.update (Thm.full_prop_of th, th)); +fun del_thm th = Data.map (Termtab.delete_safe (Thm.full_prop_of th)); + +val add = Thm.declaration_attribute add_thm; +val del = Thm.declaration_attribute del_thm; + +val setup = + Attrib.setup @{binding noatp} (Attrib.add_del add del) "sledgehammer blacklisting" #> + PureThy.add_thms_dynamic (@{binding noatp}, map #2 o Termtab.dest o Data.get); + +end; diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/Tools/res_clause.ML --- a/src/HOL/Tools/res_clause.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/HOL/Tools/res_clause.ML Wed Nov 04 11:40:59 2009 +0100 @@ -1,11 +1,12 @@ -(* Author: Jia Meng, Cambridge University Computer Laboratory - Copyright 2004 University of Cambridge +(* Title: HOL/Tools/res_clause.ML + Author: Jia Meng, Cambridge University Computer Laboratory -Storing/printing FOL clauses and arity clauses. -Typed equality is treated differently. +Storing/printing FOL clauses and arity clauses. Typed equality is +treated differently. + +FIXME: combine with res_hol_clause! *) -(*FIXME: combine with res_hol_clause!*) signature RES_CLAUSE = sig val schematic_var_prefix: string @@ -76,7 +77,7 @@ val tptp_classrelClause : classrelClause -> string end -structure ResClause: RES_CLAUSE = +structure Res_Clause: RES_CLAUSE = struct val schematic_var_prefix = "V_"; diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/Tools/res_hol_clause.ML --- a/src/HOL/Tools/res_hol_clause.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/HOL/Tools/res_hol_clause.ML Wed Nov 04 11:40:59 2009 +0100 @@ -1,5 +1,5 @@ -(* - Author: Jia Meng, NICTA +(* Title: HOL/Tools/res_hol_clause.ML + Author: Jia Meng, NICTA FOL clauses translated from HOL formulae. *) @@ -17,13 +17,13 @@ type polarity = bool type clause_id = int datatype combterm = - CombConst of string * ResClause.fol_type * ResClause.fol_type list (*Const and Free*) - | CombVar of string * ResClause.fol_type + CombConst of string * Res_Clause.fol_type * Res_Clause.fol_type list (*Const and Free*) + | CombVar of string * Res_Clause.fol_type | CombApp of combterm * combterm datatype literal = Literal of polarity * combterm datatype clause = Clause of {clause_id: clause_id, axiom_name: axiom_name, th: thm, - kind: ResClause.kind,literals: literal list, ctypes_sorts: typ list} - val type_of_combterm: combterm -> ResClause.fol_type + kind: Res_Clause.kind,literals: literal list, ctypes_sorts: typ list} + val type_of_combterm: combterm -> Res_Clause.fol_type val strip_comb: combterm -> combterm * combterm list val literals_of_term: theory -> term -> literal list * typ list exception TOO_TRIVIAL @@ -38,18 +38,18 @@ clause list val tptp_write_file: bool -> Path.T -> clause list * clause list * clause list * clause list * - ResClause.classrelClause list * ResClause.arityClause list -> + Res_Clause.classrelClause list * Res_Clause.arityClause list -> int * int val dfg_write_file: bool -> Path.T -> clause list * clause list * clause list * clause list * - ResClause.classrelClause list * ResClause.arityClause list -> + Res_Clause.classrelClause list * Res_Clause.arityClause list -> int * int end -structure ResHolClause: RES_HOL_CLAUSE = +structure Res_HOL_Clause: RES_HOL_CLAUSE = struct -structure RC = ResClause; +structure RC = Res_Clause; (* FIXME avoid structure alias *) (* theorems for combinators and function extensionality *) val ext = thm "HOL.ext"; @@ -404,7 +404,7 @@ else ct; fun cnf_helper_thms thy = - ResAxioms.cnf_rules_pairs thy o map ResAxioms.pairname + Res_Axioms.cnf_rules_pairs thy o map Res_Axioms.pairname fun get_helper_clauses dfg thy isFO (conjectures, axcls, user_lemmas) = if isFO then [] (*first-order*) @@ -454,7 +454,7 @@ fold count_constants_lit literals (const_min_arity, const_needs_hBOOL); fun display_arity const_needs_hBOOL (c,n) = - ResAxioms.trace_msg (fn () => "Constant: " ^ c ^ " arity:\t" ^ Int.toString n ^ + Res_Axioms.trace_msg (fn () => "Constant: " ^ c ^ " arity:\t" ^ Int.toString n ^ (if needs_hBOOL const_needs_hBOOL c then " needs hBOOL" else "")); fun count_constants (conjectures, _, extra_clauses, helper_clauses, _, _) = @@ -527,4 +527,5 @@ length helper_clauses + 1, length tfree_clss + length dfg_clss) end; -end +end; + diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/Tools/res_reconstruct.ML --- a/src/HOL/Tools/res_reconstruct.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/HOL/Tools/res_reconstruct.ML Wed Nov 04 11:40:59 2009 +0100 @@ -1,8 +1,9 @@ -(* Author: L C Paulson and Claire Quigley, Cambridge University Computer Laboratory *) +(* Title: HOL/Tools/res_reconstruct.ML + Author: Lawrence C Paulson and Claire Quigley, Cambridge University Computer Laboratory -(***************************************************************************) -(* Code to deal with the transfer of proofs from a prover process *) -(***************************************************************************) +Transfer of proofs from external provers. +*) + signature RES_RECONSTRUCT = sig val chained_hint: string @@ -23,13 +24,13 @@ string * string vector * (int * int) * Proof.context * thm * int -> string * string list end; -structure ResReconstruct : RES_RECONSTRUCT = +structure Res_Reconstruct : RES_RECONSTRUCT = struct val trace_path = Path.basic "atp_trace"; fun trace s = - if ! ResAxioms.trace then File.append (File.tmp_path trace_path) s + if ! Res_Axioms.trace then File.append (File.tmp_path trace_path) s else (); fun string_of_thm ctxt = PrintMode.setmp [] (Display.string_of_thm ctxt); @@ -106,12 +107,12 @@ (*If string s has the prefix s1, return the result of deleting it.*) fun strip_prefix s1 s = if String.isPrefix s1 s - then SOME (ResClause.undo_ascii_of (String.extract (s, size s1, NONE))) + then SOME (Res_Clause.undo_ascii_of (String.extract (s, size s1, NONE))) else NONE; (*Invert the table of translations between Isabelle and ATPs*) val type_const_trans_table_inv = - Symtab.make (map swap (Symtab.dest ResClause.type_const_trans_table)); + Symtab.make (map swap (Symtab.dest Res_Clause.type_const_trans_table)); fun invert_type_const c = case Symtab.lookup type_const_trans_table_inv c of @@ -129,15 +130,15 @@ | Br (a,ts) => let val Ts = map type_of_stree ts in - case strip_prefix ResClause.tconst_prefix a of + case strip_prefix Res_Clause.tconst_prefix a of SOME b => Type(invert_type_const b, Ts) | NONE => if not (null ts) then raise STREE t (*only tconsts have type arguments*) else - case strip_prefix ResClause.tfree_prefix a of + case strip_prefix Res_Clause.tfree_prefix a of SOME b => TFree("'" ^ b, HOLogic.typeS) | NONE => - case strip_prefix ResClause.tvar_prefix a of + case strip_prefix Res_Clause.tvar_prefix a of SOME b => make_tvar b | NONE => make_tvar a (*Variable from the ATP, say X1*) end; @@ -145,7 +146,7 @@ (*Invert the table of translations between Isabelle and ATPs*) val const_trans_table_inv = Symtab.update ("fequal", "op =") - (Symtab.make (map swap (Symtab.dest ResClause.const_trans_table))); + (Symtab.make (map swap (Symtab.dest Res_Clause.const_trans_table))); fun invert_const c = case Symtab.lookup const_trans_table_inv c of @@ -166,7 +167,7 @@ | Br ("hBOOL",[t]) => term_of_stree [] thy t (*ignore hBOOL*) | Br ("hAPP",[t,u]) => term_of_stree (u::args) thy t | Br (a,ts) => - case strip_prefix ResClause.const_prefix a of + case strip_prefix Res_Clause.const_prefix a of SOME "equal" => list_comb(Const ("op =", HOLogic.typeT), List.map (term_of_stree [] thy) ts) | SOME b => @@ -180,10 +181,10 @@ | NONE => (*a variable, not a constant*) let val T = HOLogic.typeT val opr = (*a Free variable is typically a Skolem function*) - case strip_prefix ResClause.fixed_var_prefix a of + case strip_prefix Res_Clause.fixed_var_prefix a of SOME b => Free(b,T) | NONE => - case strip_prefix ResClause.schematic_var_prefix a of + case strip_prefix Res_Clause.schematic_var_prefix a of SOME b => make_var (b,T) | NONE => make_var (a,T) (*Variable from the ATP, say X1*) in list_comb (opr, List.map (term_of_stree [] thy) (ts@args)) end; @@ -193,7 +194,7 @@ | constraint_of_stree pol t = case t of Int _ => raise STREE t | Br (a,ts) => - (case (strip_prefix ResClause.class_prefix a, map type_of_stree ts) of + (case (strip_prefix Res_Clause.class_prefix a, map type_of_stree ts) of (SOME b, [T]) => (pol, b, T) | _ => raise STREE t); @@ -443,11 +444,11 @@ val _ = trace (Int.toString (length nonnull_lines) ^ " nonnull_lines extracted\n") val (_,lines) = List.foldr (add_wanted_prfline ctxt) (0,[]) nonnull_lines val _ = trace (Int.toString (length lines) ^ " lines extracted\n") - val (ccls,fixes) = ResAxioms.neg_conjecture_clauses ctxt th sgno + val (ccls,fixes) = Res_Axioms.neg_conjecture_clauses ctxt th sgno val _ = trace (Int.toString (length ccls) ^ " conjecture clauses\n") val ccls = map forall_intr_vars ccls val _ = - if ! ResAxioms.trace then app (fn th => trace ("\nccl: " ^ string_of_thm ctxt th)) ccls + if ! Res_Axioms.trace then app (fn th => trace ("\nccl: " ^ string_of_thm ctxt th)) ccls else () val ilines = isar_lines ctxt (map prop_of ccls) (stringify_deps thm_names [] lines) val _ = trace "\ndecode_tstp_file: finishing\n" @@ -456,124 +457,128 @@ end; - (*=== EXTRACTING PROOF-TEXT === *) +(*=== EXTRACTING PROOF-TEXT === *) - val begin_proof_strings = ["# SZS output start CNFRefutation.", - "=========== Refutation ==========", +val begin_proof_strings = ["# SZS output start CNFRefutation.", + "=========== Refutation ==========", "Here is a proof"]; - val end_proof_strings = ["# SZS output end CNFRefutation", - "======= End of refutation =======", + +val end_proof_strings = ["# SZS output end CNFRefutation", + "======= End of refutation =======", "Formulae used in the proof"]; - fun get_proof_extract proof = - let + +fun get_proof_extract proof = + let (*splits to_split by the first possible of a list of splitters*) val (begin_string, end_string) = (find_first (fn s => String.isSubstring s proof) begin_proof_strings, find_first (fn s => String.isSubstring s proof) end_proof_strings) - in - if is_none begin_string orelse is_none end_string - then error "Could not extract proof (no substring indicating a proof)" - else proof |> first_field (the begin_string) |> the |> snd - |> first_field (the end_string) |> the |> fst end; + in + if is_none begin_string orelse is_none end_string + then error "Could not extract proof (no substring indicating a proof)" + else proof |> first_field (the begin_string) |> the |> snd + |> first_field (the end_string) |> the |> fst + end; (* ==== CHECK IF PROOF OF E OR VAMPIRE WAS SUCCESSFUL === *) - val failure_strings_E = ["SZS status: Satisfiable","SZS status Satisfiable", - "SZS status: ResourceOut","SZS status ResourceOut","# Cannot determine problem status"]; - val failure_strings_vampire = ["Satisfiability detected", "Refutation not found", "CANNOT PROVE"]; - val failure_strings_SPASS = ["SPASS beiseite: Completion found.", - "SPASS beiseite: Ran out of time.", "SPASS beiseite: Maximal number of loops exceeded."]; - val failure_strings_remote = ["Remote-script could not extract proof"]; - fun find_failure proof = - let val failures = - map_filter (fn s => if String.isSubstring s proof then SOME s else NONE) - (failure_strings_E @ failure_strings_vampire @ failure_strings_SPASS @ failure_strings_remote) - val correct = null failures andalso - exists (fn s => String.isSubstring s proof) begin_proof_strings andalso - exists (fn s => String.isSubstring s proof) end_proof_strings - in - if correct then NONE - else if null failures then SOME "Output of ATP not in proper format" - else SOME (hd failures) end; +val failure_strings_E = ["SZS status: Satisfiable","SZS status Satisfiable", + "SZS status: ResourceOut","SZS status ResourceOut","# Cannot determine problem status"]; +val failure_strings_vampire = ["Satisfiability detected", "Refutation not found", "CANNOT PROVE"]; +val failure_strings_SPASS = ["SPASS beiseite: Completion found.", + "SPASS beiseite: Ran out of time.", "SPASS beiseite: Maximal number of loops exceeded."]; +val failure_strings_remote = ["Remote-script could not extract proof"]; +fun find_failure proof = + let val failures = + map_filter (fn s => if String.isSubstring s proof then SOME s else NONE) + (failure_strings_E @ failure_strings_vampire @ failure_strings_SPASS @ failure_strings_remote) + val correct = null failures andalso + exists (fn s => String.isSubstring s proof) begin_proof_strings andalso + exists (fn s => String.isSubstring s proof) end_proof_strings + in + if correct then NONE + else if null failures then SOME "Output of ATP not in proper format" + else SOME (hd failures) end; - (* === EXTRACTING LEMMAS === *) - (* lines have the form "cnf(108, axiom, ...", - the number (108) has to be extracted)*) - fun get_step_nums false proofextract = - let val toks = String.tokens (not o Char.isAlphaNum) - fun inputno ("cnf"::ntok::"axiom"::_) = Int.fromString ntok - | inputno ("cnf"::ntok::"negated"::"conjecture"::_) = Int.fromString ntok - | inputno _ = NONE - val lines = split_lines proofextract - in map_filter (inputno o toks) lines end - (*String contains multiple lines. We want those of the form - "253[0:Inp] et cetera..." - A list consisting of the first number in each line is returned. *) - | get_step_nums true proofextract = - let val toks = String.tokens (not o Char.isAlphaNum) - fun inputno (ntok::"0"::"Inp"::_) = Int.fromString ntok - | inputno _ = NONE - val lines = split_lines proofextract - in map_filter (inputno o toks) lines end - - (*extracting lemmas from tstp-output between the lines from above*) - fun extract_lemmas get_step_nums (proof, thm_names, conj_count, _, _, _) = +(* === EXTRACTING LEMMAS === *) +(* lines have the form "cnf(108, axiom, ...", +the number (108) has to be extracted)*) +fun get_step_nums false proofextract = + let val toks = String.tokens (not o Char.isAlphaNum) + fun inputno ("cnf"::ntok::"axiom"::_) = Int.fromString ntok + | inputno ("cnf"::ntok::"negated"::"conjecture"::_) = Int.fromString ntok + | inputno _ = NONE + val lines = split_lines proofextract + in map_filter (inputno o toks) lines end +(*String contains multiple lines. We want those of the form + "253[0:Inp] et cetera..." + A list consisting of the first number in each line is returned. *) +| get_step_nums true proofextract = + let val toks = String.tokens (not o Char.isAlphaNum) + fun inputno (ntok::"0"::"Inp"::_) = Int.fromString ntok + | inputno _ = NONE + val lines = split_lines proofextract + in map_filter (inputno o toks) lines end + +(*extracting lemmas from tstp-output between the lines from above*) +fun extract_lemmas get_step_nums (proof, thm_names, conj_count, _, _, _) = + let + (* get the names of axioms from their numbers*) + fun get_axiom_names thm_names step_nums = let - (* get the names of axioms from their numbers*) - fun get_axiom_names thm_names step_nums = - let - val last_axiom = Vector.length thm_names - fun is_axiom n = n <= last_axiom - fun is_conj n = n >= fst conj_count andalso n < fst conj_count + snd conj_count - fun getname i = Vector.sub(thm_names, i-1) - in - (sort_distinct string_ord (filter (fn x => x <> "??.unknown") - (map getname (filter is_axiom step_nums))), - exists is_conj step_nums) - end - val proofextract = get_proof_extract proof + val last_axiom = Vector.length thm_names + fun is_axiom n = n <= last_axiom + fun is_conj n = n >= fst conj_count andalso n < fst conj_count + snd conj_count + fun getname i = Vector.sub(thm_names, i-1) in - get_axiom_names thm_names (get_step_nums proofextract) - end; + (sort_distinct string_ord (filter (fn x => x <> "??.unknown") + (map getname (filter is_axiom step_nums))), + exists is_conj step_nums) + end + val proofextract = get_proof_extract proof + in + get_axiom_names thm_names (get_step_nums proofextract) + end; - (*Used to label theorems chained into the sledgehammer call*) - val chained_hint = "CHAINED"; - val nochained = filter_out (fn y => y = chained_hint) - - (* metis-command *) - fun metis_line [] = "apply metis" - | metis_line xs = "apply (metis " ^ space_implode " " xs ^ ")" +(*Used to label theorems chained into the sledgehammer call*) +val chained_hint = "CHAINED"; +val nochained = filter_out (fn y => y = chained_hint) + +(* metis-command *) +fun metis_line [] = "apply metis" + | metis_line xs = "apply (metis " ^ space_implode " " xs ^ ")" - (* atp_minimize [atp=] *) - fun minimize_line _ [] = "" - | minimize_line name lemmas = "For minimizing the number of lemmas try this command:\n" ^ - (Markup.markup Markup.sendback) ("atp_minimize [atp=" ^ name ^ "] " ^ - space_implode " " (nochained lemmas)) +(* atp_minimize [atp=] *) +fun minimize_line _ [] = "" + | minimize_line name lemmas = "For minimizing the number of lemmas try this command:\n" ^ + (Markup.markup Markup.sendback) ("atp_minimize [atp=" ^ name ^ "] " ^ + space_implode " " (nochained lemmas)) - fun sendback_metis_nochained lemmas = - (Markup.markup Markup.sendback o metis_line) (nochained lemmas) +fun sendback_metis_nochained lemmas = + (Markup.markup Markup.sendback o metis_line) (nochained lemmas) - fun lemma_list dfg name result = - let val (lemmas, used_conj) = extract_lemmas (get_step_nums dfg) result - in (sendback_metis_nochained lemmas ^ "\n" ^ minimize_line name lemmas ^ - (if used_conj then "" - else "\nWarning: Goal is provable because context is inconsistent."), - nochained lemmas) - end; +fun lemma_list dfg name result = + let val (lemmas, used_conj) = extract_lemmas (get_step_nums dfg) result + in (sendback_metis_nochained lemmas ^ "\n" ^ minimize_line name lemmas ^ + (if used_conj then "" + else "\nWarning: Goal is provable because context is inconsistent."), + nochained lemmas) + end; - (* === Extracting structured Isar-proof === *) - fun structured_proof name (result as (proof, thm_names, conj_count, ctxt, goal, subgoalno)) = - let - (*Could use split_lines, but it can return blank lines...*) - val lines = String.tokens (equal #"\n"); - val nospaces = String.translate (fn c => if Char.isSpace c then "" else str c) - val proofextract = get_proof_extract proof - val cnfs = filter (String.isPrefix "cnf(") (map nospaces (lines proofextract)) - val (one_line_proof, lemma_names) = lemma_list false name result - val structured = if chained_hint mem_string (String.tokens (fn c => c = #" ") one_line_proof) then "" - else decode_tstp_file cnfs ctxt goal subgoalno thm_names - in - (one_line_proof ^ "\n\n" ^ (Markup.markup Markup.sendback) structured, lemma_names) - end +(* === Extracting structured Isar-proof === *) +fun structured_proof name (result as (proof, thm_names, conj_count, ctxt, goal, subgoalno)) = + let + (*Could use split_lines, but it can return blank lines...*) + val lines = String.tokens (equal #"\n"); + val nospaces = String.translate (fn c => if Char.isSpace c then "" else str c) + val proofextract = get_proof_extract proof + val cnfs = filter (String.isPrefix "cnf(") (map nospaces (lines proofextract)) + val (one_line_proof, lemma_names) = lemma_list false name result + val structured = + if chained_hint mem_string (String.tokens (fn c => c = #" ") one_line_proof) then "" + else decode_tstp_file cnfs ctxt goal subgoalno thm_names + in + (one_line_proof ^ "\n\n" ^ Markup.markup Markup.sendback structured, lemma_names) +end end; diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/Tools/rewrite_hol_proof.ML --- a/src/HOL/Tools/rewrite_hol_proof.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/HOL/Tools/rewrite_hol_proof.ML Wed Nov 04 11:40:59 2009 +0100 @@ -15,8 +15,8 @@ open Proofterm; -val rews = map (pairself (ProofSyntax.proof_of_term @{theory} true) o - Logic.dest_equals o Logic.varify o ProofSyntax.read_term @{theory} propT) +val rews = map (pairself (Proof_Syntax.proof_of_term @{theory} true) o + Logic.dest_equals o Logic.varify o Proof_Syntax.read_term @{theory} propT) (** eliminate meta-equality rules **) diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/Tools/sat_solver.ML --- a/src/HOL/Tools/sat_solver.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/HOL/Tools/sat_solver.ML Wed Nov 04 11:40:59 2009 +0100 @@ -272,7 +272,7 @@ else parse_lines lines in - (parse_lines o (List.filter (fn l => l <> "")) o split_lines o File.read) path + (parse_lines o filter (fn l => l <> "") o split_lines o File.read) path end; (* ------------------------------------------------------------------------- *) @@ -352,7 +352,7 @@ o (map int_from_string) o (maps (String.fields (fn c => c mem [#" ", #"\t", #"\n"]))) o filter_preamble - o (List.filter (fn l => l <> "")) + o filter (fn l => l <> "") o split_lines o File.read) path diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/Tools/simpdata.ML --- a/src/HOL/Tools/simpdata.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/HOL/Tools/simpdata.ML Wed Nov 04 11:40:59 2009 +0100 @@ -64,8 +64,8 @@ else let val Ps = map (fn k => Free ("P" ^ string_of_int k, propT)) (1 upto j); - fun mk_simp_implies Q = List.foldr (fn (R, S) => - Const ("HOL.simp_implies", propT --> propT --> propT) $ R $ S) Q Ps + fun mk_simp_implies Q = fold_rev (fn R => fn S => + Const ("HOL.simp_implies", propT --> propT --> propT) $ R $ S) Ps Q val aT = TFree ("'a", HOLogic.typeS); val x = Free ("x", aT); val y = Free ("y", aT) diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/Tools/split_rule.ML --- a/src/HOL/Tools/split_rule.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/HOL/Tools/split_rule.ML Wed Nov 04 11:40:59 2009 +0100 @@ -118,7 +118,7 @@ fst (fold_rev complete_split_rule_var vars (rl, xs)) |> remove_internal_split |> Drule.standard - |> RuleCases.save rl + |> Rule_Cases.save rl end; @@ -138,7 +138,7 @@ |> fold_index one_goal xss |> Simplifier.full_simplify split_rule_ss |> Drule.standard - |> RuleCases.save rl + |> Rule_Cases.save rl end; diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/Tools/typecopy.ML --- a/src/HOL/Tools/typecopy.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/HOL/Tools/typecopy.ML Wed Nov 04 11:40:59 2009 +0100 @@ -44,8 +44,8 @@ (* interpretation of type copies *) -structure TypecopyInterpretation = InterpretationFun(type T = string val eq = op =); -val interpretation = TypecopyInterpretation.interpretation; +structure Typecopy_Interpretation = Interpretation(type T = string val eq = op =); +val interpretation = Typecopy_Interpretation.interpretation; (* introducing typecopies *) @@ -76,7 +76,7 @@ in thy |> (TypecopyData.map o Symtab.update_new) (tyco, info) - |> TypecopyInterpretation.data tyco + |> Typecopy_Interpretation.data tyco |> pair (tyco, info) end in @@ -126,7 +126,7 @@ end; val setup = - TypecopyInterpretation.init + Typecopy_Interpretation.init #> interpretation add_default_code end; diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/Tools/typedef.ML --- a/src/HOL/Tools/typedef.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/HOL/Tools/typedef.ML Wed Nov 04 11:40:59 2009 +0100 @@ -53,8 +53,8 @@ fun declare_type_name a = Variable.declare_constraints (Logic.mk_type (TFree (a, dummyS))); -structure TypedefInterpretation = InterpretationFun(type T = string val eq = op =); -val interpretation = TypedefInterpretation.interpretation; +structure Typedef_Interpretation = Interpretation(type T = string val eq = op =); +val interpretation = Typedef_Interpretation.interpretation; fun prepare_typedef prep_term def name (t, vs, mx) raw_set opt_morphs thy = let @@ -115,7 +115,7 @@ theory |> Sign.add_consts_i [(name, setT', NoSyn)] |> PureThy.add_defs false [Thm.no_attributes (apfst (Binding.name) - (PrimitiveDefs.mk_defpair (setC, set)))] + (Primitive_Defs.mk_defpair (setC, set)))] |-> (fn [th] => pair (SOME th)) else (NONE, theory); fun contract_def NONE th = th @@ -152,13 +152,13 @@ ((Binding.suffix_name "_inject" Rep_name, make @{thm type_definition.Rep_inject}), []), ((Binding.suffix_name "_inject" Abs_name, make @{thm type_definition.Abs_inject}), []), ((Binding.suffix_name "_cases" Rep_name, make @{thm type_definition.Rep_cases}), - [RuleCases.case_names [Binding.name_of Rep_name], Induct.cases_pred full_name]), + [Rule_Cases.case_names [Binding.name_of Rep_name], Induct.cases_pred full_name]), ((Binding.suffix_name "_cases" Abs_name, make @{thm type_definition.Abs_cases}), - [RuleCases.case_names [Binding.name_of Abs_name], Induct.cases_type full_tname]), + [Rule_Cases.case_names [Binding.name_of Abs_name], Induct.cases_type full_tname]), ((Binding.suffix_name "_induct" Rep_name, make @{thm type_definition.Rep_induct}), - [RuleCases.case_names [Binding.name_of Rep_name], Induct.induct_pred full_name]), + [Rule_Cases.case_names [Binding.name_of Rep_name], Induct.induct_pred full_name]), ((Binding.suffix_name "_induct" Abs_name, make @{thm type_definition.Abs_induct}), - [RuleCases.case_names [Binding.name_of Abs_name], Induct.induct_type full_tname])] + [Rule_Cases.case_names [Binding.name_of Abs_name], Induct.induct_type full_tname])] ||> Sign.parent_path; val info = {rep_type = oldT, abs_type = newT, Rep_name = full Rep_name, Abs_name = full Abs_name, @@ -169,7 +169,7 @@ in thy2 |> put_info full_tname info - |> TypedefInterpretation.data full_tname + |> Typedef_Interpretation.data full_tname |> pair (full_tname, info) end); @@ -264,6 +264,6 @@ end; -val setup = TypedefInterpretation.init; +val setup = Typedef_Interpretation.init; end; diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/Typerep.thy --- a/src/HOL/Typerep.thy Thu Oct 29 16:23:57 2009 +0100 +++ b/src/HOL/Typerep.thy Wed Nov 04 11:40:59 2009 +0100 @@ -29,7 +29,7 @@ | typerep_tr' _ T ts = raise Match; in Sign.add_syntax_i - [("_TYPEREP", SimpleSyntax.read_typ "type => logic", Delimfix "(1TYPEREP/(1'(_')))")] + [("_TYPEREP", Simple_Syntax.read_typ "type => logic", Delimfix "(1TYPEREP/(1'(_')))")] #> Sign.add_trfuns ([], [("_TYPEREP", typerep_tr)], [], []) #> Sign.add_trfunsT [(@{const_syntax typerep}, typerep_tr')] end diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/ex/Codegenerator_Candidates.thy --- a/src/HOL/ex/Codegenerator_Candidates.thy Thu Oct 29 16:23:57 2009 +0100 +++ b/src/HOL/ex/Codegenerator_Candidates.thy Wed Nov 04 11:40:59 2009 +0100 @@ -9,7 +9,6 @@ AssocList Binomial Fset - Commutative_Ring Enum List_Prefix Nat_Infinity @@ -22,7 +21,7 @@ Tree While_Combinator Word - "~~/src/HOL/ex/Commutative_Ring_Complete" + "~~/src/HOL/Decision_Procs/Commutative_Ring_Complete" "~~/src/HOL/ex/Records" begin diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/ex/Commutative_RingEx.thy --- a/src/HOL/ex/Commutative_RingEx.thy Thu Oct 29 16:23:57 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,50 +0,0 @@ -(* ID: $Id$ - Author: Bernhard Haeupler -*) - -header {* Some examples demonstrating the comm-ring method *} - -theory Commutative_RingEx -imports Commutative_Ring -begin - -lemma "4*(x::int)^5*y^3*x^2*3 + x*z + 3^5 = 12*x^7*y^3 + z*x + 243" -by comm_ring - -lemma "((x::int) + y)^2 = x^2 + y^2 + 2*x*y" -by comm_ring - -lemma "((x::int) + y)^3 = x^3 + y^3 + 3*x^2*y + 3*y^2*x" -by comm_ring - -lemma "((x::int) - y)^3 = x^3 + 3*x*y^2 + (-3)*y*x^2 - y^3" -by comm_ring - -lemma "((x::int) - y)^2 = x^2 + y^2 - 2*x*y" -by comm_ring - -lemma " ((a::int) + b + c)^2 = a^2 + b^2 + c^2 + 2*a*b + 2*b*c + 2*a*c" -by comm_ring - -lemma "((a::int) - b - c)^2 = a^2 + b^2 + c^2 - 2*a*b + 2*b*c - 2*a*c" -by comm_ring - -lemma "(a::int)*b + a*c = a*(b+c)" -by comm_ring - -lemma "(a::int)^2 - b^2 = (a - b) * (a + b)" -by comm_ring - -lemma "(a::int)^3 - b^3 = (a - b) * (a^2 + a*b + b^2)" -by comm_ring - -lemma "(a::int)^3 + b^3 = (a + b) * (a^2 - a*b + b^2)" -by comm_ring - -lemma "(a::int)^4 - b^4 = (a - b) * (a + b)*(a^2 + b^2)" -by comm_ring - -lemma "(a::int)^10 - b^10 = (a - b) * (a^9 + a^8*b + a^7*b^2 + a^6*b^3 + a^5*b^4 + a^4*b^5 + a^3*b^6 + a^2*b^7 + a*b^8 + b^9 )" -by comm_ring - -end diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/ex/Commutative_Ring_Complete.thy --- a/src/HOL/ex/Commutative_Ring_Complete.thy Thu Oct 29 16:23:57 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,391 +0,0 @@ -(* Author: Bernhard Haeupler - -This theory is about of the relative completeness of method comm-ring -method. As long as the reified atomic polynomials of type 'a pol are -in normal form, the cring method is complete. -*) - -header {* Proof of the relative completeness of method comm-ring *} - -theory Commutative_Ring_Complete -imports Commutative_Ring -begin - -text {* Formalization of normal form *} -fun - isnorm :: "('a::{comm_ring}) pol \ bool" -where - "isnorm (Pc c) \ True" - | "isnorm (Pinj i (Pc c)) \ False" - | "isnorm (Pinj i (Pinj j Q)) \ False" - | "isnorm (Pinj 0 P) \ False" - | "isnorm (Pinj i (PX Q1 j Q2)) \ isnorm (PX Q1 j Q2)" - | "isnorm (PX P 0 Q) \ False" - | "isnorm (PX (Pc c) i Q) \ c \ 0 \ isnorm Q" - | "isnorm (PX (PX P1 j (Pc c)) i Q) \ c \ 0 \ isnorm (PX P1 j (Pc c)) \ isnorm Q" - | "isnorm (PX P i Q) \ isnorm P \ isnorm Q" - -(* Some helpful lemmas *) -lemma norm_Pinj_0_False:"isnorm (Pinj 0 P) = False" -by(cases P, auto) - -lemma norm_PX_0_False:"isnorm (PX (Pc 0) i Q) = False" -by(cases i, auto) - -lemma norm_Pinj:"isnorm (Pinj i Q) \ isnorm Q" -by(cases i,simp add: norm_Pinj_0_False norm_PX_0_False,cases Q) auto - -lemma norm_PX2:"isnorm (PX P i Q) \ isnorm Q" -by(cases i, auto, cases P, auto, case_tac pol2, auto) - -lemma norm_PX1:"isnorm (PX P i Q) \ isnorm P" -by(cases i, auto, cases P, auto, case_tac pol2, auto) - -lemma mkPinj_cn:"\y~=0; isnorm Q\ \ isnorm (mkPinj y Q)" -apply(auto simp add: mkPinj_def norm_Pinj_0_False split: pol.split) -apply(case_tac nat, auto simp add: norm_Pinj_0_False) -by(case_tac pol, auto) (case_tac y, auto) - -lemma norm_PXtrans: - assumes A:"isnorm (PX P x Q)" and "isnorm Q2" - shows "isnorm (PX P x Q2)" -proof(cases P) - case (PX p1 y p2) from prems show ?thesis by(cases x, auto, cases p2, auto) -next - case Pc from prems show ?thesis by(cases x, auto) -next - case Pinj from prems show ?thesis by(cases x, auto) -qed - -lemma norm_PXtrans2: assumes A:"isnorm (PX P x Q)" and "isnorm Q2" shows "isnorm (PX P (Suc (n+x)) Q2)" -proof(cases P) - case (PX p1 y p2) - from prems show ?thesis by(cases x, auto, cases p2, auto) -next - case Pc - from prems show ?thesis by(cases x, auto) -next - case Pinj - from prems show ?thesis by(cases x, auto) -qed - -text {* mkPX conserves normalizedness (@{text "_cn"}) *} -lemma mkPX_cn: - assumes "x \ 0" and "isnorm P" and "isnorm Q" - shows "isnorm (mkPX P x Q)" -proof(cases P) - case (Pc c) - from prems show ?thesis by (cases x) (auto simp add: mkPinj_cn mkPX_def) -next - case (Pinj i Q) - from prems show ?thesis by (cases x) (auto simp add: mkPinj_cn mkPX_def) -next - case (PX P1 y P2) - from prems have Y0:"y>0" by(cases y, auto) - from prems have "isnorm P1" "isnorm P2" by (auto simp add: norm_PX1[of P1 y P2] norm_PX2[of P1 y P2]) - with prems Y0 show ?thesis by (cases x, auto simp add: mkPX_def norm_PXtrans2[of P1 y _ Q _], cases P2, auto) -qed - -text {* add conserves normalizedness *} -lemma add_cn:"isnorm P \ isnorm Q \ isnorm (P \ Q)" -proof(induct P Q rule: add.induct) - case (2 c i P2) thus ?case by (cases P2, simp_all, cases i, simp_all) -next - case (3 i P2 c) thus ?case by (cases P2, simp_all, cases i, simp_all) -next - case (4 c P2 i Q2) - from prems have "isnorm P2" "isnorm Q2" by (auto simp only: norm_PX1[of P2 i Q2] norm_PX2[of P2 i Q2]) - with prems show ?case by(cases i, simp, cases P2, auto, case_tac pol2, auto) -next - case (5 P2 i Q2 c) - from prems have "isnorm P2" "isnorm Q2" by (auto simp only: norm_PX1[of P2 i Q2] norm_PX2[of P2 i Q2]) - with prems show ?case by(cases i, simp, cases P2, auto, case_tac pol2, auto) -next - case (6 x P2 y Q2) - from prems have Y0:"y>0" by (cases y, auto simp add: norm_Pinj_0_False) - from prems have X0:"x>0" by (cases x, auto simp add: norm_Pinj_0_False) - have "x < y \ x = y \ x > y" by arith - moreover - { assume "xy" hence "EX d. x=d+y" by arith - then obtain d where "x=d+y".. - moreover - note prems Y0 - moreover - from prems have "isnorm P2" "isnorm Q2" by (auto simp add: norm_Pinj[of _ P2] norm_Pinj[of _ Q2]) - moreover - with prems have "isnorm (Pinj d P2)" by (cases d, simp, cases P2, auto) - ultimately have ?case by (simp add: mkPinj_cn)} - ultimately show ?case by blast -next - case (7 x P2 Q2 y R) - have "x=0 \ (x = 1) \ (x > 1)" by arith - moreover - { assume "x=0" with prems have ?case by (auto simp add: norm_Pinj_0_False)} - moreover - { assume "x=1" - from prems have "isnorm R" "isnorm P2" by (auto simp add: norm_Pinj[of _ P2] norm_PX2[of Q2 y R]) - with prems have "isnorm (R \ P2)" by simp - with prems have ?case by (simp add: norm_PXtrans[of Q2 y _]) } - moreover - { assume "x > 1" hence "EX d. x=Suc (Suc d)" by arith - then obtain d where X:"x=Suc (Suc d)" .. - from prems have NR:"isnorm R" "isnorm P2" by (auto simp add: norm_Pinj[of _ P2] norm_PX2[of Q2 y R]) - with prems have "isnorm (Pinj (x - 1) P2)" by(cases P2, auto) - with prems NR have "isnorm (R \ Pinj (x - 1) P2)" "isnorm (PX Q2 y R)" by simp fact - with X have ?case by (simp add: norm_PXtrans[of Q2 y _]) } - ultimately show ?case by blast -next - case (8 Q2 y R x P2) - have "x = 0 \ x = 1 \ x > 1" by arith - moreover - { assume "x=0" with prems have ?case by (auto simp add: norm_Pinj_0_False)} - moreover - { assume "x=1" - from prems have "isnorm R" "isnorm P2" by (auto simp add: norm_Pinj[of _ P2] norm_PX2[of Q2 y R]) - with prems have "isnorm (R \ P2)" by simp - with prems have ?case by (simp add: norm_PXtrans[of Q2 y _]) } - moreover - { assume "x > 1" hence "EX d. x=Suc (Suc d)" by arith - then obtain d where X:"x=Suc (Suc d)" .. - from prems have NR:"isnorm R" "isnorm P2" by (auto simp add: norm_Pinj[of _ P2] norm_PX2[of Q2 y R]) - with prems have "isnorm (Pinj (x - 1) P2)" by(cases P2, auto) - with prems NR have "isnorm (R \ Pinj (x - 1) P2)" "isnorm (PX Q2 y R)" by simp fact - with X have ?case by (simp add: norm_PXtrans[of Q2 y _]) } - ultimately show ?case by blast -next - case (9 P1 x P2 Q1 y Q2) - from prems have Y0:"y>0" by(cases y, auto) - from prems have X0:"x>0" by(cases x, auto) - from prems have NP1:"isnorm P1" and NP2:"isnorm P2" by (auto simp add: norm_PX1[of P1 _ P2] norm_PX2[of P1 _ P2]) - from prems have NQ1:"isnorm Q1" and NQ2:"isnorm Q2" by (auto simp add: norm_PX1[of Q1 _ Q2] norm_PX2[of Q1 _ Q2]) - have "y < x \ x = y \ x < y" by arith - moreover - {assume sm1:"y < x" hence "EX d. x=d+y" by arith - then obtain d where sm2:"x=d+y".. - note prems NQ1 NP1 NP2 NQ2 sm1 sm2 - moreover - have "isnorm (PX P1 d (Pc 0))" - proof(cases P1) - case (PX p1 y p2) - with prems show ?thesis by(cases d, simp,cases p2, auto) - next case Pc from prems show ?thesis by(cases d, auto) - next case Pinj from prems show ?thesis by(cases d, auto) - qed - ultimately have "isnorm (P2 \ Q2)" "isnorm (PX P1 (x - y) (Pc 0) \ Q1)" by auto - with Y0 sm1 sm2 have ?case by (simp add: mkPX_cn)} - moreover - {assume "x=y" - from prems NP1 NP2 NQ1 NQ2 have "isnorm (P2 \ Q2)" "isnorm (P1 \ Q1)" by auto - with Y0 prems have ?case by (simp add: mkPX_cn) } - moreover - {assume sm1:"x Q2)" "isnorm (PX Q1 (y - x) (Pc 0) \ P1)" by auto - with X0 sm1 sm2 have ?case by (simp add: mkPX_cn)} - ultimately show ?case by blast -qed simp - -text {* mul concerves normalizedness *} -lemma mul_cn :"isnorm P \ isnorm Q \ isnorm (P \ Q)" -proof(induct P Q rule: mul.induct) - case (2 c i P2) thus ?case - by (cases P2, simp_all) (cases "i",simp_all add: mkPinj_cn) -next - case (3 i P2 c) thus ?case - by (cases P2, simp_all) (cases "i",simp_all add: mkPinj_cn) -next - case (4 c P2 i Q2) - from prems have "isnorm P2" "isnorm Q2" by (auto simp only: norm_PX1[of P2 i Q2] norm_PX2[of P2 i Q2]) - with prems show ?case - by - (case_tac "c=0",simp_all,case_tac "i=0",simp_all add: mkPX_cn) -next - case (5 P2 i Q2 c) - from prems have "isnorm P2" "isnorm Q2" by (auto simp only: norm_PX1[of P2 i Q2] norm_PX2[of P2 i Q2]) - with prems show ?case - by - (case_tac "c=0",simp_all,case_tac "i=0",simp_all add: mkPX_cn) -next - case (6 x P2 y Q2) - have "x < y \ x = y \ x > y" by arith - moreover - { assume "x0" by (cases x, auto simp add: norm_Pinj_0_False) - moreover - from prems have "isnorm P2" "isnorm Q2" by (auto simp add: norm_Pinj[of _ P2] norm_Pinj[of _ Q2]) - moreover - with prems have "isnorm (Pinj d Q2)" by (cases d, simp, cases Q2, auto) - ultimately have ?case by (simp add: mkPinj_cn)} - moreover - { assume "x=y" - moreover - from prems have "isnorm P2" "isnorm Q2" by(auto simp add: norm_Pinj[of _ P2] norm_Pinj[of _ Q2]) - moreover - with prems have "y>0" by (cases y, auto simp add: norm_Pinj_0_False) - moreover - note prems - moreover - ultimately have ?case by (simp add: mkPinj_cn) } - moreover - { assume "x>y" hence "EX d. x=d+y" by arith - then obtain d where "x=d+y".. - moreover - note prems - moreover - from prems have "y>0" by (cases y, auto simp add: norm_Pinj_0_False) - moreover - from prems have "isnorm P2" "isnorm Q2" by (auto simp add: norm_Pinj[of _ P2] norm_Pinj[of _ Q2]) - moreover - with prems have "isnorm (Pinj d P2)" by (cases d, simp, cases P2, auto) - ultimately have ?case by (simp add: mkPinj_cn) } - ultimately show ?case by blast -next - case (7 x P2 Q2 y R) - from prems have Y0:"y>0" by(cases y, auto) - have "x=0 \ (x = 1) \ (x > 1)" by arith - moreover - { assume "x=0" with prems have ?case by (auto simp add: norm_Pinj_0_False)} - moreover - { assume "x=1" - from prems have "isnorm R" "isnorm P2" by (auto simp add: norm_Pinj[of _ P2] norm_PX2[of Q2 y R]) - with prems have "isnorm (R \ P2)" "isnorm Q2" by (auto simp add: norm_PX1[of Q2 y R]) - with Y0 prems have ?case by (simp add: mkPX_cn)} - moreover - { assume "x > 1" hence "EX d. x=Suc (Suc d)" by arith - then obtain d where X:"x=Suc (Suc d)" .. - from prems have NR:"isnorm R" "isnorm Q2" by (auto simp add: norm_PX2[of Q2 y R] norm_PX1[of Q2 y R]) - moreover - from prems have "isnorm (Pinj (x - 1) P2)" by(cases P2, auto) - moreover - from prems have "isnorm (Pinj x P2)" by(cases P2, auto) - moreover - note prems - ultimately have "isnorm (R \ Pinj (x - 1) P2)" "isnorm (Pinj x P2 \ Q2)" by auto - with Y0 X have ?case by (simp add: mkPX_cn)} - ultimately show ?case by blast -next - case (8 Q2 y R x P2) - from prems have Y0:"y>0" by(cases y, auto) - have "x=0 \ (x = 1) \ (x > 1)" by arith - moreover - { assume "x=0" with prems have ?case by (auto simp add: norm_Pinj_0_False)} - moreover - { assume "x=1" - from prems have "isnorm R" "isnorm P2" by (auto simp add: norm_Pinj[of _ P2] norm_PX2[of Q2 y R]) - with prems have "isnorm (R \ P2)" "isnorm Q2" by (auto simp add: norm_PX1[of Q2 y R]) - with Y0 prems have ?case by (simp add: mkPX_cn) } - moreover - { assume "x > 1" hence "EX d. x=Suc (Suc d)" by arith - then obtain d where X:"x=Suc (Suc d)" .. - from prems have NR:"isnorm R" "isnorm Q2" by (auto simp add: norm_PX2[of Q2 y R] norm_PX1[of Q2 y R]) - moreover - from prems have "isnorm (Pinj (x - 1) P2)" by(cases P2, auto) - moreover - from prems have "isnorm (Pinj x P2)" by(cases P2, auto) - moreover - note prems - ultimately have "isnorm (R \ Pinj (x - 1) P2)" "isnorm (Pinj x P2 \ Q2)" by auto - with Y0 X have ?case by (simp add: mkPX_cn) } - ultimately show ?case by blast -next - case (9 P1 x P2 Q1 y Q2) - from prems have X0:"x>0" by(cases x, auto) - from prems have Y0:"y>0" by(cases y, auto) - note prems - moreover - from prems have "isnorm P1" "isnorm P2" by (auto simp add: norm_PX1[of P1 x P2] norm_PX2[of P1 x P2]) - moreover - from prems have "isnorm Q1" "isnorm Q2" by (auto simp add: norm_PX1[of Q1 y Q2] norm_PX2[of Q1 y Q2]) - ultimately have "isnorm (P1 \ Q1)" "isnorm (P2 \ Q2)" - "isnorm (P1 \ mkPinj 1 Q2)" "isnorm (Q1 \ mkPinj 1 P2)" - by (auto simp add: mkPinj_cn) - with prems X0 Y0 have - "isnorm (mkPX (P1 \ Q1) (x + y) (P2 \ Q2))" - "isnorm (mkPX (P1 \ mkPinj (Suc 0) Q2) x (Pc 0))" - "isnorm (mkPX (Q1 \ mkPinj (Suc 0) P2) y (Pc 0))" - by (auto simp add: mkPX_cn) - thus ?case by (simp add: add_cn) -qed(simp) - -text {* neg conserves normalizedness *} -lemma neg_cn: "isnorm P \ isnorm (neg P)" -proof (induct P) - case (Pinj i P2) - from prems have "isnorm P2" by (simp add: norm_Pinj[of i P2]) - with prems show ?case by(cases P2, auto, cases i, auto) -next - case (PX P1 x P2) - from prems have "isnorm P2" "isnorm P1" by (auto simp add: norm_PX1[of P1 x P2] norm_PX2[of P1 x P2]) - with prems show ?case - proof(cases P1) - case (PX p1 y p2) - with prems show ?thesis by(cases x, auto, cases p2, auto) - next - case Pinj - with prems show ?thesis by(cases x, auto) - qed(cases x, auto) -qed(simp) - -text {* sub conserves normalizedness *} -lemma sub_cn:"isnorm p \ isnorm q \ isnorm (p \ q)" -by (simp add: sub_def add_cn neg_cn) - -text {* sqr conserves normalizizedness *} -lemma sqr_cn:"isnorm P \ isnorm (sqr P)" -proof(induct P) - case (Pinj i Q) - from prems show ?case by(cases Q, auto simp add: mkPX_cn mkPinj_cn, cases i, auto simp add: mkPX_cn mkPinj_cn) -next - case (PX P1 x P2) - from prems have "x+x~=0" "isnorm P2" "isnorm P1" by (cases x, auto simp add: norm_PX1[of P1 x P2] norm_PX2[of P1 x P2]) - with prems have - "isnorm (mkPX (Pc (1 + 1) \ P1 \ mkPinj (Suc 0) P2) x (Pc 0))" - and "isnorm (mkPX (sqr P1) (x + x) (sqr P2))" - by (auto simp add: add_cn mkPX_cn mkPinj_cn mul_cn) - thus ?case by (auto simp add: add_cn mkPX_cn mkPinj_cn mul_cn) -qed simp - -text {* pow conserves normalizedness *} -lemma pow_cn:"isnorm P \ isnorm (pow n P)" -proof (induct n arbitrary: P rule: nat_less_induct) - case (1 k) - show ?case - proof (cases "k=0") - case False - then have K2:"k div 2 < k" by (cases k, auto) - from prems have "isnorm (sqr P)" by (simp add: sqr_cn) - with prems K2 show ?thesis - by (simp add: allE[of _ "(k div 2)" _] allE[of _ "(sqr P)" _], cases k, auto simp add: mul_cn) - qed simp -qed - -end diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/ex/Numeral.thy --- a/src/HOL/ex/Numeral.thy Thu Oct 29 16:23:57 2009 +0100 +++ b/src/HOL/ex/Numeral.thy Wed Nov 04 11:40:59 2009 +0100 @@ -5,7 +5,7 @@ header {* An experimental alternative numeral representation. *} theory Numeral -imports Plain Divides +imports Main begin subsection {* The @{text num} type *} diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/ex/Predicate_Compile_Alternative_Defs.thy --- a/src/HOL/ex/Predicate_Compile_Alternative_Defs.thy Thu Oct 29 16:23:57 2009 +0100 +++ b/src/HOL/ex/Predicate_Compile_Alternative_Defs.thy Wed Nov 04 11:40:59 2009 +0100 @@ -45,7 +45,7 @@ unfolding mem_def[symmetric, of _ a2] apply auto unfolding mem_def - apply auto + apply fastsimp done qed diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/ex/Predicate_Compile_Quickcheck_ex.thy --- a/src/HOL/ex/Predicate_Compile_Quickcheck_ex.thy Thu Oct 29 16:23:57 2009 +0100 +++ b/src/HOL/ex/Predicate_Compile_Quickcheck_ex.thy Wed Nov 04 11:40:59 2009 +0100 @@ -5,6 +5,31 @@ section {* Sets *} +lemma "x \ {(1::nat)} ==> False" +quickcheck[generator=predicate_compile] +oops + +lemma "x \ {Suc 0, Suc (Suc 0)} ==> x \ Suc 0" +oops + +lemma "x \ {Suc 0, Suc (Suc 0)} ==> x = Suc 0" +quickcheck[generator=predicate_compile] +oops + +lemma "x \ {Suc 0, Suc (Suc 0)} ==> x <= Suc 0" +quickcheck[generator=predicate_compile] +oops + +section {* Numerals *} + +lemma + "x \ {1, 2, (3::nat)} ==> x < 3" +(*quickcheck[generator=predicate_compile]*) +oops +lemma + "x \ {1, 2} \ {3, 4} ==> x > 4" +(*quickcheck[generator=predicate_compile]*) +oops section {* Context Free Grammar *} @@ -32,10 +57,10 @@ | "w \ S\<^isub>2 \ b # w \ B\<^isub>2" | "\v \ B\<^isub>2; v \ B\<^isub>2\ \ a # v @ w \ B\<^isub>2" -code_pred [inductify, rpred] S\<^isub>2 . -thm S\<^isub>2.rpred_equation -thm A\<^isub>2.rpred_equation -thm B\<^isub>2.rpred_equation +code_pred [inductify, random] S\<^isub>2 . +thm S\<^isub>2.random_equation +thm A\<^isub>2.random_equation +thm B\<^isub>2.random_equation values [random] 10 "{x. S\<^isub>2 x}" diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/ex/Predicate_Compile_ex.thy --- a/src/HOL/ex/Predicate_Compile_ex.thy Thu Oct 29 16:23:57 2009 +0100 +++ b/src/HOL/ex/Predicate_Compile_ex.thy Wed Nov 04 11:40:59 2009 +0100 @@ -6,9 +6,9 @@ inductive False' :: "bool" -code_pred (mode: []) False' . +code_pred (mode : []) False' . code_pred [depth_limited] False' . -code_pred [rpred] False' . +code_pred [random] False' . inductive EmptySet :: "'a \ bool" @@ -17,7 +17,7 @@ definition EmptySet' :: "'a \ bool" where "EmptySet' = {}" -code_pred (mode: [], [1]) [inductify, show_intermediate_results] EmptySet' . +code_pred (mode: [], [1]) [inductify] EmptySet' . inductive EmptyRel :: "'a \ 'b \ bool" @@ -26,7 +26,13 @@ inductive EmptyClosure :: "('a \ 'a \ bool) \ 'a \ 'a \ bool" for r :: "'a \ 'a \ bool" -code_pred (mode: [], [1], [2], [1, 2])EmptyClosure . +code_pred + (mode: [] ==> [], [] ==> [1], [] ==> [2], [] ==> [1, 2], + [1] ==> [], [1] ==> [1], [1] ==> [2], [1] ==> [1, 2], + [2] ==> [], [2] ==> [1], [2] ==> [2], [2] ==> [1, 2], + [1, 2] ==> [], [1, 2] ==> [1], [1, 2] ==> [2], [1, 2] ==> [1, 2]) + EmptyClosure . + thm EmptyClosure.equation (* TODO: inductive package is broken! inductive False'' :: "bool" @@ -60,19 +66,99 @@ where "zerozero (0, 0)" -code_pred zerozero . -code_pred [rpred, show_compilation] zerozero . +code_pred (mode: [i], [(i, o)], [(o, i)], [o]) zerozero . +code_pred [random] zerozero . + +subsection {* Alternative Rules *} + +datatype char = C | D | E | F | G | H + +inductive is_C_or_D +where + "(x = C) \ (x = D) ==> is_C_or_D x" + +code_pred (mode: [1]) is_C_or_D . +thm is_C_or_D.equation + +inductive is_D_or_E +where + "(x = D) \ (x = E) ==> is_D_or_E x" + +lemma [code_pred_intros]: + "is_D_or_E D" +by (auto intro: is_D_or_E.intros) + +lemma [code_pred_intros]: + "is_D_or_E E" +by (auto intro: is_D_or_E.intros) + +code_pred (mode: [], [1]) is_D_or_E +proof - + case is_D_or_E + from this(1) show thesis + proof + fix x + assume x: "a1 = x" + assume "x = D \ x = E" + from this show thesis + proof + assume "x = D" from this x is_D_or_E(2) show thesis by simp + next + assume "x = E" from this x is_D_or_E(3) show thesis by simp + qed + qed +qed + +thm is_D_or_E.equation + +inductive is_F_or_G +where + "x = F \ x = G ==> is_F_or_G x" + +lemma [code_pred_intros]: + "is_F_or_G F" +by (auto intro: is_F_or_G.intros) + +lemma [code_pred_intros]: + "is_F_or_G G" +by (auto intro: is_F_or_G.intros) + +inductive is_FGH +where + "is_F_or_G x ==> is_FGH x" +| "is_FGH H" + +text {* Compilation of is_FGH requires elimination rule for is_F_or_G *} + +code_pred (mode: [], [1]) is_FGH +proof - + case is_F_or_G + from this(1) show thesis + proof + fix x + assume x: "a1 = x" + assume "x = F \ x = G" + from this show thesis + proof + assume "x = F" + from this x is_F_or_G(2) show thesis by simp + next + assume "x = G" + from this x is_F_or_G(3) show thesis by simp + qed + qed +qed subsection {* Preprocessor Inlining *} definition "equals == (op =)" -inductive zerozero' where +inductive zerozero' :: "nat * nat => bool" where "equals (x, y) (0, 0) ==> zerozero' (x, y)" code_pred (mode: [1]) zerozero' . -lemma zerozero'_eq: "zerozero' == zerozero" +lemma zerozero'_eq: "zerozero' x == zerozero x" proof - have "zerozero' = zerozero" apply (auto simp add: mem_def) @@ -81,7 +167,7 @@ apply (cases rule: zerozero.cases) apply (auto simp add: equals_def intro: zerozero'.intros) done - from this show "zerozero' == zerozero" by auto + from this show "zerozero' x == zerozero x" by auto qed declare zerozero'_eq [code_pred_inline] @@ -89,9 +175,27 @@ definition "zerozero'' x == zerozero' x" text {* if preprocessing fails, zerozero'' will not have all modes. *} -ML {* Predicate_Compile_Inline_Defs.get @{context} *} -(* TODO: *) -code_pred (mode: [1]) [inductify, show_intermediate_results] zerozero'' . + +code_pred (mode: [o], [(i, o)], [(o,i)], [i]) [inductify] zerozero'' . + +subsection {* Numerals *} + +definition + "one_or_two == {Suc 0, (Suc (Suc 0))}" + +(*code_pred [inductify] one_or_two .*) +code_pred [inductify, random] one_or_two . +(*values "{x. one_or_two x}"*) +values [random] "{x. one_or_two x}" + +definition one_or_two': + "one_or_two' == {1, (2::nat)}" + +code_pred [inductify] one_or_two' . +thm one_or_two'.equation +(* TODO: handling numerals *) +(*values "{x. one_or_two' x}"*) + subsection {* even predicate *} @@ -102,14 +206,14 @@ code_pred (mode: [], [1]) even . code_pred [depth_limited] even . -code_pred [rpred] even . +code_pred [random] even . thm odd.equation thm even.equation thm odd.depth_limited_equation thm even.depth_limited_equation -thm even.rpred_equation -thm odd.rpred_equation +thm even.random_equation +thm odd.random_equation values "{x. even 2}" values "{x. odd 2}" @@ -123,9 +227,9 @@ definition odd' where "odd' x == \ even x" -code_pred [inductify] odd' . +code_pred (mode: [1]) [inductify] odd' . code_pred [inductify, depth_limited] odd' . -code_pred [inductify, rpred] odd' . +code_pred [inductify, random] odd' . thm odd'.depth_limited_equation values [depth_limit = 2] "{x. odd' 7}" @@ -135,7 +239,7 @@ where "n mod 2 = 0 \ is_even n" -code_pred is_even . +code_pred (mode: [1]) is_even . subsection {* append predicate *} @@ -145,11 +249,11 @@ code_pred (mode: [1, 2], [3], [2, 3], [1, 3], [1, 2, 3]) append . code_pred [depth_limited] append . -code_pred [rpred] append . +code_pred [random] append . thm append.equation thm append.depth_limited_equation -thm append.rpred_equation +thm append.random_equation values "{(ys, xs). append xs ys [0, Suc 0, 2]}" values "{zs. append [0, Suc 0, 2] [17, 8] zs}" @@ -172,10 +276,19 @@ lemmas [code_pred_intros] = append2_Nil append2.intros(2) -code_pred append2 +code_pred (mode: [1, 2], [3], [2, 3], [1, 3], [1, 2, 3]) append2 proof - case append2 - from append2.cases[OF append2(1)] append2(2-3) show thesis by blast + from append2(1) show thesis + proof + fix xs + assume "a1 = []" "a2 = xs" "a3 = xs" + from this append2(2) show thesis by simp + next + fix xs ys zs x + assume "a1 = x # xs" "a2 = ys" "a3 = x # zs" "append2 xs ys zs" + from this append2(3) show thesis by fastsimp + qed qed inductive tupled_append :: "'a list \ 'a list \ 'a list \ bool" @@ -183,8 +296,8 @@ "tupled_append ([], xs, xs)" | "tupled_append (xs, ys, zs) \ tupled_append (x # xs, ys, x # zs)" -code_pred tupled_append . -code_pred [rpred] tupled_append . +code_pred (mode: [(i,i,o)], [(i,o,i)], [(o,i,i)], [(o,o,i)], [i]) tupled_append . +code_pred [random] tupled_append . thm tupled_append.equation (* TODO: values with tupled modes @@ -197,7 +310,7 @@ | "[| ys = fst (xa, y); x # zs = snd (xa, y); tupled_append' (xs, ys, zs) |] ==> tupled_append' (x # xs, xa, y)" -code_pred tupled_append' . +code_pred (mode: [(i,i,o)], [(i,o,i)], [(o,i,i)], [(o,o,i)], [i]) tupled_append' . thm tupled_append'.equation inductive tupled_append'' :: "'a list \ 'a list \ 'a list \ bool" @@ -205,9 +318,7 @@ "tupled_append'' ([], xs, xs)" | "ys = fst yszs ==> x # zs = snd yszs ==> tupled_append'' (xs, ys, zs) \ tupled_append'' (x # xs, yszs)" -thm tupled_append''.cases - -code_pred [inductify] tupled_append'' . +code_pred (mode: [(i,i,o)], [(i,o,i)], [(o,i,i)], [(o,o,i)], [i]) [inductify] tupled_append'' . thm tupled_append''.equation inductive tupled_append''' :: "'a list \ 'a list \ 'a list \ bool" @@ -215,7 +326,7 @@ "tupled_append''' ([], xs, xs)" | "yszs = (ys, zs) ==> tupled_append''' (xs, yszs) \ tupled_append''' (x # xs, ys, x # zs)" -code_pred [inductify] tupled_append''' . +code_pred (mode: [(i,i,o)], [(i,o,i)], [(o,i,i)], [(o,o,i)], [i]) [inductify] tupled_append''' . thm tupled_append'''.equation subsection {* map_ofP predicate *} @@ -237,9 +348,9 @@ | "P x ==> filter1 P xs ys ==> filter1 P (x#xs) (x#ys)" | "\ P x ==> filter1 P xs ys ==> filter1 P (x#xs) ys" -code_pred (mode: [1], [1, 2]) filter1 . +code_pred (mode: [1] ==> [1], [1] ==> [1, 2]) filter1 . code_pred [depth_limited] filter1 . -code_pred [rpred] filter1 . +code_pred [random] filter1 . thm filter1.equation @@ -251,16 +362,16 @@ code_pred (mode: [1, 2, 3], [1, 2]) filter2 . code_pred [depth_limited] filter2 . -code_pred [rpred] filter2 . +code_pred [random] filter2 . thm filter2.equation -thm filter2.rpred_equation +thm filter2.random_equation inductive filter3 for P where "List.filter P xs = ys ==> filter3 P xs ys" -code_pred filter3 . +code_pred (mode: [] ==> [1], [] ==> [1, 2], [1] ==> [1], [1] ==> [1, 2]) filter3 . code_pred [depth_limited] filter3 . thm filter3.depth_limited_equation @@ -268,9 +379,9 @@ where "List.filter P xs = ys ==> filter4 P xs ys" -code_pred filter4 . +code_pred (mode: [1, 2], [1, 2, 3]) filter4 . code_pred [depth_limited] filter4 . -code_pred [rpred] filter4 . +code_pred [random] filter4 . subsection {* reverse predicate *} @@ -288,7 +399,7 @@ "tupled_rev ([], [])" | "tupled_rev (xs, xs') \ tupled_append (xs', [x], ys) \ tupled_rev (x#xs, ys)" -code_pred tupled_rev . +code_pred (mode: [(i, o)], [(o, i)], [i]) tupled_rev . thm tupled_rev.equation subsection {* partition predicate *} @@ -299,9 +410,9 @@ | "f x \ partition f xs ys zs \ partition f (x # xs) (x # ys) zs" | "\ f x \ partition f xs ys zs \ partition f (x # xs) ys (x # zs)" -code_pred (mode: [1], [2, 3], [1, 2], [1, 3], [1, 2, 3]) partition . +code_pred (mode: [1] ==> [1], [1] ==> [2, 3], [1] ==> [1, 2], [1] ==> [1, 3], [1] ==> [1, 2, 3]) partition . code_pred [depth_limited] partition . -code_pred [rpred] partition . +code_pred [random] partition . values 10 "{(ys, zs). partition is_even [0, Suc 0, 2, 3, 4, 5, 6, 7] ys zs}" @@ -314,7 +425,7 @@ | "f x \ tupled_partition f (xs, ys, zs) \ tupled_partition f (x # xs, x # ys, zs)" | "\ f x \ tupled_partition f (xs, ys, zs) \ tupled_partition f (x # xs, ys, x # zs)" -code_pred tupled_partition . +code_pred (mode: [i] ==> [i], [i] ==> [(i, i, o)], [i] ==> [(i, o, i)], [i] ==> [(o, i, i)], [i] ==> [(i, o, o)]) tupled_partition . thm tupled_partition.equation @@ -325,25 +436,25 @@ subsection {* transitive predicate *} -code_pred tranclp +code_pred (mode: [1] ==> [1, 2], [1] ==> [1], [2] ==> [1, 2], [2] ==> [2], [] ==> [1, 2], [] ==> [1], [] ==> [2], [] ==> []) tranclp proof - case tranclp from this converse_tranclpE[OF this(1)] show thesis by metis qed code_pred [depth_limited] tranclp . -code_pred [rpred] tranclp . +code_pred [random] tranclp . thm tranclp.equation -thm tranclp.rpred_equation +thm tranclp.random_equation inductive succ :: "nat \ nat \ bool" where "succ 0 1" | "succ m n \ succ (Suc m) (Suc n)" code_pred succ . -code_pred [rpred] succ . +code_pred [random] succ . thm succ.equation -thm succ.rpred_equation +thm succ.random_equation values 10 "{(m, n). succ n m}" values "{m. succ 0 m}" @@ -456,9 +567,9 @@ subsection {* Lexicographic order *} code_pred [inductify] lexord . -code_pred [inductify, rpred] lexord . +code_pred [inductify, random] lexord . thm lexord.equation -thm lexord.rpred_equation +thm lexord.random_equation inductive less_than_nat :: "nat * nat => bool" where @@ -468,16 +579,16 @@ code_pred less_than_nat . code_pred [depth_limited] less_than_nat . -code_pred [rpred] less_than_nat . +code_pred [random] less_than_nat . inductive test_lexord :: "nat list * nat list => bool" where "lexord less_than_nat (xs, ys) ==> test_lexord (xs, ys)" -code_pred [rpred] test_lexord . +code_pred [random] test_lexord . code_pred [depth_limited] test_lexord . thm test_lexord.depth_limited_equation -thm test_lexord.rpred_equation +thm test_lexord.random_equation values "{x. test_lexord ([1, 2, 3], [1, 2, 5])}" values [depth_limit = 5] "{x. test_lexord ([1, 2, 3], [1, 2, 5])}" @@ -487,15 +598,15 @@ code_pred [inductify] lexn . thm lexn.equation -code_pred [rpred] lexn . +code_pred [random] lexn . -thm lexn.rpred_equation +thm lexn.random_equation -code_pred [inductify, show_steps] lenlex . +code_pred [inductify] lenlex . thm lenlex.equation -code_pred [inductify, rpred] lenlex . -thm lenlex.rpred_equation +code_pred [inductify, random] lenlex . +thm lenlex.random_equation code_pred [inductify] lists . thm lists.equation @@ -516,8 +627,8 @@ code_pred [inductify] avl . thm avl.equation -code_pred [rpred] avl . -thm avl.rpred_equation +code_pred [random] avl . +thm avl.random_equation fun set_of where @@ -573,12 +684,12 @@ subsection {* Inverting list functions *} -code_pred [inductify] length . -code_pred [inductify, rpred] length . +code_pred [inductify, show_intermediate_results] length . +code_pred [inductify, random] length . thm size_listP.equation -thm size_listP.rpred_equation +thm size_listP.random_equation -values [random] 20 "{xs. size_listP (xs::nat list) (5::nat)}" +(*values [random] 1 "{xs. size_listP (xs::nat list) (5::nat)}"*) code_pred [inductify] concat . code_pred [inductify] hd . @@ -600,7 +711,7 @@ code_pred [inductify] foldr . code_pred [inductify] foldl . code_pred [inductify] filter . -code_pred [inductify, rpred] filter . +code_pred [inductify, random] filter . subsection {* Context Free Grammar *} @@ -615,9 +726,9 @@ | "\v \ B\<^isub>1; v \ B\<^isub>1\ \ a # v @ w \ B\<^isub>1" code_pred [inductify] S\<^isub>1p . -code_pred [inductify, rpred] S\<^isub>1p . +code_pred [inductify, random] S\<^isub>1p . thm S\<^isub>1p.equation -thm S\<^isub>1p.rpred_equation +thm S\<^isub>1p.random_equation values [random] 5 "{x. S\<^isub>1p x}" @@ -629,10 +740,10 @@ | "w \ S\<^isub>2 \ b # w \ B\<^isub>2" | "\v \ B\<^isub>2; v \ B\<^isub>2\ \ a # v @ w \ B\<^isub>2" -code_pred [inductify, rpred] S\<^isub>2 . -thm S\<^isub>2.rpred_equation -thm A\<^isub>2.rpred_equation -thm B\<^isub>2.rpred_equation +code_pred [inductify, random] S\<^isub>2 . +thm S\<^isub>2.random_equation +thm A\<^isub>2.random_equation +thm B\<^isub>2.random_equation values [random] 10 "{x. S\<^isub>2 x}" @@ -658,6 +769,8 @@ | "w \ S\<^isub>4 \ b # w \ B\<^isub>4" | "\v \ B\<^isub>4; w \ B\<^isub>4\ \ a # v @ w \ B\<^isub>4" +code_pred (mode: [], [1]) S\<^isub>4p . + subsection {* Lambda *} datatype type = @@ -708,4 +821,10 @@ | appR [simp, intro!]: "s \\<^sub>\ t ==> u \ s \\<^sub>\ u \ t" | abs [simp, intro!]: "s \\<^sub>\ t ==> Abs T s \\<^sub>\ Abs T t" +code_pred (mode: [1, 2], [1, 2, 3]) typing . +thm typing.equation + +code_pred (mode: [1], [1, 2]) beta . +thm beta.equation + end \ No newline at end of file diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/ex/ROOT.ML --- a/src/HOL/ex/ROOT.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/HOL/ex/ROOT.ML Wed Nov 04 11:40:59 2009 +0100 @@ -45,8 +45,6 @@ "Groebner_Examples", "MT", "Unification", - "Commutative_RingEx", - "Commutative_Ring_Complete", "Primrec", "Tarski", "Adder", diff -r 2b5b0f9e271c -r 1464ddca182b src/HOL/ex/predicate_compile.ML --- a/src/HOL/ex/predicate_compile.ML Thu Oct 29 16:23:57 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,2182 +0,0 @@ -(* Author: Lukas Bulwahn, TU Muenchen - -(Prototype of) A compiler from predicates specified by intro/elim rules -to equations. -*) - -signature PREDICATE_COMPILE = -sig - type mode = int list option list * int list - (*val add_equations_of: bool -> string list -> theory -> theory *) - val register_predicate : (thm list * thm * int) -> theory -> theory - val is_registered : theory -> string -> bool - (* val fetch_pred_data : theory -> string -> (thm list * thm * int) *) - val predfun_intro_of: theory -> string -> mode -> thm - val predfun_elim_of: theory -> string -> mode -> thm - val strip_intro_concl: int -> term -> term * (term list * term list) - val predfun_name_of: theory -> string -> mode -> string - val all_preds_of : theory -> string list - val modes_of: theory -> string -> mode list - val string_of_mode : mode -> string - val intros_of: theory -> string -> thm list - val nparams_of: theory -> string -> int - val add_intro: thm -> theory -> theory - val set_elim: thm -> theory -> theory - val setup: theory -> theory - val code_pred: string -> Proof.context -> Proof.state - val code_pred_cmd: string -> Proof.context -> Proof.state - val print_stored_rules: theory -> unit - val print_all_modes: theory -> unit - val do_proofs: bool Unsynchronized.ref - val mk_casesrule : Proof.context -> int -> thm list -> term - val analyze_compr: theory -> term -> term - val eval_ref: (unit -> term Predicate.pred) option Unsynchronized.ref - val add_equations : string list -> theory -> theory - val code_pred_intros_attrib : attribute - (* used by Quickcheck_Generator *) - (*val funT_of : mode -> typ -> typ - val mk_if_pred : term -> term - val mk_Eval : term * term -> term*) - val mk_tupleT : typ list -> typ -(* val mk_predT : typ -> typ *) - (* temporary for testing of the compilation *) - datatype indprem = Prem of term list * term | Negprem of term list * term | Sidecond of term | - GeneratorPrem of term list * term | Generator of (string * typ); - val prepare_intrs: theory -> string list -> - (string * typ) list * int * string list * string list * (string * mode list) list * - (string * (term list * indprem list) list) list * (string * (int option list * int)) list - datatype compilation_funs = CompilationFuns of { - mk_predT : typ -> typ, - dest_predT : typ -> typ, - mk_bot : typ -> term, - mk_single : term -> term, - mk_bind : term * term -> term, - mk_sup : term * term -> term, - mk_if : term -> term, - mk_not : term -> term, - mk_map : typ -> typ -> term -> term -> term, - lift_pred : term -> term - }; - datatype tmode = Mode of mode * int list * tmode option list; - type moded_clause = term list * (indprem * tmode) list - type 'a pred_mode_table = (string * (mode * 'a) list) list - val infer_modes : bool -> theory -> (string * (int list option list * int list) list) list - -> (string * (int option list * int)) list -> string list - -> (string * (term list * indprem list) list) list - -> (moded_clause list) pred_mode_table - val infer_modes_with_generator : theory -> (string * (int list option list * int list) list) list - -> (string * (int option list * int)) list -> string list - -> (string * (term list * indprem list) list) list - -> (moded_clause list) pred_mode_table - (*val compile_preds : theory -> compilation_funs -> string list -> string list - -> (string * typ) list -> (moded_clause list) pred_mode_table -> term pred_mode_table - val rpred_create_definitions :(string * typ) list -> string * mode list - -> theory -> theory - val split_smode : int list -> term list -> (term list * term list) *) - val print_moded_clauses : - theory -> (moded_clause list) pred_mode_table -> unit - val print_compiled_terms : theory -> term pred_mode_table -> unit - (*val rpred_prove_preds : theory -> term pred_mode_table -> thm pred_mode_table*) - val rpred_compfuns : compilation_funs - val dest_funT : typ -> typ * typ - (* val depending_preds_of : theory -> thm list -> string list *) - val add_quickcheck_equations : string list -> theory -> theory - val add_sizelim_equations : string list -> theory -> theory - val is_inductive_predicate : theory -> string -> bool - val terms_vs : term list -> string list - val subsets : int -> int -> int list list - val check_mode_clause : bool -> theory -> string list -> - (string * mode list) list -> (string * mode list) list -> mode -> (term list * indprem list) - -> (term list * (indprem * tmode) list) option - val string_of_moded_prem : theory -> (indprem * tmode) -> string - val all_modes_of : theory -> (string * mode list) list - val all_generator_modes_of : theory -> (string * mode list) list - val compile_clause : compilation_funs -> term option -> (term list -> term) -> - theory -> string list -> string list -> mode -> term -> moded_clause -> term - val preprocess_intro : theory -> thm -> thm - val is_constrt : theory -> term -> bool - val is_predT : typ -> bool - val guess_nparams : typ -> int -end; - -structure Predicate_Compile : PREDICATE_COMPILE = -struct - -(** auxiliary **) - -(* debug stuff *) - -fun tracing s = (if ! Toplevel.debug then tracing s else ()); - -fun print_tac s = Seq.single; (* (if ! Toplevel.debug then Tactical.print_tac s else Seq.single); *) -fun debug_tac msg = Seq.single; (* (fn st => (tracing msg; Seq.single st)); *) - -val do_proofs = Unsynchronized.ref true; - -fun mycheat_tac thy i st = - (Tactic.rtac (Skip_Proof.make_thm thy (Var (("A", 0), propT))) i) st - -fun remove_last_goal thy st = - (Tactic.rtac (Skip_Proof.make_thm thy (Var (("A", 0), propT))) (nprems_of st)) st - -(* reference to preprocessing of InductiveSet package *) - -val ind_set_codegen_preproc = Inductive_Set.codegen_preproc; - -(** fundamentals **) - -(* syntactic operations *) - -fun mk_eq (x, xs) = - let fun mk_eqs _ [] = [] - | mk_eqs a (b::cs) = - HOLogic.mk_eq (Free (a, fastype_of b), b) :: mk_eqs a cs - in mk_eqs x xs end; - -fun mk_tupleT [] = HOLogic.unitT - | mk_tupleT Ts = foldr1 HOLogic.mk_prodT Ts; - -fun dest_tupleT (Type (@{type_name Product_Type.unit}, [])) = [] - | dest_tupleT (Type (@{type_name "*"}, [T1, T2])) = T1 :: (dest_tupleT T2) - | dest_tupleT t = [t] - -fun mk_tuple [] = HOLogic.unit - | mk_tuple ts = foldr1 HOLogic.mk_prod ts; - -fun dest_tuple (Const (@{const_name Product_Type.Unity}, _)) = [] - | dest_tuple (Const (@{const_name Pair}, _) $ t1 $ t2) = t1 :: (dest_tuple t2) - | dest_tuple t = [t] - -fun mk_scomp (t, u) = - let - val T = fastype_of t - val U = fastype_of u - val [A] = binder_types T - val D = body_type U - in - Const (@{const_name "scomp"}, T --> U --> A --> D) $ t $ u - end; - -fun dest_funT (Type ("fun",[S, T])) = (S, T) - | dest_funT T = raise TYPE ("dest_funT", [T], []) - -fun mk_fun_comp (t, u) = - let - val (_, B) = dest_funT (fastype_of t) - val (C, A) = dest_funT (fastype_of u) - in - Const(@{const_name "Fun.comp"}, (A --> B) --> (C --> A) --> C --> B) $ t $ u - end; - -fun dest_randomT (Type ("fun", [@{typ Random.seed}, - Type ("*", [Type ("*", [T, @{typ "unit => Code_Evaluation.term"}]) ,@{typ Random.seed}])])) = T - | dest_randomT T = raise TYPE ("dest_randomT", [T], []) - -(* destruction of intro rules *) - -(* FIXME: look for other place where this functionality was used before *) -fun strip_intro_concl nparams intro = let - val _ $ u = Logic.strip_imp_concl intro - val (pred, all_args) = strip_comb u - val (params, args) = chop nparams all_args -in (pred, (params, args)) end - -(** data structures **) - -type smode = int list; -type mode = smode option list * smode; -datatype tmode = Mode of mode * int list * tmode option list; - -fun split_smode is ts = - let - fun split_smode' _ _ [] = ([], []) - | split_smode' is i (t::ts) = (if i mem is then apfst else apsnd) (cons t) - (split_smode' is (i+1) ts) - in split_smode' is 1 ts end - -fun split_mode (iss, is) ts = - let - val (t1, t2) = chop (length iss) ts - in (t1, split_smode is t2) end - -fun string_of_mode (iss, is) = space_implode " -> " (map - (fn NONE => "X" - | SOME js => enclose "[" "]" (commas (map string_of_int js))) - (iss @ [SOME is])); - -fun string_of_tmode (Mode (predmode, termmode, param_modes)) = - "predmode: " ^ (string_of_mode predmode) ^ - (if null param_modes then "" else - "; " ^ "params: " ^ commas (map (the_default "NONE" o Option.map string_of_tmode) param_modes)) - -datatype indprem = Prem of term list * term | Negprem of term list * term | Sidecond of term | - GeneratorPrem of term list * term | Generator of (string * typ); - -type moded_clause = term list * (indprem * tmode) list -type 'a pred_mode_table = (string * (mode * 'a) list) list - -datatype predfun_data = PredfunData of { - name : string, - definition : thm, - intro : thm, - elim : thm -}; - -fun rep_predfun_data (PredfunData data) = data; -fun mk_predfun_data (name, definition, intro, elim) = - PredfunData {name = name, definition = definition, intro = intro, elim = elim} - -datatype function_data = FunctionData of { - name : string, - equation : thm option (* is not used at all? *) -}; - -fun rep_function_data (FunctionData data) = data; -fun mk_function_data (name, equation) = - FunctionData {name = name, equation = equation} - -datatype pred_data = PredData of { - intros : thm list, - elim : thm option, - nparams : int, - functions : (mode * predfun_data) list, - generators : (mode * function_data) list, - sizelim_functions : (mode * function_data) list -}; - -fun rep_pred_data (PredData data) = data; -fun mk_pred_data ((intros, elim, nparams), (functions, generators, sizelim_functions)) = - PredData {intros = intros, elim = elim, nparams = nparams, - functions = functions, generators = generators, sizelim_functions = sizelim_functions} -fun map_pred_data f (PredData {intros, elim, nparams, functions, generators, sizelim_functions}) = - mk_pred_data (f ((intros, elim, nparams), (functions, generators, sizelim_functions))) - -fun eq_option eq (NONE, NONE) = true - | eq_option eq (SOME x, SOME y) = eq (x, y) - | eq_option eq _ = false - -fun eq_pred_data (PredData d1, PredData d2) = - eq_list (Thm.eq_thm) (#intros d1, #intros d2) andalso - eq_option (Thm.eq_thm) (#elim d1, #elim d2) andalso - #nparams d1 = #nparams d2 - -structure PredData = TheoryDataFun -( - type T = pred_data Graph.T; - val empty = Graph.empty; - val copy = I; - val extend = I; - fun merge _ = Graph.merge eq_pred_data; -); - -(* queries *) - -fun lookup_pred_data thy name = - Option.map rep_pred_data (try (Graph.get_node (PredData.get thy)) name) - -fun the_pred_data thy name = case lookup_pred_data thy name - of NONE => error ("No such predicate " ^ quote name) - | SOME data => data; - -val is_registered = is_some oo lookup_pred_data - -val all_preds_of = Graph.keys o PredData.get - -val intros_of = #intros oo the_pred_data - -fun the_elim_of thy name = case #elim (the_pred_data thy name) - of NONE => error ("No elimination rule for predicate " ^ quote name) - | SOME thm => thm - -val has_elim = is_some o #elim oo the_pred_data; - -val nparams_of = #nparams oo the_pred_data - -val modes_of = (map fst) o #functions oo the_pred_data - -fun all_modes_of thy = map (fn name => (name, modes_of thy name)) (all_preds_of thy) - -val is_compiled = not o null o #functions oo the_pred_data - -fun lookup_predfun_data thy name mode = - Option.map rep_predfun_data (AList.lookup (op =) - (#functions (the_pred_data thy name)) mode) - -fun the_predfun_data thy name mode = case lookup_predfun_data thy name mode - of NONE => error ("No function defined for mode " ^ string_of_mode mode ^ " of predicate " ^ name) - | SOME data => data; - -val predfun_name_of = #name ooo the_predfun_data - -val predfun_definition_of = #definition ooo the_predfun_data - -val predfun_intro_of = #intro ooo the_predfun_data - -val predfun_elim_of = #elim ooo the_predfun_data - -fun lookup_generator_data thy name mode = - Option.map rep_function_data (AList.lookup (op =) - (#generators (the_pred_data thy name)) mode) - -fun the_generator_data thy name mode = case lookup_generator_data thy name mode - of NONE => error ("No generator defined for mode " ^ string_of_mode mode ^ " of predicate " ^ name) - | SOME data => data - -val generator_name_of = #name ooo the_generator_data - -val generator_modes_of = (map fst) o #generators oo the_pred_data - -fun all_generator_modes_of thy = - map (fn name => (name, generator_modes_of thy name)) (all_preds_of thy) - -fun lookup_sizelim_function_data thy name mode = - Option.map rep_function_data (AList.lookup (op =) - (#sizelim_functions (the_pred_data thy name)) mode) - -fun the_sizelim_function_data thy name mode = case lookup_sizelim_function_data thy name mode - of NONE => error ("No size-limited function defined for mode " ^ string_of_mode mode - ^ " of predicate " ^ name) - | SOME data => data - -val sizelim_function_name_of = #name ooo the_sizelim_function_data - -(*val generator_modes_of = (map fst) o #generators oo the_pred_data*) - -(* diagnostic display functions *) - -fun print_modes modes = tracing ("Inferred modes:\n" ^ - cat_lines (map (fn (s, ms) => s ^ ": " ^ commas (map - string_of_mode ms)) modes)); - -fun print_pred_mode_table string_of_entry thy pred_mode_table = - let - fun print_mode pred (mode, entry) = "mode : " ^ (string_of_mode mode) - ^ (string_of_entry pred mode entry) - fun print_pred (pred, modes) = - "predicate " ^ pred ^ ": " ^ cat_lines (map (print_mode pred) modes) - val _ = tracing (cat_lines (map print_pred pred_mode_table)) - in () end; - -fun string_of_moded_prem thy (Prem (ts, p), tmode) = - (Syntax.string_of_term_global thy (list_comb (p, ts))) ^ - "(" ^ (string_of_tmode tmode) ^ ")" - | string_of_moded_prem thy (GeneratorPrem (ts, p), Mode (predmode, is, _)) = - (Syntax.string_of_term_global thy (list_comb (p, ts))) ^ - "(generator_mode: " ^ (string_of_mode predmode) ^ ")" - | string_of_moded_prem thy (Generator (v, T), _) = - "Generator for " ^ v ^ " of Type " ^ (Syntax.string_of_typ_global thy T) - | string_of_moded_prem thy (Negprem (ts, p), Mode (_, is, _)) = - (Syntax.string_of_term_global thy (list_comb (p, ts))) ^ - "(negative mode: " ^ (space_implode ", " (map string_of_int is)) ^ ")" - | string_of_moded_prem thy (Sidecond t, Mode (_, is, _)) = - (Syntax.string_of_term_global thy t) ^ - "(sidecond mode: " ^ (space_implode ", " (map string_of_int is)) ^ ")" - | string_of_moded_prem _ _ = error "string_of_moded_prem: unimplemented" - -fun print_moded_clauses thy = - let - fun string_of_clause pred mode clauses = - cat_lines (map (fn (ts, prems) => (space_implode " --> " - (map (string_of_moded_prem thy) prems)) ^ " --> " ^ pred ^ " " - ^ (space_implode " " (map (Syntax.string_of_term_global thy) ts))) clauses) - in print_pred_mode_table string_of_clause thy end; - -fun print_compiled_terms thy = - print_pred_mode_table (fn _ => fn _ => Syntax.string_of_term_global thy) thy - -fun print_stored_rules thy = - let - val preds = (Graph.keys o PredData.get) thy - fun print pred () = let - val _ = writeln ("predicate: " ^ pred) - val _ = writeln ("number of parameters: " ^ string_of_int (nparams_of thy pred)) - val _ = writeln ("introrules: ") - val _ = fold (fn thm => fn u => writeln (Display.string_of_thm_global thy thm)) - (rev (intros_of thy pred)) () - in - if (has_elim thy pred) then - writeln ("elimrule: " ^ Display.string_of_thm_global thy (the_elim_of thy pred)) - else - writeln ("no elimrule defined") - end - in - fold print preds () - end; - -fun print_all_modes thy = - let - val _ = writeln ("Inferred modes:") - fun print (pred, modes) u = - let - val _ = writeln ("predicate: " ^ pred) - val _ = writeln ("modes: " ^ (commas (map string_of_mode modes))) - in u end - in - fold print (all_modes_of thy) () - end - -(** preprocessing rules **) - -fun imp_prems_conv cv ct = - case Thm.term_of ct of - Const ("==>", _) $ _ $ _ => Conv.combination_conv (Conv.arg_conv cv) (imp_prems_conv cv) ct - | _ => Conv.all_conv ct - -fun Trueprop_conv cv ct = - case Thm.term_of ct of - Const ("Trueprop", _) $ _ => Conv.arg_conv cv ct - | _ => error "Trueprop_conv" - -fun preprocess_intro thy rule = - Conv.fconv_rule - (imp_prems_conv - (Trueprop_conv (Conv.try_conv (Conv.rewr_conv (Thm.symmetric @{thm Predicate.eq_is_eq}))))) - (Thm.transfer thy rule) - -fun preprocess_elim thy nparams elimrule = - let - fun replace_eqs (Const ("Trueprop", _) $ (Const ("op =", T) $ lhs $ rhs)) = - HOLogic.mk_Trueprop (Const (@{const_name Predicate.eq}, T) $ lhs $ rhs) - | replace_eqs t = t - val prems = Thm.prems_of elimrule - val nargs = length (snd (strip_comb (HOLogic.dest_Trueprop (hd prems)))) - nparams - fun preprocess_case t = - let - val params = Logic.strip_params t - val (assums1, assums2) = chop nargs (Logic.strip_assums_hyp t) - val assums_hyp' = assums1 @ (map replace_eqs assums2) - in - list_all (params, Logic.list_implies (assums_hyp', Logic.strip_assums_concl t)) - end - val cases' = map preprocess_case (tl prems) - val elimrule' = Logic.list_implies ((hd prems) :: cases', Thm.concl_of elimrule) - in - Thm.equal_elim - (Thm.symmetric (Conv.implies_concl_conv (MetaSimplifier.rewrite true [@{thm eq_is_eq}]) - (cterm_of thy elimrule'))) - elimrule - end; - -(* special case: predicate with no introduction rule *) -fun noclause thy predname elim = let - val T = (Logic.unvarifyT o Sign.the_const_type thy) predname - val Ts = binder_types T - val names = Name.variant_list [] - (map (fn i => "x" ^ (string_of_int i)) (1 upto (length Ts))) - val vs = map2 (curry Free) names Ts - val clausehd = HOLogic.mk_Trueprop (list_comb (Const (predname, T), vs)) - val intro_t = Logic.mk_implies (@{prop False}, clausehd) - val P = HOLogic.mk_Trueprop (Free ("P", HOLogic.boolT)) - val elim_t = Logic.list_implies ([clausehd, Logic.mk_implies (@{prop False}, P)], P) - val intro = Goal.prove (ProofContext.init thy) names [] intro_t - (fn {...} => etac @{thm FalseE} 1) - val elim = Goal.prove (ProofContext.init thy) ("P" :: names) [] elim_t - (fn {...} => etac elim 1) -in - ([intro], elim) -end - -fun fetch_pred_data thy name = - case try (Inductive.the_inductive (ProofContext.init thy)) name of - SOME (info as (_, result)) => - let - fun is_intro_of intro = - let - val (const, _) = strip_comb (HOLogic.dest_Trueprop (concl_of intro)) - in (fst (dest_Const const) = name) end; - val intros = ind_set_codegen_preproc thy ((map (preprocess_intro thy)) - (filter is_intro_of (#intrs result))) - val pre_elim = nth (#elims result) (find_index (fn s => s = name) (#names (fst info))) - val nparams = length (Inductive.params_of (#raw_induct result)) - val elim = singleton (ind_set_codegen_preproc thy) (preprocess_elim thy nparams pre_elim) - val (intros, elim) = if null intros then noclause thy name elim else (intros, elim) - in - mk_pred_data ((intros, SOME elim, nparams), ([], [], [])) - end - | NONE => error ("No such predicate: " ^ quote name) - -(* updaters *) - -fun apfst3 f (x, y, z) = (f x, y, z) -fun apsnd3 f (x, y, z) = (x, f y, z) -fun aptrd3 f (x, y, z) = (x, y, f z) - -fun add_predfun name mode data = - let - val add = (apsnd o apfst3 o cons) (mode, mk_predfun_data data) - in PredData.map (Graph.map_node name (map_pred_data add)) end - -fun is_inductive_predicate thy name = - is_some (try (Inductive.the_inductive (ProofContext.init thy)) name) - -fun depending_preds_of thy (key, value) = - let - val intros = (#intros o rep_pred_data) value - in - fold Term.add_const_names (map Thm.prop_of intros) [] - |> filter (fn c => (not (c = key)) andalso (is_inductive_predicate thy c orelse is_registered thy c)) - end; - - -(* code dependency graph *) -(* -fun dependencies_of thy name = - let - val (intros, elim, nparams) = fetch_pred_data thy name - val data = mk_pred_data ((intros, SOME elim, nparams), ([], [], [])) - val keys = depending_preds_of thy intros - in - (data, keys) - end; -*) -(* guessing number of parameters *) -fun find_indexes pred xs = - let - fun find is n [] = is - | find is n (x :: xs) = find (if pred x then (n :: is) else is) (n + 1) xs; - in rev (find [] 0 xs) end; - -fun is_predT (T as Type("fun", [_, _])) = (snd (strip_type T) = HOLogic.boolT) - | is_predT _ = false - -fun guess_nparams T = - let - val argTs = binder_types T - val nparams = fold Integer.max - (map (fn x => x + 1) (find_indexes is_predT argTs)) 0 - in nparams end; - -fun add_intro thm thy = let - val (name, T) = dest_Const (fst (strip_intro_concl 0 (prop_of thm))) - fun cons_intro gr = - case try (Graph.get_node gr) name of - SOME pred_data => Graph.map_node name (map_pred_data - (apfst (fn (intro, elim, nparams) => (thm::intro, elim, nparams)))) gr - | NONE => - let - val nparams = the_default (guess_nparams T) (try (#nparams o rep_pred_data o (fetch_pred_data thy)) name) - in Graph.new_node (name, mk_pred_data (([thm], NONE, nparams), ([], [], []))) gr end; - in PredData.map cons_intro thy end - -fun set_elim thm = let - val (name, _) = dest_Const (fst - (strip_comb (HOLogic.dest_Trueprop (hd (prems_of thm))))) - fun set (intros, _, nparams) = (intros, SOME thm, nparams) - in PredData.map (Graph.map_node name (map_pred_data (apfst set))) end - -fun set_nparams name nparams = let - fun set (intros, elim, _ ) = (intros, elim, nparams) - in PredData.map (Graph.map_node name (map_pred_data (apfst set))) end - -fun register_predicate (pre_intros, pre_elim, nparams) thy = let - val (name, _) = dest_Const (fst (strip_intro_concl nparams (prop_of (hd pre_intros)))) - (* preprocessing *) - val intros = ind_set_codegen_preproc thy (map (preprocess_intro thy) pre_intros) - val elim = singleton (ind_set_codegen_preproc thy) (preprocess_elim thy nparams pre_elim) - in - PredData.map - (Graph.new_node (name, mk_pred_data ((intros, SOME elim, nparams), ([], [], [])))) thy - end - -fun set_generator_name pred mode name = - let - val set = (apsnd o apsnd3 o cons) (mode, mk_function_data (name, NONE)) - in - PredData.map (Graph.map_node pred (map_pred_data set)) - end - -fun set_sizelim_function_name pred mode name = - let - val set = (apsnd o aptrd3 o cons) (mode, mk_function_data (name, NONE)) - in - PredData.map (Graph.map_node pred (map_pred_data set)) - end - -(** data structures for generic compilation for different monads **) - -(* maybe rename functions more generic: - mk_predT -> mk_monadT; dest_predT -> dest_monadT - mk_single -> mk_return (?) -*) -datatype compilation_funs = CompilationFuns of { - mk_predT : typ -> typ, - dest_predT : typ -> typ, - mk_bot : typ -> term, - mk_single : term -> term, - mk_bind : term * term -> term, - mk_sup : term * term -> term, - mk_if : term -> term, - mk_not : term -> term, -(* funT_of : mode -> typ -> typ, *) -(* mk_fun_of : theory -> (string * typ) -> mode -> term, *) - mk_map : typ -> typ -> term -> term -> term, - lift_pred : term -> term -}; - -fun mk_predT (CompilationFuns funs) = #mk_predT funs -fun dest_predT (CompilationFuns funs) = #dest_predT funs -fun mk_bot (CompilationFuns funs) = #mk_bot funs -fun mk_single (CompilationFuns funs) = #mk_single funs -fun mk_bind (CompilationFuns funs) = #mk_bind funs -fun mk_sup (CompilationFuns funs) = #mk_sup funs -fun mk_if (CompilationFuns funs) = #mk_if funs -fun mk_not (CompilationFuns funs) = #mk_not funs -(*fun funT_of (CompilationFuns funs) = #funT_of funs*) -(*fun mk_fun_of (CompilationFuns funs) = #mk_fun_of funs*) -fun mk_map (CompilationFuns funs) = #mk_map funs -fun lift_pred (CompilationFuns funs) = #lift_pred funs - -fun funT_of compfuns (iss, is) T = - let - val Ts = binder_types T - val (paramTs, (inargTs, outargTs)) = split_mode (iss, is) Ts - val paramTs' = map2 (fn NONE => I | SOME is => funT_of compfuns ([], is)) iss paramTs - in - (paramTs' @ inargTs) ---> (mk_predT compfuns (mk_tupleT outargTs)) - end; - -fun sizelim_funT_of compfuns (iss, is) T = - let - val Ts = binder_types T - val (paramTs, (inargTs, outargTs)) = split_mode (iss, is) Ts - val paramTs' = map2 (fn SOME is => sizelim_funT_of compfuns ([], is) | NONE => I) iss paramTs - in - (paramTs' @ inargTs @ [@{typ "code_numeral"}]) ---> (mk_predT compfuns (mk_tupleT outargTs)) - end; - -fun mk_fun_of compfuns thy (name, T) mode = - Const (predfun_name_of thy name mode, funT_of compfuns mode T) - -fun mk_sizelim_fun_of compfuns thy (name, T) mode = - Const (sizelim_function_name_of thy name mode, sizelim_funT_of compfuns mode T) - -fun mk_generator_of compfuns thy (name, T) mode = - Const (generator_name_of thy name mode, sizelim_funT_of compfuns mode T) - - -structure PredicateCompFuns = -struct - -fun mk_predT T = Type (@{type_name "Predicate.pred"}, [T]) - -fun dest_predT (Type (@{type_name "Predicate.pred"}, [T])) = T - | dest_predT T = raise TYPE ("dest_predT", [T], []); - -fun mk_bot T = Const (@{const_name Orderings.bot}, mk_predT T); - -fun mk_single t = - let val T = fastype_of t - in Const(@{const_name Predicate.single}, T --> mk_predT T) $ t end; - -fun mk_bind (x, f) = - let val T as Type ("fun", [_, U]) = fastype_of f - in - Const (@{const_name Predicate.bind}, fastype_of x --> T --> U) $ x $ f - end; - -val mk_sup = HOLogic.mk_binop @{const_name sup}; - -fun mk_if cond = Const (@{const_name Predicate.if_pred}, - HOLogic.boolT --> mk_predT HOLogic.unitT) $ cond; - -fun mk_not t = let val T = mk_predT HOLogic.unitT - in Const (@{const_name Predicate.not_pred}, T --> T) $ t end - -fun mk_Enum f = - let val T as Type ("fun", [T', _]) = fastype_of f - in - Const (@{const_name Predicate.Pred}, T --> mk_predT T') $ f - end; - -fun mk_Eval (f, x) = - let - val T = fastype_of x - in - Const (@{const_name Predicate.eval}, mk_predT T --> T --> HOLogic.boolT) $ f $ x - end; - -fun mk_map T1 T2 tf tp = Const (@{const_name Predicate.map}, - (T1 --> T2) --> mk_predT T1 --> mk_predT T2) $ tf $ tp; - -val lift_pred = I - -val compfuns = CompilationFuns {mk_predT = mk_predT, dest_predT = dest_predT, mk_bot = mk_bot, - mk_single = mk_single, mk_bind = mk_bind, mk_sup = mk_sup, mk_if = mk_if, mk_not = mk_not, - mk_map = mk_map, lift_pred = lift_pred}; - -end; - -(* termify_code: -val termT = Type ("Code_Evaluation.term", []); -fun termifyT T = HOLogic.mk_prodT (T, HOLogic.unitT --> termT) -*) -(* -fun lift_random random = - let - val T = dest_randomT (fastype_of random) - in - mk_scomp (random, - mk_fun_comp (HOLogic.pair_const (PredicateCompFuns.mk_predT T) @{typ Random.seed}, - mk_fun_comp (Const (@{const_name Predicate.single}, T --> (PredicateCompFuns.mk_predT T)), - Const (@{const_name "fst"}, HOLogic.mk_prodT (T, @{typ "unit => term"}) --> T)))) - end; -*) - -structure RPredCompFuns = -struct - -fun mk_rpredT T = - @{typ "Random.seed"} --> HOLogic.mk_prodT (PredicateCompFuns.mk_predT T, @{typ "Random.seed"}) - -fun dest_rpredT (Type ("fun", [_, - Type (@{type_name "*"}, [Type (@{type_name "Predicate.pred"}, [T]), _])])) = T - | dest_rpredT T = raise TYPE ("dest_rpredT", [T], []); - -fun mk_bot T = Const(@{const_name RPred.bot}, mk_rpredT T) - -fun mk_single t = - let - val T = fastype_of t - in - Const (@{const_name RPred.return}, T --> mk_rpredT T) $ t - end; - -fun mk_bind (x, f) = - let - val T as (Type ("fun", [_, U])) = fastype_of f - in - Const (@{const_name RPred.bind}, fastype_of x --> T --> U) $ x $ f - end - -val mk_sup = HOLogic.mk_binop @{const_name RPred.supp} - -fun mk_if cond = Const (@{const_name RPred.if_rpred}, - HOLogic.boolT --> mk_rpredT HOLogic.unitT) $ cond; - -fun mk_not t = error "Negation is not defined for RPred" - -fun mk_map t = error "FIXME" (*FIXME*) - -fun lift_pred t = - let - val T = PredicateCompFuns.dest_predT (fastype_of t) - val lift_predT = PredicateCompFuns.mk_predT T --> mk_rpredT T - in - Const (@{const_name "RPred.lift_pred"}, lift_predT) $ t - end; - -val compfuns = CompilationFuns {mk_predT = mk_rpredT, dest_predT = dest_rpredT, mk_bot = mk_bot, - mk_single = mk_single, mk_bind = mk_bind, mk_sup = mk_sup, mk_if = mk_if, mk_not = mk_not, - mk_map = mk_map, lift_pred = lift_pred}; - -end; -(* for external use with interactive mode *) -val rpred_compfuns = RPredCompFuns.compfuns; - -fun lift_random random = - let - val T = dest_randomT (fastype_of random) - in - Const (@{const_name lift_random}, (@{typ Random.seed} --> - HOLogic.mk_prodT (HOLogic.mk_prodT (T, @{typ "unit => term"}), @{typ Random.seed})) --> - RPredCompFuns.mk_rpredT T) $ random - end; - -(* Mode analysis *) - -(*** check if a term contains only constructor functions ***) -fun is_constrt thy = - let - val cnstrs = flat (maps - (map (fn (_, (Tname, _, cs)) => map (apsnd (rpair Tname o length)) cs) o #descr o snd) - (Symtab.dest (Datatype.get_all thy))); - fun check t = (case strip_comb t of - (Free _, []) => true - | (Const (s, T), ts) => (case (AList.lookup (op =) cnstrs s, body_type T) of - (SOME (i, Tname), Type (Tname', _)) => length ts = i andalso Tname = Tname' andalso forall check ts - | _ => false) - | _ => false) - in check end; - -(*** check if a type is an equality type (i.e. doesn't contain fun) - FIXME this is only an approximation ***) -fun is_eqT (Type (s, Ts)) = s <> "fun" andalso forall is_eqT Ts - | is_eqT _ = true; - -fun term_vs tm = fold_aterms (fn Free (x, T) => cons x | _ => I) tm []; -val terms_vs = distinct (op =) o maps term_vs; - -(** collect all Frees in a term (with duplicates!) **) -fun term_vTs tm = - fold_aterms (fn Free xT => cons xT | _ => I) tm []; - -(*FIXME this function should not be named merge... make it local instead*) -fun merge xs [] = xs - | merge [] ys = ys - | merge (x::xs) (y::ys) = if length x >= length y then x::merge xs (y::ys) - else y::merge (x::xs) ys; - -fun subsets i j = if i <= j then - let val is = subsets (i+1) j - in merge (map (fn ks => i::ks) is) is end - else [[]]; - -(* FIXME: should be in library - map_prod *) -fun cprod ([], ys) = [] - | cprod (x :: xs, ys) = map (pair x) ys @ cprod (xs, ys); - -fun cprods xss = List.foldr (map op :: o cprod) [[]] xss; - - - -(*TODO: cleanup function and put together with modes_of_term *) -(* -fun modes_of_param default modes t = let - val (vs, t') = strip_abs t - val b = length vs - fun mk_modes name args = Option.map (maps (fn (m as (iss, is)) => - let - val (args1, args2) = - if length args < length iss then - error ("Too few arguments for inductive predicate " ^ name) - else chop (length iss) args; - val k = length args2; - val perm = map (fn i => (find_index_eq (Bound (b - i)) args2) + 1) - (1 upto b) - val partial_mode = (1 upto k) \\ perm - in - if not (partial_mode subset is) then [] else - let - val is' = - (fold_index (fn (i, j) => if j mem is then cons (i + 1) else I) perm []) - |> fold (fn i => if i > k then cons (i - k + b) else I) is - - val res = map (fn x => Mode (m, is', x)) (cprods (map - (fn (NONE, _) => [NONE] - | (SOME js, arg) => map SOME (filter - (fn Mode (_, js', _) => js=js') (modes_of_term modes arg))) - (iss ~~ args1))) - in res end - end)) (AList.lookup op = modes name) - in case strip_comb t' of - (Const (name, _), args) => the_default default (mk_modes name args) - | (Var ((name, _), _), args) => the (mk_modes name args) - | (Free (name, _), args) => the (mk_modes name args) - | _ => default end - -and -*) -fun modes_of_term modes t = - let - val ks = 1 upto length (binder_types (fastype_of t)); - val default = [Mode (([], ks), ks, [])]; - fun mk_modes name args = Option.map (maps (fn (m as (iss, is)) => - let - val (args1, args2) = - if length args < length iss then - error ("Too few arguments for inductive predicate " ^ name) - else chop (length iss) args; - val k = length args2; - val prfx = 1 upto k - in - if not (is_prefix op = prfx is) then [] else - let val is' = map (fn i => i - k) (List.drop (is, k)) - in map (fn x => Mode (m, is', x)) (cprods (map - (fn (NONE, _) => [NONE] - | (SOME js, arg) => map SOME (filter - (fn Mode (_, js', _) => js=js') (modes_of_term modes arg))) - (iss ~~ args1))) - end - end)) (AList.lookup op = modes name) - - in - case strip_comb (Envir.eta_contract t) of - (Const (name, _), args) => the_default default (mk_modes name args) - | (Var ((name, _), _), args) => the (mk_modes name args) - | (Free (name, _), args) => the (mk_modes name args) - | (Abs _, []) => error "Abs at param position" (* modes_of_param default modes t *) - | _ => default - end - -fun select_mode_prem thy modes vs ps = - find_first (is_some o snd) (ps ~~ map - (fn Prem (us, t) => find_first (fn Mode (_, is, _) => - let - val (in_ts, out_ts) = split_smode is us; - val (out_ts', in_ts') = List.partition (is_constrt thy) out_ts; - val vTs = maps term_vTs out_ts'; - val dupTs = map snd (duplicates (op =) vTs) @ - map_filter (AList.lookup (op =) vTs) vs; - in - subset (op =) (terms_vs (in_ts @ in_ts'), vs) andalso - forall (is_eqT o fastype_of) in_ts' andalso - subset (op =) (term_vs t, vs) andalso - forall is_eqT dupTs - end) - (modes_of_term modes t handle Option => - error ("Bad predicate: " ^ Syntax.string_of_term_global thy t)) - | Negprem (us, t) => find_first (fn Mode (_, is, _) => - length us = length is andalso - subset (op =) (terms_vs us, vs) andalso - subset (op =) (term_vs t, vs) - (modes_of_term modes t handle Option => - error ("Bad predicate: " ^ Syntax.string_of_term_global thy t)) - | Sidecond t => if subset (op =) (term_vs t, vs) then SOME (Mode (([], []), [], [])) - else NONE - ) ps); - -fun fold_prem f (Prem (args, _)) = fold f args - | fold_prem f (Negprem (args, _)) = fold f args - | fold_prem f (Sidecond t) = f t - -fun all_subsets [] = [[]] - | all_subsets (x::xs) = let val xss' = all_subsets xs in xss' @ (map (cons x) xss') end - -fun generator vTs v = - let - val T = the (AList.lookup (op =) vTs v) - in - (Generator (v, T), Mode (([], []), [], [])) - end; - -fun gen_prem (Prem (us, t)) = GeneratorPrem (us, t) - | gen_prem _ = error "gen_prem : invalid input for gen_prem" - -fun param_gen_prem param_vs (p as Prem (us, t as Free (v, _))) = - if member (op =) param_vs v then - GeneratorPrem (us, t) - else p - | param_gen_prem param_vs p = p - -fun check_mode_clause with_generator thy param_vs modes gen_modes (iss, is) (ts, ps) = - let - val modes' = modes @ map_filter - (fn (_, NONE) => NONE | (v, SOME js) => SOME (v, [([], js)])) - (param_vs ~~ iss); - val gen_modes' = gen_modes @ map_filter - (fn (_, NONE) => NONE | (v, SOME js) => SOME (v, [([], js)])) - (param_vs ~~ iss); - val vTs = distinct (op =) ((fold o fold_prem) Term.add_frees ps (fold Term.add_frees ts [])) - val prem_vs = distinct (op =) ((fold o fold_prem) Term.add_free_names ps []) - fun check_mode_prems acc_ps vs [] = SOME (acc_ps, vs) - | check_mode_prems acc_ps vs ps = (case select_mode_prem thy modes' vs ps of - NONE => - (if with_generator then - (case select_mode_prem thy gen_modes' vs ps of - SOME (p, SOME mode) => check_mode_prems ((gen_prem p, mode) :: acc_ps) - (case p of Prem (us, _) => union (op =) vs (terms_vs us) | _ => vs) - (filter_out (equal p) ps) - | NONE => - let - val all_generator_vs = all_subsets (prem_vs \\ vs) |> sort (int_ord o (pairself length)) - in - case (find_first (fn generator_vs => is_some - (select_mode_prem thy modes' (union (op =) vs generator_vs) ps)) all_generator_vs) of - SOME generator_vs => check_mode_prems ((map (generator vTs) generator_vs) @ acc_ps) - (union (op =) vs generator_vs) ps - | NONE => NONE - end) - else - NONE) - | SOME (p, SOME mode) => check_mode_prems ((if with_generator then param_gen_prem param_vs p else p, mode) :: acc_ps) - (case p of Prem (us, _) => union (op =) vs (terms_vs us) | _ => vs) - (filter_out (equal p) ps)) - val (in_ts, in_ts') = List.partition (is_constrt thy) (fst (split_smode is ts)); - val in_vs = terms_vs in_ts; - val concl_vs = terms_vs ts - in - if forall is_eqT (map snd (duplicates (op =) (maps term_vTs in_ts))) andalso - forall (is_eqT o fastype_of) in_ts' then - case check_mode_prems [] (union (op =) param_vs in_vs) ps of - NONE => NONE - | SOME (acc_ps, vs) => - if with_generator then - SOME (ts, (rev acc_ps) @ (map (generator vTs) (concl_vs \\ vs))) - else - if subset (op =) (concl_vs, vs) then SOME (ts, rev acc_ps) else NONE - else NONE - end; - -fun check_modes_pred with_generator thy param_vs preds modes gen_modes (p, ms) = - let val SOME rs = AList.lookup (op =) preds p - in (p, List.filter (fn m => case find_index - (is_none o check_mode_clause with_generator thy param_vs modes gen_modes m) rs of - ~1 => true - | i => (tracing ("Clause " ^ string_of_int (i + 1) ^ " of " ^ - p ^ " violates mode " ^ string_of_mode m); false)) ms) - end; - -fun get_modes_pred with_generator thy param_vs preds modes gen_modes (p, ms) = - let - val SOME rs = AList.lookup (op =) preds p - in - (p, map (fn m => - (m, map (the o check_mode_clause with_generator thy param_vs modes gen_modes m) rs)) ms) - end; - -fun fixp f (x : (string * mode list) list) = - let val y = f x - in if x = y then x else fixp f y end; - -fun modes_of_arities arities = - (map (fn (s, (ks, k)) => (s, cprod (cprods (map - (fn NONE => [NONE] - | SOME k' => map SOME (subsets 1 k')) ks), - subsets 1 k))) arities) - -fun infer_modes with_generator thy extra_modes arities param_vs preds = - let - val modes = - fixp (fn modes => - map (check_modes_pred with_generator thy param_vs preds (modes @ extra_modes) []) modes) - (modes_of_arities arities) - in - map (get_modes_pred with_generator thy param_vs preds (modes @ extra_modes) []) modes - end; - -fun remove_from rem [] = [] - | remove_from rem ((k, vs) :: xs) = - (case AList.lookup (op =) rem k of - NONE => (k, vs) - | SOME vs' => (k, vs \\ vs')) - :: remove_from rem xs - -fun infer_modes_with_generator thy extra_modes arities param_vs preds = - let - val prednames = map fst preds - val extra_modes = all_modes_of thy - val gen_modes = all_generator_modes_of thy - |> filter_out (fn (name, _) => member (op =) prednames name) - val starting_modes = remove_from extra_modes (modes_of_arities arities) - val modes = - fixp (fn modes => - map (check_modes_pred true thy param_vs preds extra_modes (gen_modes @ modes)) modes) - starting_modes - in - map (get_modes_pred true thy param_vs preds extra_modes (gen_modes @ modes)) modes - end; - -(* term construction *) - -fun mk_v (names, vs) s T = (case AList.lookup (op =) vs s of - NONE => (Free (s, T), (names, (s, [])::vs)) - | SOME xs => - let - val s' = Name.variant names s; - val v = Free (s', T) - in - (v, (s'::names, AList.update (op =) (s, v::xs) vs)) - end); - -fun distinct_v (Free (s, T)) nvs = mk_v nvs s T - | distinct_v (t $ u) nvs = - let - val (t', nvs') = distinct_v t nvs; - val (u', nvs'') = distinct_v u nvs'; - in (t' $ u', nvs'') end - | distinct_v x nvs = (x, nvs); - -fun compile_match thy compfuns eqs eqs' out_ts success_t = - let - val eqs'' = maps mk_eq eqs @ eqs' - val names = fold Term.add_free_names (success_t :: eqs'' @ out_ts) []; - val name = Name.variant names "x"; - val name' = Name.variant (name :: names) "y"; - val T = mk_tupleT (map fastype_of out_ts); - val U = fastype_of success_t; - val U' = dest_predT compfuns U; - val v = Free (name, T); - val v' = Free (name', T); - in - lambda v (fst (Datatype.make_case - (ProofContext.init thy) false [] v - [(mk_tuple out_ts, - if null eqs'' then success_t - else Const (@{const_name HOL.If}, HOLogic.boolT --> U --> U --> U) $ - foldr1 HOLogic.mk_conj eqs'' $ success_t $ - mk_bot compfuns U'), - (v', mk_bot compfuns U')])) - end; - -(*FIXME function can be removed*) -fun mk_funcomp f t = - let - val names = Term.add_free_names t []; - val Ts = binder_types (fastype_of t); - val vs = map Free - (Name.variant_list names (replicate (length Ts) "x") ~~ Ts) - in - fold_rev lambda vs (f (list_comb (t, vs))) - end; -(* -fun compile_param_ext thy compfuns modes (NONE, t) = t - | compile_param_ext thy compfuns modes (m as SOME (Mode ((iss, is'), is, ms)), t) = - let - val (vs, u) = strip_abs t - val (ivs, ovs) = split_mode is vs - val (f, args) = strip_comb u - val (params, args') = chop (length ms) args - val (inargs, outargs) = split_mode is' args' - val b = length vs - val perm = map (fn i => (find_index_eq (Bound (b - i)) args') + 1) (1 upto b) - val outp_perm = - snd (split_mode is perm) - |> map (fn i => i - length (filter (fn x => x < i) is')) - val names = [] -- TODO - val out_names = Name.variant_list names (replicate (length outargs) "x") - val f' = case f of - Const (name, T) => - if AList.defined op = modes name then - mk_predfun_of thy compfuns (name, T) (iss, is') - else error "compile param: Not an inductive predicate with correct mode" - | Free (name, T) => Free (name, param_funT_of compfuns T (SOME is')) - val outTs = dest_tupleT (dest_predT compfuns (body_type (fastype_of f'))) - val out_vs = map Free (out_names ~~ outTs) - val params' = map (compile_param thy modes) (ms ~~ params) - val f_app = list_comb (f', params' @ inargs) - val single_t = (mk_single compfuns (mk_tuple (map (fn i => nth out_vs (i - 1)) outp_perm))) - val match_t = compile_match thy compfuns [] [] out_vs single_t - in list_abs (ivs, - mk_bind compfuns (f_app, match_t)) - end - | compile_param_ext _ _ _ _ = error "compile params" -*) - -fun compile_param size thy compfuns (NONE, t) = t - | compile_param size thy compfuns (m as SOME (Mode ((iss, is'), is, ms)), t) = - let - val (f, args) = strip_comb (Envir.eta_contract t) - val (params, args') = chop (length ms) args - val params' = map (compile_param size thy compfuns) (ms ~~ params) - val mk_fun_of = case size of NONE => mk_fun_of | SOME _ => mk_sizelim_fun_of - val funT_of = case size of NONE => funT_of | SOME _ => sizelim_funT_of - val f' = - case f of - Const (name, T) => - mk_fun_of compfuns thy (name, T) (iss, is') - | Free (name, T) => Free (name, funT_of compfuns (iss, is') T) - | _ => error ("PredicateCompiler: illegal parameter term") - in list_comb (f', params' @ args') end - -fun compile_expr size thy ((Mode (mode, is, ms)), t) = - case strip_comb t of - (Const (name, T), params) => - let - val params' = map (compile_param size thy PredicateCompFuns.compfuns) (ms ~~ params) - val mk_fun_of = case size of NONE => mk_fun_of | SOME _ => mk_sizelim_fun_of - in - list_comb (mk_fun_of PredicateCompFuns.compfuns thy (name, T) mode, params') - end - | (Free (name, T), args) => - let - val funT_of = case size of NONE => funT_of | SOME _ => sizelim_funT_of - in - list_comb (Free (name, funT_of PredicateCompFuns.compfuns ([], is) T), args) - end; - -fun compile_gen_expr size thy compfuns ((Mode (mode, is, ms)), t) = - case strip_comb t of - (Const (name, T), params) => - let - val params' = map (compile_param size thy compfuns) (ms ~~ params) - in - list_comb (mk_generator_of compfuns thy (name, T) mode, params') - end - | (Free (name, T), args) => - list_comb (Free (name, sizelim_funT_of RPredCompFuns.compfuns ([], is) T), args) - -(** specific rpred functions -- move them to the correct place in this file *) - -(* uncommented termify code; causes more trouble than expected at first *) -(* -fun mk_valtermify_term (t as Const (c, T)) = HOLogic.mk_prod (t, Abs ("u", HOLogic.unitT, HOLogic.reflect_term t)) - | mk_valtermify_term (Free (x, T)) = Free (x, termifyT T) - | mk_valtermify_term (t1 $ t2) = - let - val T = fastype_of t1 - val (T1, T2) = dest_funT T - val t1' = mk_valtermify_term t1 - val t2' = mk_valtermify_term t2 - in - Const ("Code_Evaluation.valapp", termifyT T --> termifyT T1 --> termifyT T2) $ t1' $ t2' - end - | mk_valtermify_term _ = error "Not a valid term for mk_valtermify_term" -*) - -fun compile_clause compfuns size final_term thy all_vs param_vs (iss, is) inp (ts, moded_ps) = - let - fun check_constrt t (names, eqs) = - if is_constrt thy t then (t, (names, eqs)) else - let - val s = Name.variant names "x"; - val v = Free (s, fastype_of t) - in (v, (s::names, HOLogic.mk_eq (v, t)::eqs)) end; - - val (in_ts, out_ts) = split_smode is ts; - val (in_ts', (all_vs', eqs)) = - fold_map check_constrt in_ts (all_vs, []); - - fun compile_prems out_ts' vs names [] = - let - val (out_ts'', (names', eqs')) = - fold_map check_constrt out_ts' (names, []); - val (out_ts''', (names'', constr_vs)) = fold_map distinct_v - out_ts'' (names', map (rpair []) vs); - in - (* termify code: - compile_match thy compfuns constr_vs (eqs @ eqs') out_ts''' - (mk_single compfuns (mk_tuple (map mk_valtermify_term out_ts))) - *) - compile_match thy compfuns constr_vs (eqs @ eqs') out_ts''' - (final_term out_ts) - end - | compile_prems out_ts vs names ((p, mode as Mode ((_, is), _, _)) :: ps) = - let - val vs' = distinct (op =) (flat (vs :: map term_vs out_ts)); - val (out_ts', (names', eqs)) = - fold_map check_constrt out_ts (names, []) - val (out_ts'', (names'', constr_vs')) = fold_map distinct_v - out_ts' ((names', map (rpair []) vs)) - val (compiled_clause, rest) = case p of - Prem (us, t) => - let - val (in_ts, out_ts''') = split_smode is us; - val args = case size of - NONE => in_ts - | SOME size_t => in_ts @ [size_t] - val u = lift_pred compfuns - (list_comb (compile_expr size thy (mode, t), args)) - val rest = compile_prems out_ts''' vs' names'' ps - in - (u, rest) - end - | Negprem (us, t) => - let - val (in_ts, out_ts''') = split_smode is us - val u = lift_pred compfuns - (mk_not PredicateCompFuns.compfuns (list_comb (compile_expr NONE thy (mode, t), in_ts))) - val rest = compile_prems out_ts''' vs' names'' ps - in - (u, rest) - end - | Sidecond t => - let - val rest = compile_prems [] vs' names'' ps; - in - (mk_if compfuns t, rest) - end - | GeneratorPrem (us, t) => - let - val (in_ts, out_ts''') = split_smode is us; - val args = case size of - NONE => in_ts - | SOME size_t => in_ts @ [size_t] - val u = list_comb (compile_gen_expr size thy compfuns (mode, t), args) - val rest = compile_prems out_ts''' vs' names'' ps - in - (u, rest) - end - | Generator (v, T) => - let - val u = lift_random (HOLogic.mk_random T @{term "1::code_numeral"}) - val rest = compile_prems [Free (v, T)] vs' names'' ps; - in - (u, rest) - end - in - compile_match thy compfuns constr_vs' eqs out_ts'' - (mk_bind compfuns (compiled_clause, rest)) - end - val prem_t = compile_prems in_ts' param_vs all_vs' moded_ps; - in - mk_bind compfuns (mk_single compfuns inp, prem_t) - end - -fun compile_pred compfuns mk_fun_of use_size thy all_vs param_vs s T mode moded_cls = - let - val (Ts1, (Us1, Us2)) = split_mode mode (binder_types T) - val funT_of = if use_size then sizelim_funT_of else funT_of - val Ts1' = map2 (fn NONE => I | SOME is => funT_of compfuns ([], is)) (fst mode) Ts1 - val xnames = Name.variant_list (all_vs @ param_vs) - (map (fn i => "x" ^ string_of_int i) (snd mode)); - val size_name = Name.variant (all_vs @ param_vs @ xnames) "size" - (* termify code: val xs = map2 (fn s => fn T => Free (s, termifyT T)) xnames Us1; *) - val xs = map2 (fn s => fn T => Free (s, T)) xnames Us1; - val params = map2 (fn s => fn T => Free (s, T)) param_vs Ts1' - val size = Free (size_name, @{typ "code_numeral"}) - val decr_size = - if use_size then - SOME (Const ("HOL.minus_class.minus", @{typ "code_numeral => code_numeral => code_numeral"}) - $ size $ Const ("HOL.one_class.one", @{typ "Code_Numeral.code_numeral"})) - else - NONE - val cl_ts = - map (compile_clause compfuns decr_size (fn out_ts => mk_single compfuns (mk_tuple out_ts)) - thy all_vs param_vs mode (mk_tuple xs)) moded_cls; - val t = foldr1 (mk_sup compfuns) cl_ts - val T' = mk_predT compfuns (mk_tupleT Us2) - val size_t = Const (@{const_name "If"}, @{typ bool} --> T' --> T' --> T') - $ HOLogic.mk_eq (size, @{term "0 :: code_numeral"}) - $ mk_bot compfuns (dest_predT compfuns T') $ t - val fun_const = mk_fun_of compfuns thy (s, T) mode - val eq = if use_size then - (list_comb (fun_const, params @ xs @ [size]), size_t) - else - (list_comb (fun_const, params @ xs), t) - in - HOLogic.mk_Trueprop (HOLogic.mk_eq eq) - end; - -(* special setup for simpset *) -val HOL_basic_ss' = HOL_basic_ss setSolver - (mk_solver "all_tac_solver" (fn _ => fn _ => all_tac)) - -(* Definition of executable functions and their intro and elim rules *) - -fun print_arities arities = tracing ("Arities:\n" ^ - cat_lines (map (fn (s, (ks, k)) => s ^ ": " ^ - space_implode " -> " (map - (fn NONE => "X" | SOME k' => string_of_int k') - (ks @ [SOME k]))) arities)); - -fun mk_Eval_of ((x, T), NONE) names = (x, names) - | mk_Eval_of ((x, T), SOME mode) names = let - val Ts = binder_types T - val argnames = Name.variant_list names - (map (fn i => "x" ^ string_of_int i) (1 upto (length Ts))); - val args = map Free (argnames ~~ Ts) - val (inargs, outargs) = split_smode mode args - val r = PredicateCompFuns.mk_Eval (list_comb (x, inargs), mk_tuple outargs) - val t = fold_rev lambda args r -in - (t, argnames @ names) -end; - -fun create_intro_elim_rule (mode as (iss, is)) defthm mode_id funT pred thy = -let - val Ts = binder_types (fastype_of pred) - val funtrm = Const (mode_id, funT) - val argnames = Name.variant_list [] - (map (fn i => "x" ^ string_of_int i) (1 upto (length Ts))); - val (Ts1, Ts2) = chop (length iss) Ts; - val Ts1' = map2 (fn NONE => I | SOME is => funT_of (PredicateCompFuns.compfuns) ([], is)) iss Ts1 - val args = map Free (argnames ~~ (Ts1' @ Ts2)) - val (params, ioargs) = chop (length iss) args - val (inargs, outargs) = split_smode is ioargs - val param_names = Name.variant_list argnames - (map (fn i => "p" ^ string_of_int i) (1 upto (length iss))) - val param_vs = map Free (param_names ~~ Ts1) - val (params', names) = fold_map mk_Eval_of ((params ~~ Ts1) ~~ iss) [] - val predpropI = HOLogic.mk_Trueprop (list_comb (pred, param_vs @ ioargs)) - val predpropE = HOLogic.mk_Trueprop (list_comb (pred, params' @ ioargs)) - val param_eqs = map (HOLogic.mk_Trueprop o HOLogic.mk_eq) (param_vs ~~ params') - val funargs = params @ inargs - val funpropE = HOLogic.mk_Trueprop (PredicateCompFuns.mk_Eval (list_comb (funtrm, funargs), - if null outargs then Free("y", HOLogic.unitT) else mk_tuple outargs)) - val funpropI = HOLogic.mk_Trueprop (PredicateCompFuns.mk_Eval (list_comb (funtrm, funargs), - mk_tuple outargs)) - val introtrm = Logic.list_implies (predpropI :: param_eqs, funpropI) - val simprules = [defthm, @{thm eval_pred}, - @{thm "split_beta"}, @{thm "fst_conv"}, @{thm "snd_conv"}] - val unfolddef_tac = Simplifier.asm_full_simp_tac (HOL_basic_ss addsimps simprules) 1 - val introthm = Goal.prove (ProofContext.init thy) (argnames @ param_names @ ["y"]) [] introtrm (fn {...} => unfolddef_tac) - val P = HOLogic.mk_Trueprop (Free ("P", HOLogic.boolT)); - val elimtrm = Logic.list_implies ([funpropE, Logic.mk_implies (predpropE, P)], P) - val elimthm = Goal.prove (ProofContext.init thy) (argnames @ param_names @ ["y", "P"]) [] elimtrm (fn {...} => unfolddef_tac) -in - (introthm, elimthm) -end; - -fun create_constname_of_mode thy prefix name mode = - let - fun string_of_mode mode = if null mode then "0" - else space_implode "_" (map string_of_int mode) - val HOmode = space_implode "_and_" - (fold (fn NONE => I | SOME mode => cons (string_of_mode mode)) (fst mode) []) - in - (Sign.full_bname thy (prefix ^ (Long_Name.base_name name))) ^ - (if HOmode = "" then "_" else "_for_" ^ HOmode ^ "_yields_") ^ (string_of_mode (snd mode)) - end; - -fun create_definitions preds (name, modes) thy = - let - val compfuns = PredicateCompFuns.compfuns - val T = AList.lookup (op =) preds name |> the - fun create_definition (mode as (iss, is)) thy = let - val mode_cname = create_constname_of_mode thy "" name mode - val mode_cbasename = Long_Name.base_name mode_cname - val Ts = binder_types T - val (Ts1, Ts2) = chop (length iss) Ts - val (Us1, Us2) = split_smode is Ts2 - val Ts1' = map2 (fn NONE => I | SOME is => funT_of compfuns ([], is)) iss Ts1 - val funT = (Ts1' @ Us1) ---> (mk_predT compfuns (mk_tupleT Us2)) - val names = Name.variant_list [] - (map (fn i => "x" ^ string_of_int i) (1 upto (length Ts))); - val xs = map Free (names ~~ (Ts1' @ Ts2)); - val (xparams, xargs) = chop (length iss) xs; - val (xins, xouts) = split_smode is xargs - val (xparams', names') = fold_map mk_Eval_of ((xparams ~~ Ts1) ~~ iss) names - fun mk_split_lambda [] t = lambda (Free (Name.variant names' "x", HOLogic.unitT)) t - | mk_split_lambda [x] t = lambda x t - | mk_split_lambda xs t = - let - fun mk_split_lambda' (x::y::[]) t = HOLogic.mk_split (lambda x (lambda y t)) - | mk_split_lambda' (x::xs) t = HOLogic.mk_split (lambda x (mk_split_lambda' xs t)) - in - mk_split_lambda' xs t - end; - val predterm = PredicateCompFuns.mk_Enum (mk_split_lambda xouts - (list_comb (Const (name, T), xparams' @ xargs))) - val lhs = list_comb (Const (mode_cname, funT), xparams @ xins) - val def = Logic.mk_equals (lhs, predterm) - val ([definition], thy') = thy |> - Sign.add_consts_i [(Binding.name mode_cbasename, funT, NoSyn)] |> - PureThy.add_defs false [((Binding.name (mode_cbasename ^ "_def"), def), [])] - val (intro, elim) = - create_intro_elim_rule mode definition mode_cname funT (Const (name, T)) thy' - in thy' |> add_predfun name mode (mode_cname, definition, intro, elim) - |> PureThy.store_thm (Binding.name (mode_cbasename ^ "I"), intro) |> snd - |> PureThy.store_thm (Binding.name (mode_cbasename ^ "E"), elim) |> snd - |> Theory.checkpoint - end; - in - fold create_definition modes thy - end; - -fun sizelim_create_definitions preds (name, modes) thy = - let - val T = AList.lookup (op =) preds name |> the - fun create_definition mode thy = - let - val mode_cname = create_constname_of_mode thy "sizelim_" name mode - val funT = sizelim_funT_of PredicateCompFuns.compfuns mode T - in - thy |> Sign.add_consts_i [(Binding.name (Long_Name.base_name mode_cname), funT, NoSyn)] - |> set_sizelim_function_name name mode mode_cname - end; - in - fold create_definition modes thy - end; - -fun rpred_create_definitions preds (name, modes) thy = - let - val T = AList.lookup (op =) preds name |> the - fun create_definition mode thy = - let - val mode_cname = create_constname_of_mode thy "gen_" name mode - val funT = sizelim_funT_of RPredCompFuns.compfuns mode T - in - thy |> Sign.add_consts_i [(Binding.name (Long_Name.base_name mode_cname), funT, NoSyn)] - |> set_generator_name name mode mode_cname - end; - in - fold create_definition modes thy - end; - -(* Proving equivalence of term *) - -fun is_Type (Type _) = true - | is_Type _ = false - -(* returns true if t is an application of an datatype constructor *) -(* which then consequently would be splitted *) -(* else false *) -fun is_constructor thy t = - if (is_Type (fastype_of t)) then - (case Datatype.get_info thy ((fst o dest_Type o fastype_of) t) of - NONE => false - | SOME info => (let - val constr_consts = maps (fn (_, (_, _, constrs)) => map fst constrs) (#descr info) - val (c, _) = strip_comb t - in (case c of - Const (name, _) => name mem_string constr_consts - | _ => false) end)) - else false - -(* MAJOR FIXME: prove_params should be simple - - different form of introrule for parameters ? *) -fun prove_param thy (NONE, t) = TRY (rtac @{thm refl} 1) - | prove_param thy (m as SOME (Mode (mode, is, ms)), t) = - let - val (f, args) = strip_comb (Envir.eta_contract t) - val (params, _) = chop (length ms) args - val f_tac = case f of - Const (name, T) => simp_tac (HOL_basic_ss addsimps - (@{thm eval_pred}::(predfun_definition_of thy name mode):: - @{thm "Product_Type.split_conv"}::[])) 1 - | Free _ => TRY (rtac @{thm refl} 1) - | Abs _ => error "prove_param: No valid parameter term" - in - REPEAT_DETERM (etac @{thm thin_rl} 1) - THEN REPEAT_DETERM (rtac @{thm ext} 1) - THEN print_tac "prove_param" - THEN f_tac - THEN print_tac "after simplification in prove_args" - THEN (EVERY (map (prove_param thy) (ms ~~ params))) - THEN (REPEAT_DETERM (atac 1)) - end - -fun prove_expr thy (Mode (mode, is, ms), t, us) (premposition : int) = - case strip_comb t of - (Const (name, T), args) => - let - val introrule = predfun_intro_of thy name mode - val (args1, args2) = chop (length ms) args - in - rtac @{thm bindI} 1 - THEN print_tac "before intro rule:" - (* for the right assumption in first position *) - THEN rotate_tac premposition 1 - THEN debug_tac (Display.string_of_thm (ProofContext.init thy) introrule) - THEN rtac introrule 1 - THEN print_tac "after intro rule" - (* work with parameter arguments *) - THEN (atac 1) - THEN (print_tac "parameter goal") - THEN (EVERY (map (prove_param thy) (ms ~~ args1))) - THEN (REPEAT_DETERM (atac 1)) - end - | _ => rtac @{thm bindI} 1 THEN atac 1 - -fun SOLVED tac st = FILTER (fn st' => nprems_of st' = nprems_of st - 1) tac st; - -fun SOLVEDALL tac st = FILTER (fn st' => nprems_of st' = 0) tac st - -fun prove_match thy (out_ts : term list) = let - fun get_case_rewrite t = - if (is_constructor thy t) then let - val case_rewrites = (#case_rewrites (Datatype.the_info thy - ((fst o dest_Type o fastype_of) t))) - in case_rewrites @ maps get_case_rewrite (snd (strip_comb t)) end - else [] - val simprules = @{thm "unit.cases"} :: @{thm "prod.cases"} :: maps get_case_rewrite out_ts -(* replace TRY by determining if it necessary - are there equations when calling compile match? *) -in - (* make this simpset better! *) - asm_simp_tac (HOL_basic_ss' addsimps simprules) 1 - THEN print_tac "after prove_match:" - THEN (DETERM (TRY (EqSubst.eqsubst_tac (ProofContext.init thy) [0] [@{thm "HOL.if_P"}] 1 - THEN (REPEAT_DETERM (rtac @{thm conjI} 1 THEN (SOLVED (asm_simp_tac HOL_basic_ss 1)))) - THEN (SOLVED (asm_simp_tac HOL_basic_ss 1))))) - THEN print_tac "after if simplification" -end; - -(* corresponds to compile_fun -- maybe call that also compile_sidecond? *) - -fun prove_sidecond thy modes t = - let - fun preds_of t nameTs = case strip_comb t of - (f as Const (name, T), args) => - if AList.defined (op =) modes name then (name, T) :: nameTs - else fold preds_of args nameTs - | _ => nameTs - val preds = preds_of t [] - val defs = map - (fn (pred, T) => predfun_definition_of thy pred ([], (1 upto (length (binder_types T))))) - preds - in - (* remove not_False_eq_True when simpset in prove_match is better *) - simp_tac (HOL_basic_ss addsimps @{thm not_False_eq_True} :: @{thm eval_pred} :: defs) 1 - (* need better control here! *) - end - -fun prove_clause thy nargs modes (iss, is) (_, clauses) (ts, moded_ps) = - let - val (in_ts, clause_out_ts) = split_smode is ts; - fun prove_prems out_ts [] = - (prove_match thy out_ts) - THEN asm_simp_tac HOL_basic_ss' 1 - THEN (rtac (if null clause_out_ts then @{thm singleI_unit} else @{thm singleI}) 1) - | prove_prems out_ts ((p, mode as Mode ((iss, is), _, param_modes)) :: ps) = - let - val premposition = (find_index (equal p) clauses) + nargs - val rest_tac = (case p of Prem (us, t) => - let - val (_, out_ts''') = split_smode is us - val rec_tac = prove_prems out_ts''' ps - in - print_tac "before clause:" - THEN asm_simp_tac HOL_basic_ss 1 - THEN print_tac "before prove_expr:" - THEN prove_expr thy (mode, t, us) premposition - THEN print_tac "after prove_expr:" - THEN rec_tac - end - | Negprem (us, t) => - let - val (_, out_ts''') = split_smode is us - val rec_tac = prove_prems out_ts''' ps - val name = (case strip_comb t of (Const (c, _), _) => SOME c | _ => NONE) - val (_, params) = strip_comb t - in - rtac @{thm bindI} 1 - THEN (if (is_some name) then - simp_tac (HOL_basic_ss addsimps [predfun_definition_of thy (the name) (iss, is)]) 1 - THEN rtac @{thm not_predI} 1 - THEN simp_tac (HOL_basic_ss addsimps [@{thm not_False_eq_True}]) 1 - THEN (REPEAT_DETERM (atac 1)) - (* FIXME: work with parameter arguments *) - THEN (EVERY (map (prove_param thy) (param_modes ~~ params))) - else - rtac @{thm not_predI'} 1) - THEN simp_tac (HOL_basic_ss addsimps [@{thm not_False_eq_True}]) 1 - THEN rec_tac - end - | Sidecond t => - rtac @{thm bindI} 1 - THEN rtac @{thm if_predI} 1 - THEN print_tac "before sidecond:" - THEN prove_sidecond thy modes t - THEN print_tac "after sidecond:" - THEN prove_prems [] ps) - in (prove_match thy out_ts) - THEN rest_tac - end; - val prems_tac = prove_prems in_ts moded_ps - in - rtac @{thm bindI} 1 - THEN rtac @{thm singleI} 1 - THEN prems_tac - end; - -fun select_sup 1 1 = [] - | select_sup _ 1 = [rtac @{thm supI1}] - | select_sup n i = (rtac @{thm supI2})::(select_sup (n - 1) (i - 1)); - -fun prove_one_direction thy clauses preds modes pred mode moded_clauses = - let - val T = the (AList.lookup (op =) preds pred) - val nargs = length (binder_types T) - nparams_of thy pred - val pred_case_rule = the_elim_of thy pred - in - REPEAT_DETERM (CHANGED (rewtac @{thm "split_paired_all"})) - THEN etac (predfun_elim_of thy pred mode) 1 - THEN etac pred_case_rule 1 - THEN (EVERY (map - (fn i => EVERY' (select_sup (length moded_clauses) i) i) - (1 upto (length moded_clauses)))) - THEN (EVERY (map2 (prove_clause thy nargs modes mode) clauses moded_clauses)) - THEN print_tac "proved one direction" - end; - -(** Proof in the other direction **) - -fun prove_match2 thy out_ts = let - fun split_term_tac (Free _) = all_tac - | split_term_tac t = - if (is_constructor thy t) then let - val info = Datatype.the_info thy ((fst o dest_Type o fastype_of) t) - val num_of_constrs = length (#case_rewrites info) - (* special treatment of pairs -- because of fishing *) - val split_rules = case (fst o dest_Type o fastype_of) t of - "*" => [@{thm prod.split_asm}] - | _ => PureThy.get_thms thy (((fst o dest_Type o fastype_of) t) ^ ".split_asm") - val (_, ts) = strip_comb t - in - (Splitter.split_asm_tac split_rules 1) -(* THEN (Simplifier.asm_full_simp_tac HOL_basic_ss 1) - THEN (DETERM (TRY (etac @{thm Pair_inject} 1))) *) - THEN (REPEAT_DETERM_N (num_of_constrs - 1) (etac @{thm botE} 1 ORELSE etac @{thm botE} 2)) - THEN (EVERY (map split_term_tac ts)) - end - else all_tac - in - split_term_tac (mk_tuple out_ts) - THEN (DETERM (TRY ((Splitter.split_asm_tac [@{thm "split_if_asm"}] 1) THEN (etac @{thm botE} 2)))) - end - -(* VERY LARGE SIMILIRATIY to function prove_param --- join both functions -*) -(* TODO: remove function *) - -fun prove_param2 thy (NONE, t) = all_tac - | prove_param2 thy (m as SOME (Mode (mode, is, ms)), t) = let - val (f, args) = strip_comb (Envir.eta_contract t) - val (params, _) = chop (length ms) args - val f_tac = case f of - Const (name, T) => full_simp_tac (HOL_basic_ss addsimps - (@{thm eval_pred}::(predfun_definition_of thy name mode) - :: @{thm "Product_Type.split_conv"}::[])) 1 - | Free _ => all_tac - | _ => error "prove_param2: illegal parameter term" - in - print_tac "before simplification in prove_args:" - THEN f_tac - THEN print_tac "after simplification in prove_args" - THEN (EVERY (map (prove_param2 thy) (ms ~~ params))) - end - - -fun prove_expr2 thy (Mode (mode, is, ms), t) = - (case strip_comb t of - (Const (name, T), args) => - etac @{thm bindE} 1 - THEN (REPEAT_DETERM (CHANGED (rewtac @{thm "split_paired_all"}))) - THEN print_tac "prove_expr2-before" - THEN (debug_tac (Syntax.string_of_term_global thy - (prop_of (predfun_elim_of thy name mode)))) - THEN (etac (predfun_elim_of thy name mode) 1) - THEN print_tac "prove_expr2" - THEN (EVERY (map (prove_param2 thy) (ms ~~ args))) - THEN print_tac "finished prove_expr2" - | _ => etac @{thm bindE} 1) - -(* FIXME: what is this for? *) -(* replace defined by has_mode thy pred *) -(* TODO: rewrite function *) -fun prove_sidecond2 thy modes t = let - fun preds_of t nameTs = case strip_comb t of - (f as Const (name, T), args) => - if AList.defined (op =) modes name then (name, T) :: nameTs - else fold preds_of args nameTs - | _ => nameTs - val preds = preds_of t [] - val defs = map - (fn (pred, T) => predfun_definition_of thy pred ([], (1 upto (length (binder_types T))))) - preds - in - (* only simplify the one assumption *) - full_simp_tac (HOL_basic_ss' addsimps @{thm eval_pred} :: defs) 1 - (* need better control here! *) - THEN print_tac "after sidecond2 simplification" - end - -fun prove_clause2 thy modes pred (iss, is) (ts, ps) i = - let - val pred_intro_rule = nth (intros_of thy pred) (i - 1) - val (in_ts, clause_out_ts) = split_smode is ts; - fun prove_prems2 out_ts [] = - print_tac "before prove_match2 - last call:" - THEN prove_match2 thy out_ts - THEN print_tac "after prove_match2 - last call:" - THEN (etac @{thm singleE} 1) - THEN (REPEAT_DETERM (etac @{thm Pair_inject} 1)) - THEN (asm_full_simp_tac HOL_basic_ss' 1) - THEN (REPEAT_DETERM (etac @{thm Pair_inject} 1)) - THEN (asm_full_simp_tac HOL_basic_ss' 1) - THEN SOLVED (print_tac "state before applying intro rule:" - THEN (rtac pred_intro_rule 1) - (* How to handle equality correctly? *) - THEN (print_tac "state before assumption matching") - THEN (REPEAT (atac 1 ORELSE - (CHANGED (asm_full_simp_tac HOL_basic_ss' 1) - THEN print_tac "state after simp_tac:")))) - | prove_prems2 out_ts ((p, mode as Mode ((iss, is), _, param_modes)) :: ps) = - let - val rest_tac = (case p of - Prem (us, t) => - let - val (_, out_ts''') = split_smode is us - val rec_tac = prove_prems2 out_ts''' ps - in - (prove_expr2 thy (mode, t)) THEN rec_tac - end - | Negprem (us, t) => - let - val (_, out_ts''') = split_smode is us - val rec_tac = prove_prems2 out_ts''' ps - val name = (case strip_comb t of (Const (c, _), _) => SOME c | _ => NONE) - val (_, params) = strip_comb t - in - print_tac "before neg prem 2" - THEN etac @{thm bindE} 1 - THEN (if is_some name then - full_simp_tac (HOL_basic_ss addsimps [predfun_definition_of thy (the name) (iss, is)]) 1 - THEN etac @{thm not_predE} 1 - THEN simp_tac (HOL_basic_ss addsimps [@{thm not_False_eq_True}]) 1 - THEN (EVERY (map (prove_param2 thy) (param_modes ~~ params))) - else - etac @{thm not_predE'} 1) - THEN rec_tac - end - | Sidecond t => - etac @{thm bindE} 1 - THEN etac @{thm if_predE} 1 - THEN prove_sidecond2 thy modes t - THEN prove_prems2 [] ps) - in print_tac "before prove_match2:" - THEN prove_match2 thy out_ts - THEN print_tac "after prove_match2:" - THEN rest_tac - end; - val prems_tac = prove_prems2 in_ts ps - in - print_tac "starting prove_clause2" - THEN etac @{thm bindE} 1 - THEN (etac @{thm singleE'} 1) - THEN (TRY (etac @{thm Pair_inject} 1)) - THEN print_tac "after singleE':" - THEN prems_tac - end; - -fun prove_other_direction thy modes pred mode moded_clauses = - let - fun prove_clause clause i = - (if i < length moded_clauses then etac @{thm supE} 1 else all_tac) - THEN (prove_clause2 thy modes pred mode clause i) - in - (DETERM (TRY (rtac @{thm unit.induct} 1))) - THEN (REPEAT_DETERM (CHANGED (rewtac @{thm split_paired_all}))) - THEN (rtac (predfun_intro_of thy pred mode) 1) - THEN (REPEAT_DETERM (rtac @{thm refl} 2)) - THEN (EVERY (map2 prove_clause moded_clauses (1 upto (length moded_clauses)))) - end; - -(** proof procedure **) - -fun prove_pred thy clauses preds modes pred mode (moded_clauses, compiled_term) = - let - val ctxt = ProofContext.init thy - val clauses = the (AList.lookup (op =) clauses pred) - in - Goal.prove ctxt (Term.add_free_names compiled_term []) [] compiled_term - (if !do_proofs then - (fn _ => - rtac @{thm pred_iffI} 1 - THEN prove_one_direction thy clauses preds modes pred mode moded_clauses - THEN print_tac "proved one direction" - THEN prove_other_direction thy modes pred mode moded_clauses - THEN print_tac "proved other direction") - else (fn _ => mycheat_tac thy 1)) - end; - -(* composition of mode inference, definition, compilation and proof *) - -(** auxillary combinators for table of preds and modes **) - -fun map_preds_modes f preds_modes_table = - map (fn (pred, modes) => - (pred, map (fn (mode, value) => (mode, f pred mode value)) modes)) preds_modes_table - -fun join_preds_modes table1 table2 = - map_preds_modes (fn pred => fn mode => fn value => - (value, the (AList.lookup (op =) (the (AList.lookup (op =) table2 pred)) mode))) table1 - -fun maps_modes preds_modes_table = - map (fn (pred, modes) => - (pred, map (fn (mode, value) => value) modes)) preds_modes_table - -fun compile_preds compfuns mk_fun_of use_size thy all_vs param_vs preds moded_clauses = - map_preds_modes (fn pred => compile_pred compfuns mk_fun_of use_size thy all_vs param_vs pred - (the (AList.lookup (op =) preds pred))) moded_clauses - -fun prove thy clauses preds modes moded_clauses compiled_terms = - map_preds_modes (prove_pred thy clauses preds modes) - (join_preds_modes moded_clauses compiled_terms) - -fun prove_by_skip thy _ _ _ _ compiled_terms = - map_preds_modes (fn pred => fn mode => fn t => Drule.standard (Skip_Proof.make_thm thy t)) - compiled_terms - -fun prepare_intrs thy prednames = - let - val intrs = maps (intros_of thy) prednames - |> map (Logic.unvarify o prop_of) - val nparams = nparams_of thy (hd prednames) - val extra_modes = all_modes_of thy |> filter_out (fn (name, _) => member (op =) prednames name) - val preds = distinct (op =) (map (dest_Const o fst o (strip_intro_concl nparams)) intrs) - val _ $ u = Logic.strip_imp_concl (hd intrs); - val params = List.take (snd (strip_comb u), nparams); - val param_vs = maps term_vs params - val all_vs = terms_vs intrs - fun dest_prem t = - (case strip_comb t of - (v as Free _, ts) => if v mem params then Prem (ts, v) else Sidecond t - | (c as Const (@{const_name Not}, _), [t]) => (case dest_prem t of - Prem (ts, t) => Negprem (ts, t) - | Negprem _ => error ("Double negation not allowed in premise: " ^ (Syntax.string_of_term_global thy (c $ t))) - | Sidecond t => Sidecond (c $ t)) - | (c as Const (s, _), ts) => - if is_registered thy s then - let val (ts1, ts2) = chop (nparams_of thy s) ts - in Prem (ts2, list_comb (c, ts1)) end - else Sidecond t - | _ => Sidecond t) - fun add_clause intr (clauses, arities) = - let - val _ $ t = Logic.strip_imp_concl intr; - val (Const (name, T), ts) = strip_comb t; - val (ts1, ts2) = chop nparams ts; - val prems = map (dest_prem o HOLogic.dest_Trueprop) (Logic.strip_imp_prems intr); - val (Ts, Us) = chop nparams (binder_types T) - in - (AList.update op = (name, these (AList.lookup op = clauses name) @ - [(ts2, prems)]) clauses, - AList.update op = (name, (map (fn U => (case strip_type U of - (Rs as _ :: _, Type ("bool", [])) => SOME (length Rs) - | _ => NONE)) Ts, - length Us)) arities) - end; - val (clauses, arities) = fold add_clause intrs ([], []); - in (preds, nparams, all_vs, param_vs, extra_modes, clauses, arities) end; - -(** main function of predicate compiler **) - -fun add_equations_of steps prednames thy = - let - val _ = tracing ("Starting predicate compiler for predicates " ^ commas prednames ^ "...") - val (preds, nparams, all_vs, param_vs, extra_modes, clauses, arities) = - prepare_intrs thy prednames - val _ = tracing "Infering modes..." - val moded_clauses = #infer_modes steps thy extra_modes arities param_vs clauses - val modes = map (fn (p, mps) => (p, map fst mps)) moded_clauses - val _ = print_modes modes - val _ = print_moded_clauses thy moded_clauses - val _ = tracing "Defining executable functions..." - val thy' = fold (#create_definitions steps preds) modes thy - |> Theory.checkpoint - val _ = tracing "Compiling equations..." - val compiled_terms = - (#compile_preds steps) thy' all_vs param_vs preds moded_clauses - val _ = print_compiled_terms thy' compiled_terms - val _ = tracing "Proving equations..." - val result_thms = #prove steps thy' clauses preds (extra_modes @ modes) - moded_clauses compiled_terms - val qname = #qname steps - (* val attrib = gn thy => Attrib.attribute_i thy Code.add_eqn_attrib *) - val attrib = fn thy => Attrib.attribute_i thy (Attrib.internal (K (Thm.declaration_attribute - (fn thm => Context.mapping (Code.add_eqn thm) I)))) - val thy'' = fold (fn (name, result_thms) => fn thy => snd (PureThy.add_thmss - [((Binding.qualify true (Long_Name.base_name name) (Binding.name qname), result_thms), - [attrib thy ])] thy)) - (maps_modes result_thms) thy' - |> Theory.checkpoint - in - thy'' - end - -fun extend' value_of edges_of key (G, visited) = - let - val (G', v) = case try (Graph.get_node G) key of - SOME v => (G, v) - | NONE => (Graph.new_node (key, value_of key) G, value_of key) - val (G'', visited') = fold (extend' value_of edges_of) (edges_of (key, v) \\ visited) - (G', key :: visited) - in - (fold (Graph.add_edge o (pair key)) (edges_of (key, v)) G'', visited') - end; - -fun extend value_of edges_of key G = fst (extend' value_of edges_of key (G, [])) - -fun gen_add_equations steps names thy = - let - val thy' = PredData.map (fold (extend (fetch_pred_data thy) (depending_preds_of thy)) names) thy - |> Theory.checkpoint; - fun strong_conn_of gr keys = - Graph.strong_conn (Graph.subgraph (member (op =) (Graph.all_succs gr keys)) gr) - val scc = strong_conn_of (PredData.get thy') names - val thy'' = fold_rev - (fn preds => fn thy => - if #are_not_defined steps thy preds then add_equations_of steps preds thy else thy) - scc thy' |> Theory.checkpoint - in thy'' end - -(* different instantiantions of the predicate compiler *) - -val add_equations = gen_add_equations - {infer_modes = infer_modes false, - create_definitions = create_definitions, - compile_preds = compile_preds PredicateCompFuns.compfuns mk_fun_of false, - prove = prove, - are_not_defined = (fn thy => forall (null o modes_of thy)), - qname = "equation"} - -val add_sizelim_equations = gen_add_equations - {infer_modes = infer_modes false, - create_definitions = sizelim_create_definitions, - compile_preds = compile_preds PredicateCompFuns.compfuns mk_sizelim_fun_of true, - prove = prove_by_skip, - are_not_defined = (fn thy => fn preds => true), (* TODO *) - qname = "sizelim_equation" - } - -val add_quickcheck_equations = gen_add_equations - {infer_modes = infer_modes_with_generator, - create_definitions = rpred_create_definitions, - compile_preds = compile_preds RPredCompFuns.compfuns mk_generator_of true, - prove = prove_by_skip, - are_not_defined = (fn thy => fn preds => true), (* TODO *) - qname = "rpred_equation"} - -(** user interface **) - -(* generation of case rules from user-given introduction rules *) - -fun mk_casesrule ctxt nparams introrules = - let - val intros = map (Logic.unvarify o prop_of) introrules - val (pred, (params, args)) = strip_intro_concl nparams (hd intros) - val ([propname], ctxt1) = Variable.variant_fixes ["thesis"] ctxt - val prop = HOLogic.mk_Trueprop (Free (propname, HOLogic.boolT)) - val (argnames, ctxt2) = Variable.variant_fixes - (map (fn i => "a" ^ string_of_int i) (1 upto (length args))) ctxt1 - val argvs = map2 (curry Free) argnames (map fastype_of args) - fun mk_case intro = - let - val (_, (_, args)) = strip_intro_concl nparams intro - val prems = Logic.strip_imp_prems intro - val eqprems = map (HOLogic.mk_Trueprop o HOLogic.mk_eq) (argvs ~~ args) - val frees = (fold o fold_aterms) - (fn t as Free _ => - if member (op aconv) params t then I else insert (op aconv) t - | _ => I) (args @ prems) [] - in fold Logic.all frees (Logic.list_implies (eqprems @ prems, prop)) end - val assm = HOLogic.mk_Trueprop (list_comb (pred, params @ argvs)) - val cases = map mk_case intros - in Logic.list_implies (assm :: cases, prop) end; - -(* code_pred_intro attribute *) - -fun attrib f = Thm.declaration_attribute (fn thm => Context.mapping (f thm) I); - -val code_pred_intros_attrib = attrib add_intro; - -local - -(* TODO: make TheoryDataFun to GenericDataFun & remove duplication of local theory and theory *) -(* TODO: must create state to prove multiple cases *) -fun generic_code_pred prep_const raw_const lthy = - let - val thy = ProofContext.theory_of lthy - val const = prep_const thy raw_const - val lthy' = LocalTheory.theory (PredData.map - (extend (fetch_pred_data thy) (depending_preds_of thy) const)) lthy - |> LocalTheory.checkpoint - val thy' = ProofContext.theory_of lthy' - val preds = Graph.all_preds (PredData.get thy') [const] |> filter_out (has_elim thy') - fun mk_cases const = - let - val nparams = nparams_of thy' const - val intros = intros_of thy' const - in mk_casesrule lthy' nparams intros end - val cases_rules = map mk_cases preds - val cases = - map (fn case_rule => RuleCases.Case {fixes = [], - assumes = [("", Logic.strip_imp_prems case_rule)], - binds = [], cases = []}) cases_rules - val case_env = map2 (fn p => fn c => (Long_Name.base_name p, SOME c)) preds cases - val lthy'' = lthy' - |> fold Variable.auto_fixes cases_rules - |> ProofContext.add_cases true case_env - fun after_qed thms goal_ctxt = - let - val global_thms = ProofContext.export goal_ctxt - (ProofContext.init (ProofContext.theory_of goal_ctxt)) (map the_single thms) - in - goal_ctxt |> LocalTheory.theory (fold set_elim global_thms #> add_equations [const]) - end - in - Proof.theorem_i NONE after_qed (map (single o (rpair [])) cases_rules) lthy'' - end; - -structure P = OuterParse - -in - -val code_pred = generic_code_pred (K I); -val code_pred_cmd = generic_code_pred Code.read_const - -val setup = PredData.put (Graph.empty) #> - Attrib.setup @{binding code_pred_intros} (Scan.succeed (attrib add_intro)) - "adding alternative introduction rules for code generation of inductive predicates" -(* Attrib.setup @{binding code_ind_cases} (Scan.succeed add_elim_attrib) - "adding alternative elimination rules for code generation of inductive predicates"; - *) - (*FIXME name discrepancy in attribs and ML code*) - (*FIXME intros should be better named intro*) - (*FIXME why distinguished attribute for cases?*) - -val _ = OuterSyntax.local_theory_to_proof "code_pred" - "prove equations for predicate specified by intro/elim rules" - OuterKeyword.thy_goal (P.term_group >> code_pred_cmd) - -end - -(*FIXME -- Naming of auxiliary rules necessary? -- add default code equations P x y z = P_i_i_i x y z -*) - -(* transformation for code generation *) - -val eval_ref = Unsynchronized.ref (NONE : (unit -> term Predicate.pred) option); - -(*FIXME turn this into an LCF-guarded preprocessor for comprehensions*) -fun analyze_compr thy t_compr = - let - val split = case t_compr of (Const (@{const_name Collect}, _) $ t) => t - | _ => error ("Not a set comprehension: " ^ Syntax.string_of_term_global thy t_compr); - val (body, Ts, fp) = HOLogic.strip_psplits split; - val (pred as Const (name, T), all_args) = strip_comb body; - val (params, args) = chop (nparams_of thy name) all_args; - val user_mode = map_filter I (map_index - (fn (i, t) => case t of Bound j => if j < length Ts then NONE - else SOME (i+1) | _ => SOME (i+1)) args); (*FIXME dangling bounds should not occur*) - val modes = filter (fn Mode (_, is, _) => is = user_mode) - (modes_of_term (all_modes_of thy) (list_comb (pred, params))); - val m = case modes - of [] => error ("No mode possible for comprehension " - ^ Syntax.string_of_term_global thy t_compr) - | [m] => m - | m :: _ :: _ => (warning ("Multiple modes possible for comprehension " - ^ Syntax.string_of_term_global thy t_compr); m); - val (inargs, outargs) = split_smode user_mode args; - val t_pred = list_comb (compile_expr NONE thy (m, list_comb (pred, params)), inargs); - val t_eval = if null outargs then t_pred else let - val outargs_bounds = map (fn Bound i => i) outargs; - val outargsTs = map (nth Ts) outargs_bounds; - val T_pred = HOLogic.mk_tupleT outargsTs; - val T_compr = HOLogic.mk_ptupleT fp Ts; - val arrange_bounds = map_index I outargs_bounds - |> sort (prod_ord (K EQUAL) int_ord) - |> map fst; - val arrange = funpow (length outargs_bounds - 1) HOLogic.mk_split - (Term.list_abs (map (pair "") outargsTs, - HOLogic.mk_ptuple fp T_compr (map Bound arrange_bounds))) - in mk_map PredicateCompFuns.compfuns T_pred T_compr arrange t_pred end - in t_eval end; - -fun eval thy t_compr = - let - val t = analyze_compr thy t_compr; - val T = dest_predT PredicateCompFuns.compfuns (fastype_of t); - val t' = mk_map PredicateCompFuns.compfuns T HOLogic.termT (HOLogic.term_of_const T) t; - in (T, Code_ML.eval NONE ("Predicate_Compile.eval_ref", eval_ref) Predicate.map thy t' []) end; - -fun values ctxt k t_compr = - let - val thy = ProofContext.theory_of ctxt; - val (T, t) = eval thy t_compr; - val setT = HOLogic.mk_setT T; - val (ts, _) = Predicate.yieldn k t; - val elemsT = HOLogic.mk_set T ts; - in if k = ~1 orelse length ts < k then elemsT - else Const (@{const_name Lattices.sup}, setT --> setT --> setT) $ elemsT $ t_compr - end; - -fun values_cmd modes k raw_t state = - let - val ctxt = Toplevel.context_of state; - val t = Syntax.read_term ctxt raw_t; - val t' = values ctxt k t; - val ty' = Term.type_of t'; - val ctxt' = Variable.auto_fixes t' ctxt; - val p = PrintMode.with_modes modes (fn () => - Pretty.block [Pretty.quote (Syntax.pretty_term ctxt' t'), Pretty.fbrk, - Pretty.str "::", Pretty.brk 1, Pretty.quote (Syntax.pretty_typ ctxt' ty')]) (); - in Pretty.writeln p end; - -local structure P = OuterParse in - -val opt_modes = Scan.optional (P.$$$ "(" |-- P.!!! (Scan.repeat1 P.xname --| P.$$$ ")")) []; - -val _ = OuterSyntax.improper_command "values" "enumerate and print comprehensions" OuterKeyword.diag - (opt_modes -- Scan.optional P.nat ~1 -- P.term - >> (fn ((modes, k), t) => Toplevel.no_timing o Toplevel.keep - (values_cmd modes k t))); - -end; - -end; - diff -r 2b5b0f9e271c -r 1464ddca182b src/HOLCF/Cprod.thy --- a/src/HOLCF/Cprod.thy Thu Oct 29 16:23:57 2009 +0100 +++ b/src/HOLCF/Cprod.thy Wed Nov 04 11:40:59 2009 +0100 @@ -52,6 +52,7 @@ translations "\(CONST cpair\x\y). t" == "CONST csplit\(\ x y. t)" + "\(CONST Pair x y). t" => "CONST csplit\(\ x y. t)" subsection {* Convert all lemmas to the continuous versions *} @@ -150,6 +151,9 @@ lemma csplit3 [simp]: "csplit\cpair\z = z" by (simp add: csplit_def cpair_cfst_csnd) +lemma csplit_Pair [simp]: "csplit\f\(x, y) = f\x\y" +by (simp add: csplit_def cfst_def csnd_def) + lemmas Cprod_rews = cfst_cpair csnd_cpair csplit2 subsection {* Product type is a bifinite domain *} diff -r 2b5b0f9e271c -r 1464ddca182b src/HOLCF/Domain.thy --- a/src/HOLCF/Domain.thy Thu Oct 29 16:23:57 2009 +0100 +++ b/src/HOLCF/Domain.thy Wed Nov 04 11:40:59 2009 +0100 @@ -222,7 +222,7 @@ definition cprod_fun :: "('a \ 'b) \ ('c \ 'd) \ 'a \ 'c \ 'b \ 'd" where - "cprod_fun = (\ f g. csplit\(\ x y. x, g\y>))" + "cprod_fun = (\ f g. csplit\(\ x y. (f\x, g\y)))" definition u_fun :: "('a \ 'b) \ 'a u \ 'b u" diff -r 2b5b0f9e271c -r 1464ddca182b src/HOLCF/Fixrec.thy --- a/src/HOLCF/Fixrec.thy Thu Oct 29 16:23:57 2009 +0100 +++ b/src/HOLCF/Fixrec.thy Wed Nov 04 11:40:59 2009 +0100 @@ -526,7 +526,8 @@ lemma match_cpair_simps [simp]: "match_cpair\\x, y\\k = k\x\y" -by (simp add: match_cpair_def) + "match_cpair\(x, y)\k = k\x\y" +by (simp_all add: match_cpair_def) lemma match_spair_simps [simp]: "\x \ \; y \ \\ \ match_spair\(:x, y:)\k = k\x\y" @@ -610,6 +611,7 @@ (@{const_name sinr}, @{const_name match_sinr}), (@{const_name spair}, @{const_name match_spair}), (@{const_name cpair}, @{const_name match_cpair}), + (@{const_name Pair}, @{const_name match_cpair}), (@{const_name ONE}, @{const_name match_ONE}), (@{const_name TT}, @{const_name match_TT}), (@{const_name FF}, @{const_name match_FF}), @@ -618,4 +620,23 @@ hide (open) const return bind fail run cases +lemmas [fixrec_simp] = + run_strict run_fail run_return + mplus_strict mplus_fail mplus_return + spair_strict_iff + sinl_defined_iff + sinr_defined_iff + up_defined + ONE_defined + dist_eq_tr(1,2) + match_UU_simps + match_cpair_simps + match_spair_simps + match_sinl_simps + match_sinr_simps + match_up_simps + match_ONE_simps + match_TT_simps + match_FF_simps + end diff -r 2b5b0f9e271c -r 1464ddca182b src/HOLCF/Tools/Domain/domain_axioms.ML --- a/src/HOLCF/Tools/Domain/domain_axioms.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/HOLCF/Tools/Domain/domain_axioms.ML Wed Nov 04 11:40:59 2009 +0100 @@ -66,7 +66,7 @@ Bound(1+length cons+x-y)))`(dc_rep`Bound 0))) (when_funs cons)); val copy_def = - let fun r i = cproj (Bound 0) eqs i; + let fun r i = proj (Bound 0) eqs i; in ("copy_def", %%:(dname^"_copy") == /\ "f" (dc_abs oo (copy_of_dtyp r (dtyp_of_eq eqn)) oo dc_rep)) end; @@ -128,12 +128,12 @@ (* ----- axiom and definitions concerning induction ------------------------- *) - val reach_ax = ("reach", mk_trp(cproj (mk_fix (%%:(comp_dname^"_copy"))) eqs n + val reach_ax = ("reach", mk_trp(proj (mk_fix (%%:(comp_dname^"_copy"))) eqs n `%x_name === %:x_name)); val take_def = ("take_def", %%:(dname^"_take") == - mk_lam("n",cproj + mk_lam("n",proj (mk_iterate (Bound 0, %%:(comp_dname^"_copy"), UU)) eqs n)); val finite_def = ("finite_def", @@ -184,7 +184,7 @@ fun one_con (con,args) = let val nonrec_args = filter_out is_rec args; - val rec_args = List.filter is_rec args; + val rec_args = filter is_rec args; val recs_cnt = length rec_args; val allargs = nonrec_args @ rec_args @ map (upd_vname (fn s=> s^"'")) rec_args; diff -r 2b5b0f9e271c -r 1464ddca182b src/HOLCF/Tools/Domain/domain_library.ML --- a/src/HOLCF/Tools/Domain/domain_library.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/HOLCF/Tools/Domain/domain_library.ML Wed Nov 04 11:40:59 2009 +0100 @@ -94,7 +94,6 @@ val mk_iterate : term * term * term -> term; val mk_fail : term; val mk_return : term -> term; - val cproj : term -> 'a list -> int -> term; val list_ccomb : term * term list -> term; (* val con_app : string -> ('a * 'b * string) list -> term; @@ -241,8 +240,8 @@ val upd_vname = upd_third; fun is_rec arg = rec_of arg >=0; fun is_nonlazy_rec arg = is_rec arg andalso not (is_lazy arg); -fun nonlazy args = map vname (filter_out is_lazy args); -fun nonlazy_rec args = map vname (List.filter is_nonlazy_rec args); +fun nonlazy args = map vname (filter_out is_lazy args); +fun nonlazy_rec args = map vname (filter is_nonlazy_rec args); (* ----- combinators for making dtyps ----- *) @@ -312,8 +311,6 @@ fun mk_compact t = %%: @{const_name compact} $ t; val ID = %%: @{const_name ID}; fun mk_strictify t = %%: @{const_name strictify}`t; -fun mk_cfst t = %%: @{const_name cfst}`t; -fun mk_csnd t = %%: @{const_name csnd}`t; (*val csplitN = "Cprod.csplit";*) (*val sfstN = "Sprod.sfst";*) (*val ssndN = "Sprod.ssnd";*) @@ -344,7 +341,6 @@ | prj f1 _ x (_::y::ys) 0 = f1 x y | prj f1 f2 x (y:: ys) j = prj f1 f2 (f2 x y) ys (j-1); fun proj x = prj (fn S => K(%%:"fst" $S)) (fn S => K(%%:"snd" $S)) x; -fun cproj x = prj (fn S => K(mk_cfst S)) (fn S => K(mk_csnd S)) x; fun lift tfn = Library.foldr (fn (x,t)=> (mk_trp(tfn x) ===> t)); fun /\ v T = %%: @{const_name Abs_CFun} $ mk_lam(v,T); @@ -353,7 +349,7 @@ val UU = %%: @{const_name UU}; fun strict f = f`UU === UU; fun defined t = t ~= UU; -fun cpair (t,u) = %%: @{const_name cpair}`t`u; +fun cpair (t,u) = %%: @{const_name Pair} $ t $ u; fun spair (t,u) = %%: @{const_name spair}`t`u; fun mk_ctuple [] = HOLogic.unit (* used in match_defs *) | mk_ctuple ts = foldr1 cpair ts; diff -r 2b5b0f9e271c -r 1464ddca182b src/HOLCF/Tools/Domain/domain_theorems.ML --- a/src/HOLCF/Tools/Domain/domain_theorems.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/HOLCF/Tools/Domain/domain_theorems.ML Wed Nov 04 11:40:59 2009 +0100 @@ -33,6 +33,8 @@ val antisym_less_inverse = @{thm below_antisym_inverse}; val beta_cfun = @{thm beta_cfun}; val cfun_arg_cong = @{thm cfun_arg_cong}; +val ch2ch_fst = @{thm ch2ch_fst}; +val ch2ch_snd = @{thm ch2ch_snd}; val ch2ch_Rep_CFunL = @{thm ch2ch_Rep_CFunL}; val ch2ch_Rep_CFunR = @{thm ch2ch_Rep_CFunR}; val chain_iterate = @{thm chain_iterate}; @@ -43,6 +45,14 @@ val compact_up = @{thm compact_up}; val contlub_cfun_arg = @{thm contlub_cfun_arg}; val contlub_cfun_fun = @{thm contlub_cfun_fun}; +val contlub_fst = @{thm contlub_fst}; +val contlub_snd = @{thm contlub_snd}; +val contlubE = @{thm contlubE}; +val cont_const = @{thm cont_const}; +val cont_id = @{thm cont_id}; +val cont2cont_fst = @{thm cont2cont_fst}; +val cont2cont_snd = @{thm cont2cont_snd}; +val cont2cont_Rep_CFun = @{thm cont2cont_Rep_CFun}; val fix_def2 = @{thm fix_def2}; val injection_eq = @{thm injection_eq}; val injection_less = @{thm injection_below}; @@ -116,7 +126,7 @@ val chain_tac = REPEAT_DETERM o resolve_tac - [chain_iterate, ch2ch_Rep_CFunR, ch2ch_Rep_CFunL]; + [chain_iterate, ch2ch_Rep_CFunR, ch2ch_Rep_CFunL, ch2ch_fst, ch2ch_snd]; (* ----- general proofs ----------------------------------------------------- *) @@ -446,7 +456,7 @@ val nlas = nonlazy args; val vns = map vname args; val vnn = List.nth (vns, n); - val nlas' = List.filter (fn v => v <> vnn) nlas; + val nlas' = filter (fn v => v <> vnn) nlas; val lhs = (%%:sel)`(con_app con args); val goal = lift_defined %: (nlas', mk_trp (lhs === %:vnn)); fun tacs1 ctxt = @@ -555,7 +565,7 @@ val sargs = case largs of [_] => [] | _ => nonlazy args; val prop = lift_defined %: (sargs, mk_trp (prem === concl)); in pg con_appls prop end; - val cons' = List.filter (fn (_,args) => args<>[]) cons; + val cons' = filter (fn (_,args) => args<>[]) cons; in val _ = trace " Proving inverts..."; val inverts = @@ -589,11 +599,11 @@ val lhs = dc_copy`%"f"`(con_app con args); fun one_rhs arg = if DatatypeAux.is_rec_type (dtyp_of arg) - then Domain_Axioms.copy_of_dtyp (cproj (%:"f") eqs) (dtyp_of arg) ` (%# arg) + then Domain_Axioms.copy_of_dtyp (proj (%:"f") eqs) (dtyp_of arg) ` (%# arg) else (%# arg); val rhs = con_app2 con one_rhs args; val goal = lift_defined %: (nonlazy_rec args, mk_trp (lhs === rhs)); - val args' = List.filter (fn a => not (is_rec a orelse is_lazy a)) args; + val args' = filter_out (fn a => is_rec a orelse is_lazy a) args; val stricts = abs_strict :: rep_strict :: @{thms domain_fun_stricts}; fun tacs1 ctxt = map (case_UU_tac ctxt stricts 1 o vname) args'; val rules = [ax_abs_iso] @ @{thms domain_fun_simps}; @@ -616,7 +626,7 @@ fun has_nonlazy_rec (_, args) = exists is_nonlazy_rec args; in val _ = trace " Proving copy_stricts..."; - val copy_stricts = map one_strict (List.filter has_nonlazy_rec cons); + val copy_stricts = map one_strict (filter has_nonlazy_rec cons); end; val copy_rews = copy_strict :: copy_apps @ copy_stricts; @@ -630,7 +640,8 @@ ((Binding.name "casedist" , [casedist] ), [Induct.cases_type dname]), ((Binding.name "when_rews" , when_rews ), [Simplifier.simp_add]), ((Binding.name "compacts" , con_compacts), [Simplifier.simp_add]), - ((Binding.name "con_rews" , con_rews ), [Simplifier.simp_add]), + ((Binding.name "con_rews" , con_rews ), + [Simplifier.simp_add, Fixrec.fixrec_simp_add]), ((Binding.name "sel_rews" , sel_rews ), [Simplifier.simp_add]), ((Binding.name "dis_rews" , dis_rews ), [Simplifier.simp_add]), ((Binding.name "pat_rews" , pat_rews ), [Simplifier.simp_add]), @@ -639,7 +650,8 @@ ((Binding.name "inverts" , inverts ), [Simplifier.simp_add]), ((Binding.name "injects" , injects ), [Simplifier.simp_add]), ((Binding.name "copy_rews" , copy_rews ), [Simplifier.simp_add]), - ((Binding.name "match_rews", mat_rews ), [Simplifier.simp_add])] + ((Binding.name "match_rews", mat_rews ), + [Simplifier.simp_add, Fixrec.fixrec_simp_add])] |> Sign.parent_path |> pair (iso_rews @ when_rews @ con_rews @ sel_rews @ dis_rews @ pat_rews @ dist_les @ dist_eqs @ copy_rews) @@ -722,7 +734,7 @@ in Library.foldr mk_all (map vname args, lhs === rhs) end; fun mk_eqns ((dn, _), cons) = map (mk_eqn dn) cons; val goal = mk_trp (foldr1 mk_conj (maps mk_eqns eqs)); - val simps = List.filter (has_fewer_prems 1) copy_rews; + val simps = filter (has_fewer_prems 1) copy_rews; fun con_tac ctxt (con, args) = if nonlazy_rec args = [] then all_tac @@ -747,7 +759,7 @@ let fun ind_hyp arg = %:(P_name (1 + rec_of arg)) $ bound_arg args arg; val t1 = mk_trp (%:p $ con_app2 con (bound_arg args) args); - val t2 = lift ind_hyp (List.filter is_rec args, t1); + val t2 = lift ind_hyp (filter is_rec args, t1); val t3 = lift_defined (bound_arg (map vname args)) (nonlazy args, t2); in Library.foldr mk_All (map vname args, t3) end; @@ -767,7 +779,7 @@ maps (fn (_,args) => resolve_tac prems 1 :: map (K(atac 1)) (nonlazy args) @ - map (K(atac 1)) (List.filter is_rec args)) + map (K(atac 1)) (filter is_rec args)) cons)) conss); local @@ -812,10 +824,10 @@ (List.nth (dnames, rec_of arg) ^ "_take n$" ^ vname arg); fun con_tacs (con, args) = asm_simp_tac take_ss 1 :: - map arg_tac (List.filter is_nonlazy_rec args) @ + map arg_tac (filter is_nonlazy_rec args) @ [resolve_tac prems 1] @ - map (K (atac 1)) (nonlazy args) @ - map (K (etac spec 1)) (List.filter is_rec args); + map (K (atac 1)) (nonlazy args) @ + map (K (etac spec 1)) (filter is_rec args); fun cases_tacs (cons, cases) = res_inst_tac context [(("x", 0), "x")] cases 1 :: asm_simp_tac (take_ss addsimps prems) 1 :: @@ -836,12 +848,14 @@ val rhs = dc_take dn $ Bound 0 `%(x_name n^"'"); val concl = mk_trp (%:(x_name n) === %:(x_name n^"'")); val goal = mk_All ("n", mk_trp (lhs === rhs)) ===> concl; + val rules = [contlub_fst RS contlubE RS ssubst, + contlub_snd RS contlubE RS ssubst]; fun tacf {prems, context} = [ res_inst_tac context [(("t", 0), x_name n )] (ax_reach RS subst) 1, res_inst_tac context [(("t", 0), x_name n^"'")] (ax_reach RS subst) 1, stac fix_def2 1, REPEAT (CHANGED - (rtac (contlub_cfun_arg RS ssubst) 1 THEN chain_tac 1)), + (resolve_tac rules 1 THEN chain_tac 1)), stac contlub_cfun_fun 1, stac contlub_cfun_fun 2, rtac lub_equal 3, @@ -955,6 +969,9 @@ fun one_adm n _ = mk_trp (mk_adm (%:(P_name n))); fun concf n dn = %:(P_name n) $ %:(x_name n); in Logic.list_implies (mapn one_adm 1 dnames, ind_term concf) end; + val cont_rules = + [cont_id, cont_const, cont2cont_Rep_CFun, + cont2cont_fst, cont2cont_snd]; fun tacf {prems, context} = map (fn ax_reach => rtac (ax_reach RS subst) 1) axs_reach @ [ quant_tac context 1, @@ -963,13 +980,14 @@ REPEAT_DETERM ( TRY (rtac adm_conj 1) THEN rtac adm_subst 1 THEN - cont_tacR 1 THEN resolve_tac prems 1), + REPEAT (resolve_tac cont_rules 1) THEN + resolve_tac prems 1), strip_tac 1, rtac (rewrite_rule axs_take_def finite_ind) 1, ind_prems_tac prems]; val ind = (pg'' thy [] goal tacf handle ERROR _ => - (warning "Cannot prove infinite induction rule"; refl)); + (warning "Cannot prove infinite induction rule"; TrueI)); in (finites, ind) end ) handle THM _ => diff -r 2b5b0f9e271c -r 1464ddca182b src/HOLCF/Tools/fixrec.ML --- a/src/HOLCF/Tools/fixrec.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/HOLCF/Tools/fixrec.ML Wed Nov 04 11:40:59 2009 +0100 @@ -13,6 +13,9 @@ val add_fixpat: Thm.binding * term list -> theory -> theory val add_fixpat_cmd: Attrib.binding * string list -> theory -> theory val add_matchers: (string * string) list -> theory -> theory + val fixrec_simp_add: Thm.attribute + val fixrec_simp_del: Thm.attribute + val fixrec_simp_tac: Proof.context -> int -> tactic val setup: theory -> theory end; @@ -36,20 +39,9 @@ infixr 6 ->>; val (op ->>) = cfunT; -fun cfunsT (Ts, U) = List.foldr cfunT U Ts; - fun dest_cfunT (Type(@{type_name "->"}, [T, U])) = (T, U) | dest_cfunT T = raise TYPE ("dest_cfunT", [T], []); -fun binder_cfun (Type(@{type_name "->"},[T, U])) = T :: binder_cfun U - | binder_cfun _ = []; - -fun body_cfun (Type(@{type_name "->"},[T, U])) = body_cfun U - | body_cfun T = T; - -fun strip_cfun T : typ list * typ = - (binder_cfun T, body_cfun T); - fun maybeT T = Type(@{type_name "maybe"}, [T]); fun dest_maybeT (Type(@{type_name "maybe"}, [T])) = T @@ -59,9 +51,27 @@ | tupleT [T] = T | tupleT (T :: Ts) = HOLogic.mk_prodT (T, tupleT Ts); +local + +fun binder_cfun (Type(@{type_name "->"},[T, U])) = T :: binder_cfun U + | binder_cfun (Type(@{type_name "fun"},[T, U])) = T :: binder_cfun U + | binder_cfun _ = []; + +fun body_cfun (Type(@{type_name "->"},[T, U])) = body_cfun U + | body_cfun (Type(@{type_name "fun"},[T, U])) = body_cfun U + | body_cfun T = T; + +fun strip_cfun T : typ list * typ = + (binder_cfun T, body_cfun T); + +fun cfunsT (Ts, U) = List.foldr cfunT U Ts; + +in + fun matchT (T, U) = body_cfun T ->> cfunsT (binder_cfun T, U) ->> U; +end (*************************************************************************) (***************************** building terms ****************************) @@ -150,6 +160,31 @@ (************* fixed-point definitions and unfolding theorems ************) (*************************************************************************) +structure FixrecUnfoldData = GenericDataFun ( + type T = thm Symtab.table; + val empty = Symtab.empty; + val copy = I; + val extend = I; + val merge = K (Symtab.merge (K true)); +); + +local + +fun name_of (Const (n, T)) = n + | name_of (Free (n, T)) = n + | name_of _ = fixrec_err "name_of" + +val lhs_name = + name_of o fst o HOLogic.dest_eq o HOLogic.dest_Trueprop o prop_of; + +in + +val add_unfold : Thm.attribute = + Thm.declaration_attribute + (fn th => FixrecUnfoldData.map (Symtab.insert (K true) (lhs_name th, th))); + +end + fun add_fixdefs (fixes : ((binding * typ) * mixfix) list) (spec : (Attrib.binding * term) list) @@ -187,16 +222,28 @@ val tuple_unfold_thm = (def_cont_fix_eq OF [tuple_fixdef_thm, cont_thm]) |> LocalDefs.unfold lthy' @{thms split_conv}; fun unfolds [] thm = [] - | unfolds (n::[]) thm = [(n^"_unfold", thm)] + | unfolds (n::[]) thm = [(n, thm)] | unfolds (n::ns) thm = let val thmL = thm RS @{thm Pair_eqD1}; val thmR = thm RS @{thm Pair_eqD2}; - in (n^"_unfold", thmL) :: unfolds ns thmR end; + in (n, thmL) :: unfolds ns thmR end; val unfold_thms = unfolds names tuple_unfold_thm; - fun mk_note (n, thm) = ((Binding.name n, []), [thm]); + val induct_note : Attrib.binding * Thm.thm list = + let + val thm_name = Binding.name (all_names ^ "_induct"); + in + ((thm_name, []), [tuple_induct_thm]) + end; + fun unfold_note (name, thm) : Attrib.binding * Thm.thm list = + let + val thm_name = Binding.name (name ^ "_unfold"); + val src = Attrib.internal (K add_unfold); + in + ((thm_name, [src]), [thm]) + end; val (thmss, lthy'') = lthy' - |> fold_map (LocalTheory.note Thm.generatedK o mk_note) - ((all_names ^ "_induct", tuple_induct_thm) :: unfold_thms); + |> fold_map (LocalTheory.note Thm.generatedK) + (induct_note :: map unfold_note unfold_thms); in (lthy'', names, fixdef_thms, map snd unfold_thms) end; @@ -210,7 +257,7 @@ val empty = Symtab.empty; val copy = I; val extend = I; - fun merge _ tabs : T = Symtab.merge (K true) tabs; + val merge = K (Symtab.merge (K true)); ); (* associate match functions with pattern constants *) @@ -235,10 +282,16 @@ | Const(@{const_name Rep_CFun},_)$f$x => let val (rhs', v, taken') = pre_build match_name x rhs [] taken; in pre_build match_name f rhs' (v::vs) taken' end + | f$(v as Free(n,T)) => + pre_build match_name f rhs (v::vs) taken + | f$x => + let val (rhs', v, taken') = pre_build match_name x rhs [] taken; + in pre_build match_name f rhs' (v::vs) taken' end | Const(c,T) => let val n = Name.variant taken "v"; fun result_type (Type(@{type_name "->"},[_,T])) (x::xs) = result_type T xs + | result_type (Type (@{type_name "fun"},[_,T])) (x::xs) = result_type T xs | result_type T _ = T; val v = Free(n, result_type T vs); val m = Const(match_name c, matchT (T, fastype_of rhs)); @@ -310,6 +363,45 @@ (********************** Proving associated theorems **********************) (*************************************************************************) +structure FixrecSimpData = GenericDataFun ( + type T = simpset; + val empty = + HOL_basic_ss + addsimps simp_thms + addsimps [@{thm beta_cfun}] + addsimprocs [@{simproc cont_proc}]; + val copy = I; + val extend = I; + val merge = K merge_ss; +); + +fun fixrec_simp_tac ctxt = + let + val tab = FixrecUnfoldData.get (Context.Proof ctxt); + val ss = FixrecSimpData.get (Context.Proof ctxt); + fun concl t = + if Logic.is_all t then concl (snd (Logic.dest_all t)) + else HOLogic.dest_Trueprop (Logic.strip_imp_concl t); + fun tac (t, i) = + let + val Const (c, T) = chead_of (fst (HOLogic.dest_eq (concl t))); + val unfold_thm = the (Symtab.lookup tab c); + val rule = unfold_thm RS @{thm ssubst_lhs}; + in + CHANGED (rtac rule i THEN asm_simp_tac ss i) + end + in + SUBGOAL (fn ti => tac ti handle _ => no_tac) + end; + +val fixrec_simp_add : Thm.attribute = + Thm.declaration_attribute + (fn th => FixrecSimpData.map (fn ss => ss addsimps [th])); + +val fixrec_simp_del : Thm.attribute = + Thm.declaration_attribute + (fn th => FixrecSimpData.map (fn ss => ss delsimps [th])); + (* proves a block of pattern matching equations as theorems, using unfold *) fun make_simps ctxt (unfold_thm, eqns : (Attrib.binding * term) list) = let @@ -427,9 +519,16 @@ val _ = OuterSyntax.command "fixpat" "define rewrites for fixrec functions" K.thy_decl (SpecParse.specs >> (Toplevel.theory o add_fixpat_cmd)); - + end; -val setup = FixrecMatchData.init; +val setup = + FixrecMatchData.init + #> Attrib.setup @{binding fixrec_simp} + (Attrib.add_del fixrec_simp_add fixrec_simp_del) + "declaration of fixrec simp rule" + #> Method.setup @{binding fixrec_simp} + (Scan.succeed (SIMPLE_METHOD' o fixrec_simp_tac)) + "pattern prover for fixrec constants"; end; diff -r 2b5b0f9e271c -r 1464ddca182b src/HOLCF/ex/Fixrec_ex.thy --- a/src/HOLCF/ex/Fixrec_ex.thy Thu Oct 29 16:23:57 2009 +0100 +++ b/src/HOLCF/ex/Fixrec_ex.thy Wed Nov 04 11:40:59 2009 +0100 @@ -31,8 +31,13 @@ fixrec from_sinl_up :: "'a u \ 'b \ 'a" where "from_sinl_up\(sinl\(up\x)) = x" +text {* Fixrec also works with the HOL pair constructor. *} -subsection {* Examples using @{text fixpat} *} +fixrec down2 :: "'a u \ 'b u \ 'a \ 'b" +where "down2\(up\x, up\y) = (x, y)" + + +subsection {* Examples using @{text fixrec_simp} *} text {* A type of lazy lists. *} @@ -48,28 +53,30 @@ "lzip\(lCons\x\xs)\(lCons\y\ys) = lCons\\(lzip\xs\ys)" | "lzip\lNil\lNil = lNil" -text {* @{text fixpat} is useful for producing strictness theorems. *} +text {* @{text fixrec_simp} is useful for producing strictness theorems. *} text {* Note that pattern matching is done in left-to-right order. *} -fixpat lzip_stricts [simp]: - "lzip\\\ys" - "lzip\lNil\\" - "lzip\(lCons\x\xs)\\" +lemma lzip_stricts [simp]: + "lzip\\\ys = \" + "lzip\lNil\\ = \" + "lzip\(lCons\x\xs)\\ = \" +by fixrec_simp+ -text {* @{text fixpat} can also produce rules for missing cases. *} +text {* @{text fixrec_simp} can also produce rules for missing cases. *} -fixpat lzip_undefs [simp]: - "lzip\lNil\(lCons\y\ys)" - "lzip\(lCons\x\xs)\lNil" +lemma lzip_undefs [simp]: + "lzip\lNil\(lCons\y\ys) = \" + "lzip\(lCons\x\xs)\lNil = \" +by fixrec_simp+ subsection {* Pattern matching with bottoms *} text {* - As an alternative to using @{text fixpat}, it is also possible to - use bottom as a constructor pattern. When using a bottom pattern, - the right-hand-side must also be bottom; otherwise, @{text fixrec} - will not be able to prove the equation. + As an alternative to using @{text fixrec_simp}, it is also possible + to use bottom as a constructor pattern. When using a bottom + pattern, the right-hand-side must also be bottom; otherwise, @{text + fixrec} will not be able to prove the equation. *} fixrec @@ -83,7 +90,7 @@ pattern does not change the meaning of the function. For example, in the definition of @{term from_sinr_up}, the first equation is actually redundant, and could have been proven separately by - @{text fixpat}. + @{text fixrec_simp}. *} text {* @@ -120,17 +127,19 @@ does not produce any simp rules. *} -text {* Simp rules can be generated later using @{text fixpat}. *} +text {* Simp rules can be generated later using @{text fixrec_simp}. *} -fixpat lzip2_simps [simp]: - "lzip2\(lCons\x\xs)\(lCons\y\ys)" - "lzip2\(lCons\x\xs)\lNil" - "lzip2\lNil\(lCons\y\ys)" - "lzip2\lNil\lNil" +lemma lzip2_simps [simp]: + "lzip2\(lCons\x\xs)\(lCons\y\ys) = lCons\\(lzip\xs\ys)" + "lzip2\(lCons\x\xs)\lNil = lNil" + "lzip2\lNil\(lCons\y\ys) = lNil" + "lzip2\lNil\lNil = lNil" +by fixrec_simp+ -fixpat lzip2_stricts [simp]: - "lzip2\\\ys" - "lzip2\(lCons\x\xs)\\" +lemma lzip2_stricts [simp]: + "lzip2\\\ys = \" + "lzip2\(lCons\x\xs)\\ = \" +by fixrec_simp+ subsection {* Mutual recursion with @{text fixrec} *} @@ -156,8 +165,11 @@ | "ts \ \ \ map_forest\f\(Trees\t\ts) = Trees\(map_tree\f\t)\(map_forest\f\ts)" -fixpat map_tree_strict [simp]: "map_tree\f\\" -fixpat map_forest_strict [simp]: "map_forest\f\\" +lemma map_tree_strict [simp]: "map_tree\f\\ = \" +by fixrec_simp + +lemma map_forest_strict [simp]: "map_forest\f\\ = \" +by fixrec_simp text {* Theorems generated: diff -r 2b5b0f9e271c -r 1464ddca182b src/Provers/Arith/fast_lin_arith.ML --- a/src/Provers/Arith/fast_lin_arith.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/Provers/Arith/fast_lin_arith.ML Wed Nov 04 11:40:59 2009 +0100 @@ -340,7 +340,7 @@ case ty of Eq => k <> 0 | Le => k > 0 | Lt => k >= 0; fun calc_blowup l = - let val (p,n) = List.partition (curry (op <) 0) (List.filter (curry (op <>) 0) l) + let val (p,n) = List.partition (curry (op <) 0) (filter (curry (op <>) 0) l) in length p * length n end; (* ------------------------------------------------------------------------- *) diff -r 2b5b0f9e271c -r 1464ddca182b src/Provers/blast.ML --- a/src/Provers/blast.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/Provers/blast.ML Wed Nov 04 11:40:59 2009 +0100 @@ -55,7 +55,7 @@ swrappers: (string * wrapper) list, uwrappers: (string * wrapper) list, safe0_netpair: netpair, safep_netpair: netpair, - haz_netpair: netpair, dup_netpair: netpair, xtra_netpair: ContextRules.netpair} + haz_netpair: netpair, dup_netpair: netpair, xtra_netpair: Context_Rules.netpair} val cla_modifiers: Method.modifier parser list val cla_meth': (claset -> int -> tactic) -> thm list -> Proof.context -> Proof.method end; diff -r 2b5b0f9e271c -r 1464ddca182b src/Provers/clasimp.ML --- a/src/Provers/clasimp.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/Provers/clasimp.ML Wed Nov 04 11:40:59 2009 +0100 @@ -139,9 +139,9 @@ fun modifier att (x, ths) = fst (Library.foldl_map (Library.apply [att]) (x, rev ths)); -val addXIs = modifier (ContextRules.intro_query NONE); -val addXEs = modifier (ContextRules.elim_query NONE); -val delXs = modifier ContextRules.rule_del; +val addXIs = modifier (Context_Rules.intro_query NONE); +val addXEs = modifier (Context_Rules.elim_query NONE); +val delXs = modifier Context_Rules.rule_del; in diff -r 2b5b0f9e271c -r 1464ddca182b src/Provers/classical.ML --- a/src/Provers/classical.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/Provers/classical.ML Wed Nov 04 11:40:59 2009 +0100 @@ -45,7 +45,7 @@ uwrappers: (string * wrapper) list, safe0_netpair: netpair, safep_netpair: netpair, haz_netpair: netpair, dup_netpair: netpair, - xtra_netpair: ContextRules.netpair} + xtra_netpair: Context_Rules.netpair} val merge_cs : claset * claset -> claset val addDs : claset * thm list -> claset val addEs : claset * thm list -> claset @@ -198,10 +198,10 @@ (*Uses introduction rules in the normal way, or on negated assumptions, trying rules in order. *) fun swap_res_tac rls = - let fun addrl (rl,brls) = (false, rl) :: (true, rl RSN (2, Data.swap)) :: brls + let fun addrl rl brls = (false, rl) :: (true, rl RSN (2, Data.swap)) :: brls in assume_tac ORELSE' contr_tac ORELSE' - biresolve_tac (List.foldr addrl [] rls) + biresolve_tac (fold_rev addrl rls []) end; (*Duplication of hazardous rules, for complete provers*) @@ -224,7 +224,7 @@ safep_netpair : netpair, (*nets for >0 subgoals*) haz_netpair : netpair, (*nets for unsafe rules*) dup_netpair : netpair, (*nets for duplication*) - xtra_netpair : ContextRules.netpair}; (*nets for extra rules*) + xtra_netpair : Context_Rules.netpair}; (*nets for extra rules*) (*Desired invariants are safe0_netpair = build safe0_brls, @@ -670,7 +670,7 @@ (*version of bimatch_from_nets_tac that only applies rules that create precisely n subgoals.*) fun n_bimatch_from_nets_tac n = - biresolution_from_nets_tac (order_list o List.filter (nsubgoalsP n)) true; + biresolution_from_nets_tac (order_list o filter (nsubgoalsP n)) true; fun eq_contr_tac i = ematch_tac [not_elim] i THEN eq_assume_tac i; val eq_assume_contr_tac = eq_assume_tac ORELSE' eq_contr_tac; @@ -898,7 +898,7 @@ fun haz_dest w = attrib (addE w o make_elim); val haz_elim = attrib o addE; val haz_intro = attrib o addI; -val rule_del = attrib delrule o ContextRules.rule_del; +val rule_del = attrib delrule o Context_Rules.rule_del; end; @@ -914,11 +914,11 @@ val setup_attrs = Attrib.setup @{binding swapped} (Scan.succeed swapped) "classical swap of introduction rule" #> - Attrib.setup @{binding dest} (ContextRules.add safe_dest haz_dest ContextRules.dest_query) + Attrib.setup @{binding dest} (Context_Rules.add safe_dest haz_dest Context_Rules.dest_query) "declaration of Classical destruction rule" #> - Attrib.setup @{binding elim} (ContextRules.add safe_elim haz_elim ContextRules.elim_query) + Attrib.setup @{binding elim} (Context_Rules.add safe_elim haz_elim Context_Rules.elim_query) "declaration of Classical elimination rule" #> - Attrib.setup @{binding intro} (ContextRules.add safe_intro haz_intro ContextRules.intro_query) + Attrib.setup @{binding intro} (Context_Rules.add safe_intro haz_intro Context_Rules.intro_query) "declaration of Classical introduction rule" #> Attrib.setup @{binding rule} (Scan.lift Args.del >> K rule_del) "remove declaration of intro/elim/dest rule"; @@ -931,9 +931,9 @@ fun some_rule_tac ctxt facts = SUBGOAL (fn (goal, i) => let - val [rules1, rules2, rules4] = ContextRules.find_rules false facts goal ctxt; + val [rules1, rules2, rules4] = Context_Rules.find_rules false facts goal ctxt; val CS {xtra_netpair, ...} = claset_of ctxt; - val rules3 = ContextRules.find_rules_netpair true facts goal xtra_netpair; + val rules3 = Context_Rules.find_rules_netpair true facts goal xtra_netpair; val rules = rules1 @ rules2 @ rules3 @ rules4; val ruleq = Drule.multi_resolves facts rules; in diff -r 2b5b0f9e271c -r 1464ddca182b src/Provers/splitter.ML --- a/src/Provers/splitter.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/Provers/splitter.ML Wed Nov 04 11:40:59 2009 +0100 @@ -192,7 +192,7 @@ if n > length ts then [] else let val lev = length apsns val lbnos = fold add_lbnos (Library.take (n, ts)) [] - val flbnos = List.filter (fn i => i < lev) lbnos + val flbnos = filter (fn i => i < lev) lbnos val tt = incr_boundvars (~lev) t in if null flbnos then if T = T' then [(thm,[],pos,TB,tt)] else [] diff -r 2b5b0f9e271c -r 1464ddca182b src/Provers/typedsimp.ML --- a/src/Provers/typedsimp.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/Provers/typedsimp.ML Wed Nov 04 11:40:59 2009 +0100 @@ -64,12 +64,12 @@ (*If the rule proves an equality then add both forms to simp_rls else add the rule to other_rls*) -fun add_rule (rl, (simp_rls, other_rls)) = +fun add_rule rl (simp_rls, other_rls) = (simp_rule rl :: resimp_rule rl :: simp_rls, other_rls) handle THM _ => (simp_rls, rl :: other_rls); (*Given the list rls, return the pair (simp_rls, other_rls).*) -fun process_rules rls = List.foldr add_rule ([],[]) rls; +fun process_rules rls = fold_rev add_rule rls ([], []); (*Given list of rewrite rules, return list of both forms, reject others*) fun process_rewrites rls = diff -r 2b5b0f9e271c -r 1464ddca182b src/Pure/IsaMakefile --- a/src/Pure/IsaMakefile Thu Oct 29 16:23:57 2009 +0100 +++ b/src/Pure/IsaMakefile Wed Nov 04 11:40:59 2009 +0100 @@ -68,19 +68,20 @@ Isar/outer_syntax.ML Isar/overloading.ML Isar/proof.ML \ Isar/proof_context.ML Isar/proof_display.ML Isar/proof_node.ML \ Isar/rule_cases.ML Isar/rule_insts.ML Isar/runtime.ML \ - Isar/skip_proof.ML Isar/spec_parse.ML Isar/specification.ML \ - Isar/theory_target.ML Isar/toplevel.ML Isar/value_parse.ML \ - ML/ml_antiquote.ML ML/ml_compiler.ML ML/ml_compiler_polyml-5.3.ML \ - ML/ml_context.ML ML/ml_env.ML ML/ml_lex.ML ML/ml_parse.ML \ - ML/ml_syntax.ML ML/ml_thms.ML ML-Systems/install_pp_polyml.ML \ - ML-Systems/install_pp_polyml-5.3.ML ML-Systems/use_context.ML \ - Proof/extraction.ML Proof/proof_rewrite_rules.ML \ - Proof/proof_syntax.ML Proof/proofchecker.ML Proof/reconstruct.ML \ - ProofGeneral/pgip.ML ProofGeneral/pgip_input.ML \ - ProofGeneral/pgip_isabelle.ML ProofGeneral/pgip_markup.ML \ - ProofGeneral/pgip_output.ML ProofGeneral/pgip_parser.ML \ - ProofGeneral/pgip_tests.ML ProofGeneral/pgip_types.ML \ - ProofGeneral/preferences.ML ProofGeneral/proof_general_emacs.ML \ + Isar/skip_proof.ML Isar/spec_parse.ML Isar/spec_rules.ML \ + Isar/specification.ML Isar/theory_target.ML Isar/toplevel.ML \ + Isar/value_parse.ML ML/ml_antiquote.ML ML/ml_compiler.ML \ + ML/ml_compiler_polyml-5.3.ML ML/ml_context.ML ML/ml_env.ML \ + ML/ml_lex.ML ML/ml_parse.ML ML/ml_syntax.ML ML/ml_thms.ML \ + ML-Systems/install_pp_polyml.ML ML-Systems/install_pp_polyml-5.3.ML \ + ML-Systems/use_context.ML Proof/extraction.ML \ + Proof/proof_rewrite_rules.ML Proof/proof_syntax.ML \ + Proof/proofchecker.ML Proof/reconstruct.ML ProofGeneral/pgip.ML \ + ProofGeneral/pgip_input.ML ProofGeneral/pgip_isabelle.ML \ + ProofGeneral/pgip_markup.ML ProofGeneral/pgip_output.ML \ + ProofGeneral/pgip_parser.ML ProofGeneral/pgip_tests.ML \ + ProofGeneral/pgip_types.ML ProofGeneral/preferences.ML \ + ProofGeneral/proof_general_emacs.ML \ ProofGeneral/proof_general_pgip.ML Pure.thy ROOT.ML Syntax/ast.ML \ Syntax/lexicon.ML Syntax/mixfix.ML Syntax/parser.ML \ Syntax/printer.ML Syntax/simple_syntax.ML Syntax/syn_ext.ML \ diff -r 2b5b0f9e271c -r 1464ddca182b src/Pure/Isar/attrib.ML --- a/src/Pure/Isar/attrib.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/Pure/Isar/attrib.ML Wed Nov 04 11:40:59 2009 +0100 @@ -287,15 +287,15 @@ "rename bound variables in abstractions" #> setup (Binding.name "unfolded") unfolded "unfolded definitions" #> setup (Binding.name "folded") folded "folded definitions" #> - setup (Binding.name "consumes") (Scan.lift (Scan.optional P.nat 1) >> RuleCases.consumes) + setup (Binding.name "consumes") (Scan.lift (Scan.optional P.nat 1) >> Rule_Cases.consumes) "number of consumed facts" #> - setup (Binding.name "case_names") (Scan.lift (Scan.repeat1 Args.name) >> RuleCases.case_names) + setup (Binding.name "case_names") (Scan.lift (Scan.repeat1 Args.name) >> Rule_Cases.case_names) "named rule cases" #> setup (Binding.name "case_conclusion") - (Scan.lift (Args.name -- Scan.repeat Args.name) >> RuleCases.case_conclusion) + (Scan.lift (Args.name -- Scan.repeat Args.name) >> Rule_Cases.case_conclusion) "named conclusion of rule cases" #> setup (Binding.name "params") - (Scan.lift (P.and_list1 (Scan.repeat Args.name)) >> RuleCases.params) + (Scan.lift (P.and_list1 (Scan.repeat Args.name)) >> Rule_Cases.params) "named rule parameters" #> setup (Binding.name "standard") (Scan.succeed (Thm.rule_attribute (K Drule.standard))) "result put into standard form (legacy)" #> diff -r 2b5b0f9e271c -r 1464ddca182b src/Pure/Isar/auto_bind.ML --- a/src/Pure/Isar/auto_bind.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/Pure/Isar/auto_bind.ML Wed Nov 04 11:40:59 2009 +0100 @@ -14,7 +14,7 @@ val no_facts: (indexname * term option) list end; -structure AutoBind: AUTO_BIND = +structure Auto_Bind: AUTO_BIND = struct (** bindings **) diff -r 2b5b0f9e271c -r 1464ddca182b src/Pure/Isar/calculation.ML --- a/src/Pure/Isar/calculation.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/Pure/Isar/calculation.ML Wed Nov 04 11:40:59 2009 +0100 @@ -66,15 +66,16 @@ (* add/del rules *) -val trans_add = Thm.declaration_attribute (CalculationData.map o apfst o apfst o Item_Net.insert); -val trans_del = Thm.declaration_attribute (CalculationData.map o apfst o apfst o Item_Net.delete); +val trans_add = Thm.declaration_attribute (CalculationData.map o apfst o apfst o Item_Net.update); +val trans_del = Thm.declaration_attribute (CalculationData.map o apfst o apfst o Item_Net.remove); val sym_add = Thm.declaration_attribute (CalculationData.map o apfst o apsnd o Thm.add_thm) - #> ContextRules.elim_query NONE; + #> Context_Rules.elim_query NONE; + val sym_del = Thm.declaration_attribute (CalculationData.map o apfst o apsnd o Thm.del_thm) - #> ContextRules.rule_del; + #> Context_Rules.rule_del; (* symmetric *) @@ -130,7 +131,8 @@ fun combine ths = (case opt_rules of SOME rules => rules | NONE => - (case ths of [] => Item_Net.content (#1 (get_rules state)) + (case ths of + [] => Item_Net.content (#1 (get_rules state)) | th :: _ => Item_Net.retrieve (#1 (get_rules state)) (strip_assums_concl th))) |> Seq.of_list |> Seq.maps (Drule.multi_resolve ths) |> Seq.filter (not o projection ths); diff -r 2b5b0f9e271c -r 1464ddca182b src/Pure/Isar/code.ML --- a/src/Pure/Isar/code.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/Pure/Isar/code.ML Wed Nov 04 11:40:59 2009 +0100 @@ -639,7 +639,8 @@ (* datatypes *) -structure Type_Interpretation = InterpretationFun(type T = string * serial val eq = eq_snd (op =) : T * T -> bool); +structure Type_Interpretation = + Interpretation(type T = string * serial val eq = eq_snd (op =) : T * T -> bool); fun add_datatype raw_cs thy = let diff -r 2b5b0f9e271c -r 1464ddca182b src/Pure/Isar/constdefs.ML --- a/src/Pure/Isar/constdefs.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/Pure/Isar/constdefs.ML Wed Nov 04 11:40:59 2009 +0100 @@ -60,7 +60,7 @@ val ctxt = ProofContext.init thy; val (structs, _) = prep_vars (map (fn (x, T) => (x, T, Structure)) raw_structs) ctxt; val (decls, thy') = fold_map (gen_constdef prep_vars prep_prop prep_att structs) specs thy; - in Pretty.writeln (ProofDisplay.pretty_consts ctxt (K true) decls); thy' end; + in Pretty.writeln (Proof_Display.pretty_consts ctxt (K true) decls); thy' end; val add_constdefs = gen_constdefs ProofContext.read_vars Syntax.read_prop Attrib.attribute; val add_constdefs_i = gen_constdefs ProofContext.cert_vars ProofContext.cert_prop (K I); diff -r 2b5b0f9e271c -r 1464ddca182b src/Pure/Isar/context_rules.ML --- a/src/Pure/Isar/context_rules.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/Pure/Isar/context_rules.ML Wed Nov 04 11:40:59 2009 +0100 @@ -7,8 +7,7 @@ signature CONTEXT_RULES = sig - type netpair - type T + type netpair = ((int * int) * (bool * thm)) Net.net * ((int * int) * (bool * thm)) Net.net val netpair_bang: Proof.context -> netpair val netpair: Proof.context -> netpair val orderlist: ((int * int) * 'a) list -> 'a list @@ -33,7 +32,7 @@ attribute context_parser end; -structure ContextRules: CONTEXT_RULES = +structure Context_Rules: CONTEXT_RULES = struct @@ -65,7 +64,7 @@ type netpair = ((int * int) * (bool * thm)) Net.net * ((int * int) * (bool * thm)) Net.net; val empty_netpairs: netpair list = replicate (length rule_indexes) (Net.empty, Net.empty); -datatype T = Rules of +datatype rules = Rules of {next: int, rules: (int * ((int * bool) * thm)) list, netpairs: netpair list, @@ -92,7 +91,7 @@ structure Rules = GenericDataFun ( - type T = T; + type T = rules; val empty = make_rules ~1 [] empty_netpairs ([], []); val extend = I; diff -r 2b5b0f9e271c -r 1464ddca182b src/Pure/Isar/element.ML --- a/src/Pure/Isar/element.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/Pure/Isar/element.ML Wed Nov 04 11:40:59 2009 +0100 @@ -294,7 +294,7 @@ fun witness_local_proof after_qed cmd wit_propss goal_ctxt int = gen_witness_proof (fn after_qed' => fn propss => Proof.map_context (K goal_ctxt) - #> Proof.local_goal (ProofDisplay.print_results int) (K I) ProofContext.bind_propp_i + #> Proof.local_goal (Proof_Display.print_results int) (K I) ProofContext.bind_propp_i cmd NONE after_qed' (map (pair Thm.empty_binding) propss)) (fn wits => fn _ => after_qed wits) wit_propss []; diff -r 2b5b0f9e271c -r 1464ddca182b src/Pure/Isar/expression.ML --- a/src/Pure/Isar/expression.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/Pure/Isar/expression.ML Wed Nov 04 11:40:59 2009 +0100 @@ -611,7 +611,7 @@ else raise Match); (* define one predicate including its intro rule and axioms - - bname: predicate name + - binding: predicate name - parms: locale parameters - defs: thms representing substitutions from defines elements - ts: terms representing locale assumptions (not normalised wrt. defs) @@ -619,9 +619,9 @@ - thy: the theory *) -fun def_pred bname parms defs ts norm_ts thy = +fun def_pred binding parms defs ts norm_ts thy = let - val name = Sign.full_name thy bname; + val name = Sign.full_name thy binding; val (body, bodyT, body_eq) = atomize_spec thy norm_ts; val env = Term.add_free_names body []; @@ -639,10 +639,10 @@ val ([pred_def], defs_thy) = thy |> bodyT = propT ? Sign.add_advanced_trfuns ([], [], [aprop_tr' (length args) name], []) - |> Sign.declare_const ((bname, predT), NoSyn) |> snd + |> Sign.declare_const ((Binding.conceal binding, predT), NoSyn) |> snd |> PureThy.add_defs false - [((Binding.conceal (Binding.map_name Thm.def_name bname), Logic.mk_equals (head, body)), - [Thm.kind_internal])]; + [((Binding.conceal (Binding.map_name Thm.def_name binding), + Logic.mk_equals (head, body)), [])]; val defs_ctxt = ProofContext.init defs_thy |> Variable.declare_term head; val cert = Thm.cterm_of defs_thy; @@ -667,7 +667,7 @@ (* main predicate definition function *) -fun define_preds pname parms (((exts, exts'), (ints, ints')), defs) thy = +fun define_preds binding parms (((exts, exts'), (ints, ints')), defs) thy = let val defs' = map (cterm_of thy #> Assumption.assume #> Drule.abs_def) defs; @@ -675,13 +675,13 @@ if null exts then (NONE, NONE, [], thy) else let - val aname = if null ints then pname else Binding.suffix_name ("_" ^ axiomsN) pname; + val abinding = if null ints then binding else Binding.suffix_name ("_" ^ axiomsN) binding; val ((statement, intro, axioms), thy') = thy - |> def_pred aname parms defs' exts exts'; + |> def_pred abinding parms defs' exts exts'; val (_, thy'') = thy' - |> Sign.mandatory_path (Binding.name_of aname) + |> Sign.mandatory_path (Binding.name_of abinding) |> PureThy.note_thmss Thm.internalK [((Binding.conceal (Binding.name introN), []), [([intro], [Locale.unfold_add])])] ||> Sign.restore_naming thy'; @@ -692,10 +692,10 @@ let val ((statement, intro, axioms), thy''') = thy'' - |> def_pred pname parms defs' (ints @ the_list a_pred) (ints' @ the_list a_pred); + |> def_pred binding parms defs' (ints @ the_list a_pred) (ints' @ the_list a_pred); val (_, thy'''') = thy''' - |> Sign.mandatory_path (Binding.name_of pname) + |> Sign.mandatory_path (Binding.name_of binding) |> PureThy.note_thmss Thm.internalK [((Binding.conceal (Binding.name introN), []), [([intro], [Locale.intro_add])]), ((Binding.conceal (Binding.name axiomsN), []), @@ -723,9 +723,9 @@ | defines_to_notes _ e = e; fun gen_add_locale prep_decl - bname raw_predicate_bname raw_import raw_body thy = + binding raw_predicate_binding raw_import raw_body thy = let - val name = Sign.full_name thy bname; + val name = Sign.full_name thy binding; val _ = Locale.defined thy name andalso error ("Duplicate definition of locale " ^ quote name); @@ -733,17 +733,17 @@ prep_decl raw_import I raw_body (ProofContext.init thy); val text as (((_, exts'), _), defs) = eval ctxt' deps body_elems; - val predicate_bname = - if Binding.is_empty raw_predicate_bname then bname - else raw_predicate_bname; + val predicate_binding = + if Binding.is_empty raw_predicate_binding then binding + else raw_predicate_binding; val ((a_statement, a_intro, a_axioms), (b_statement, b_intro, b_axioms), thy') = - define_preds predicate_bname parms text thy; + define_preds predicate_binding parms text thy; val extraTs = subtract (op =) (fold Term.add_tfreesT (map snd parms) []) (fold Term.add_tfrees exts' []); val _ = if null extraTs then () else warning ("Additional type variable(s) in locale specification " ^ - quote (Binding.str_of bname)); + quote (Binding.str_of binding)); val a_satisfy = Element.satisfy_morphism a_axioms; val b_satisfy = Element.satisfy_morphism b_axioms; @@ -755,7 +755,7 @@ val notes = if is_some asm then - [(Thm.internalK, [((Binding.conceal (Binding.suffix_name ("_" ^ axiomsN) bname), []), + [(Thm.internalK, [((Binding.conceal (Binding.suffix_name ("_" ^ axiomsN) binding), []), [([Assumption.assume (cterm_of thy' (the asm))], [(Attrib.internal o K) Locale.witness_add])])])] else []; @@ -772,7 +772,7 @@ val axioms = map Element.conclude_witness b_axioms; val loc_ctxt = thy' - |> Locale.register_locale bname (extraTs, params) + |> Locale.register_locale binding (extraTs, params) (asm, rev defs) (a_intro, b_intro) axioms ([], []) (rev notes) (rev deps') |> TheoryTarget.init (SOME name) |> fold (fn (kind, facts) => LocalTheory.notes kind facts #> snd) notes'; diff -r 2b5b0f9e271c -r 1464ddca182b src/Pure/Isar/isar_cmd.ML --- a/src/Pure/Isar/isar_cmd.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/Pure/Isar/isar_cmd.ML Wed Nov 04 11:40:59 2009 +0100 @@ -324,7 +324,7 @@ val print_context = Toplevel.keep Toplevel.print_state_context; fun print_theory verbose = Toplevel.unknown_theory o - Toplevel.keep (Pretty.writeln o ProofDisplay.pretty_full_theory verbose o Toplevel.theory_of); + Toplevel.keep (Pretty.writeln o Proof_Display.pretty_full_theory verbose o Toplevel.theory_of); val print_syntax = Toplevel.unknown_context o Toplevel.keep (ProofContext.print_syntax o Toplevel.context_of); @@ -346,8 +346,8 @@ val print_theorems_theory = Toplevel.keep (fn state => Toplevel.theory_of state |> (case Toplevel.previous_context_of state of - SOME prev => ProofDisplay.print_theorems_diff (ProofContext.theory_of prev) - | NONE => ProofDisplay.print_theorems)); + SOME prev => Proof_Display.print_theorems_diff (ProofContext.theory_of prev) + | NONE => Proof_Display.print_theorems)); val print_theorems = Toplevel.unknown_context o print_theorems_theory o print_theorems_proof; @@ -371,7 +371,7 @@ in Pretty.writeln (Simplifier.pretty_ss ctxt (simpset_of ctxt)) end); val print_rules = Toplevel.unknown_context o - Toplevel.keep (ContextRules.print_rules o Toplevel.context_of); + Toplevel.keep (Context_Rules.print_rules o Toplevel.context_of); val print_trans_rules = Toplevel.unknown_context o Toplevel.keep (Calculation.print_rules o Toplevel.context_of); @@ -408,7 +408,7 @@ in Present.display_graph gr end); fun thm_deps args = Toplevel.unknown_theory o Toplevel.keep (fn state => - ThmDeps.thm_deps (Proof.get_thmss (Toplevel.enter_proof_body state) args)); + Thm_Deps.thm_deps (Proof.get_thmss (Toplevel.enter_proof_body state) args)); (* find unused theorems *) @@ -419,7 +419,7 @@ val ctxt = Toplevel.context_of state; fun pretty_thm (a, th) = ProofContext.pretty_fact ctxt (a, [th]); in - ThmDeps.unused_thms + Thm_Deps.unused_thms (case opt_range of NONE => (Theory.parents_of thy, [thy]) | SOME (xs, NONE) => (map ThyInfo.get_theory xs, [thy]) @@ -460,11 +460,11 @@ val prop = Thm.full_prop_of thm; val prf' = Proofterm.rewrite_proof_notypes ([], []) prf; in - ProofSyntax.pretty_proof ctxt + Proof_Syntax.pretty_proof ctxt (if full then Reconstruct.reconstruct_proof thy prop prf' else prf') end | SOME args => Pretty.chunks - (map (ProofSyntax.pretty_proof_of (Proof.context_of state) full) + (map (Proof_Syntax.pretty_proof_of (Proof.context_of state) full) (Proof.get_thmss state args))); fun string_of_prop state s = diff -r 2b5b0f9e271c -r 1464ddca182b src/Pure/Isar/isar_syn.ML --- a/src/Pure/Isar/isar_syn.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/Pure/Isar/isar_syn.ML Wed Nov 04 11:40:59 2009 +0100 @@ -685,7 +685,7 @@ OuterSyntax.command "proof" "backward proof" (K.tag_proof K.prf_block) (Scan.option Method.parse >> (fn m => Toplevel.print o - Toplevel.actual_proof (ProofNode.applys (Proof.proof m)) o + Toplevel.actual_proof (Proof_Node.applys (Proof.proof m)) o Toplevel.skip_proof (fn i => i + 1))); @@ -720,7 +720,7 @@ val _ = OuterSyntax.command "back" "backtracking of proof command" (K.tag_proof K.prf_script) - (Scan.succeed (Toplevel.print o Toplevel.actual_proof ProofNode.back o Toplevel.skip_proof I)); + (Scan.succeed (Toplevel.print o Toplevel.actual_proof Proof_Node.back o Toplevel.skip_proof I)); (* nested commands *) diff -r 2b5b0f9e271c -r 1464ddca182b src/Pure/Isar/local_defs.ML --- a/src/Pure/Isar/local_defs.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/Pure/Isar/local_defs.ML Wed Nov 04 11:40:59 2009 +0100 @@ -47,11 +47,11 @@ quote (Syntax.string_of_term ctxt eq)); val ((lhs, _), eq') = eq |> Sign.no_vars (Syntax.pp ctxt) - |> PrimitiveDefs.dest_def ctxt Term.is_Free (Variable.is_fixed ctxt) (K true) + |> Primitive_Defs.dest_def ctxt Term.is_Free (Variable.is_fixed ctxt) (K true) handle TERM (msg, _) => err msg | ERROR msg => err msg; in (Term.dest_Free (Term.head_of lhs), eq') end; -val abs_def = PrimitiveDefs.abs_def #>> Term.dest_Free; +val abs_def = Primitive_Defs.abs_def #>> Term.dest_Free; fun mk_def ctxt args = let diff -r 2b5b0f9e271c -r 1464ddca182b src/Pure/Isar/local_syntax.ML --- a/src/Pure/Isar/local_syntax.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/Pure/Isar/local_syntax.ML Wed Nov 04 11:40:59 2009 +0100 @@ -21,7 +21,7 @@ val extern_term: T -> term -> term end; -structure LocalSyntax: LOCAL_SYNTAX = +structure Local_Syntax: LOCAL_SYNTAX = struct (* datatype T *) diff -r 2b5b0f9e271c -r 1464ddca182b src/Pure/Isar/local_theory.ML --- a/src/Pure/Isar/local_theory.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/Pure/Isar/local_theory.ML Wed Nov 04 11:40:59 2009 +0100 @@ -150,9 +150,9 @@ fun target_result f lthy = let val (res, ctxt') = target_of lthy - |> ContextPosition.set_visible false + |> Context_Position.set_visible false |> f - ||> ContextPosition.restore_visible lthy; + ||> Context_Position.restore_visible lthy; val thy' = ProofContext.theory_of ctxt'; val lthy' = lthy |> map_target (K ctxt') diff -r 2b5b0f9e271c -r 1464ddca182b src/Pure/Isar/method.ML --- a/src/Pure/Isar/method.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/Pure/Isar/method.ML Wed Nov 04 11:40:59 2009 +0100 @@ -240,7 +240,7 @@ let val rules = if not (null arg_rules) then arg_rules - else flat (ContextRules.find_rules false facts goal ctxt) + else flat (Context_Rules.find_rules false facts goal ctxt) in trace ctxt rules; tac rules facts i end); fun meth tac x = METHOD (HEADGOAL o tac x); diff -r 2b5b0f9e271c -r 1464ddca182b src/Pure/Isar/obtain.ML --- a/src/Pure/Isar/obtain.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/Pure/Isar/obtain.ML Wed Nov 04 11:40:59 2009 +0100 @@ -129,7 +129,7 @@ val _ = Variable.warn_extra_tfrees fix_ctxt asms_ctxt; (*obtain statements*) - val thesisN = Name.variant xs AutoBind.thesisN; + val thesisN = Name.variant xs Auto_Bind.thesisN; val (thesis_var, thesis) = #1 (bind_judgment fix_ctxt thesisN); val asm_frees = fold Term.add_frees asm_props []; @@ -142,7 +142,7 @@ Term.list_all_free (parms, Logic.list_implies (asm_props, thesis)) |> Library.curry Logic.list_rename_params xs; val obtain_prop = - Logic.list_rename_params ([AutoBind.thesisN], + Logic.list_rename_params ([Auto_Bind.thesisN], Term.list_all_free ([thesis_var], Logic.mk_implies (that_prop, thesis))); fun after_qed _ = @@ -157,7 +157,7 @@ |> Proof.proof (SOME Method.succeed_text) |> Seq.hd |> Proof.fix_i [(Binding.name thesisN, NONE, NoSyn)] |> Proof.assume_i - [((Binding.name that_name, [ContextRules.intro_query NONE]), [(that_prop, [])])] + [((Binding.name that_name, [Context_Rules.intro_query NONE]), [(that_prop, [])])] |> `Proof.the_facts ||> Proof.chain_facts chain_facts ||> Proof.show_i NONE after_qed [(Thm.empty_binding, [(thesis, [])])] false @@ -189,7 +189,7 @@ val thy = ProofContext.theory_of ctxt; val cert = Thm.cterm_of thy; - val ((thesis_var, thesis), thesis_ctxt) = bind_judgment ctxt AutoBind.thesisN; + val ((thesis_var, thesis), thesis_ctxt) = bind_judgment ctxt Auto_Bind.thesisN; val rule = (case SINGLE (Method.insert_tac facts 1 THEN tac thesis_ctxt) (Goal.init (cert thesis)) of NONE => raise THM ("Obtain.result: tactic failed", 0, facts) @@ -222,7 +222,7 @@ val maxidx = fold (Term.maxidx_typ o snd o fst) vars ~1; val rule = Thm.incr_indexes (maxidx + 1) raw_rule; - val params = RuleCases.strip_params (Logic.nth_prem (1, Thm.prop_of rule)); + val params = Rule_Cases.strip_params (Logic.nth_prem (1, Thm.prop_of rule)); val m = length vars; val n = length params; val _ = m <= n orelse err "More variables than parameters in obtained rule" rule; @@ -273,7 +273,7 @@ val ctxt = Proof.context_of state; val chain_facts = if can Proof.assert_chain state then Proof.the_facts state else []; - val (thesis_var, thesis) = #1 (bind_judgment ctxt AutoBind.thesisN); + val (thesis_var, thesis) = #1 (bind_judgment ctxt Auto_Bind.thesisN); val vars = ctxt |> prep_vars raw_vars |-> fold_map inferred_type |> fst |> polymorphic ctxt; fun guess_context raw_rule state' = @@ -293,12 +293,12 @@ |> Proof.fix_i (map (fn ((x, T), mx) => (Binding.name x, SOME T, mx)) parms) |> `Proof.context_of |-> (fn fix_ctxt => Proof.assm_i (obtain_export fix_ctxt rule (map cert ts)) [(Thm.empty_binding, asms)]) - |> Proof.bind_terms AutoBind.no_facts + |> Proof.bind_terms Auto_Bind.no_facts end; val goal = Var (("guess", 0), propT); fun print_result ctxt' (k, [(s, [_, th])]) = - ProofDisplay.print_results int ctxt' (k, [(s, [th])]); + Proof_Display.print_results int ctxt' (k, [(s, [th])]); val before_qed = SOME (Method.primitive_text (Goal.conclude #> MetaSimplifier.norm_hhf #> (fn th => Goal.protect (Conjunction.intr (Drule.mk_term (Thm.cprop_of th)) th)))); fun after_qed [[_, res]] = @@ -307,7 +307,7 @@ state |> Proof.enter_forward |> Proof.begin_block - |> Proof.fix_i [(Binding.name AutoBind.thesisN, NONE, NoSyn)] + |> Proof.fix_i [(Binding.name Auto_Bind.thesisN, NONE, NoSyn)] |> Proof.chain_facts chain_facts |> Proof.local_goal print_result (K I) (apsnd (rpair I)) "guess" before_qed after_qed [(Thm.empty_binding, [Logic.mk_term goal, goal])] diff -r 2b5b0f9e271c -r 1464ddca182b src/Pure/Isar/proof.ML --- a/src/Pure/Isar/proof.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/Pure/Isar/proof.ML Wed Nov 04 11:40:59 2009 +0100 @@ -225,7 +225,7 @@ fun put_facts facts = map_current (fn (ctxt, _, mode, goal) => (ctxt, facts, mode, goal)) #> - put_thms true (AutoBind.thisN, facts); + put_thms true (Auto_Bind.thisN, facts); fun these_factss more_facts (named_factss, state) = (named_factss, state |> put_facts (SOME (maps snd named_factss @ more_facts))); @@ -387,7 +387,7 @@ fun no_goal_cases st = map (rpair NONE) (goals st); fun goal_cases st = - RuleCases.make_common true (Thm.theory_of_thm st, Thm.prop_of st) (map (rpair []) (goals st)); + Rule_Cases.make_common true (Thm.theory_of_thm st, Thm.prop_of st) (map (rpair []) (goals st)); fun apply_method current_context meth state = let @@ -710,7 +710,7 @@ in state' |> assume_i assumptions - |> bind_terms AutoBind.no_facts + |> bind_terms Auto_Bind.no_facts |> `the_facts |-> (fn thms => note_thmss_i [((Binding.name name, []), [(thms, [])])]) end; @@ -941,7 +941,7 @@ local fun gen_have prep_att prepp before_qed after_qed stmt int = - local_goal (ProofDisplay.print_results int) prep_att prepp "have" before_qed after_qed stmt; + local_goal (Proof_Display.print_results int) prep_att prepp "have" before_qed after_qed stmt; fun gen_show prep_att prepp before_qed after_qed stmt int state = let @@ -949,14 +949,14 @@ val rule = Unsynchronized.ref (NONE: thm option); fun fail_msg ctxt = "Local statement will fail to refine any pending goal" :: - (case ! rule of NONE => [] | SOME th => [ProofDisplay.string_of_rule ctxt "Failed" th]) + (case ! rule of NONE => [] | SOME th => [Proof_Display.string_of_rule ctxt "Failed" th]) |> cat_lines; fun print_results ctxt res = - if ! testing then () else ProofDisplay.print_results int ctxt res; + if ! testing then () else Proof_Display.print_results int ctxt res; fun print_rule ctxt th = if ! testing then rule := SOME th - else if int then priority (ProofDisplay.string_of_rule ctxt "Successful" th) + else if int then priority (Proof_Display.string_of_rule ctxt "Successful" th) else (); val test_proof = try (local_skip_proof true) @@ -1026,10 +1026,10 @@ val goal' = Thm.adjust_maxidx_thm (Thm.maxidx_of goal) (Drule.comp_no_flatten (goal, Thm.nprems_of goal) 1 Drule.protectI); - fun after_local' [[th]] = put_thms false (AutoBind.thisN, SOME [th]); - fun after_global' [[th]] = ProofContext.put_thms false (AutoBind.thisN, SOME [th]); + fun after_local' [[th]] = put_thms false (Auto_Bind.thisN, SOME [th]); + fun after_global' [[th]] = ProofContext.put_thms false (Auto_Bind.thisN, SOME [th]); val after_qed' = (after_local', after_global'); - val this_name = ProofContext.full_name goal_ctxt (Binding.name AutoBind.thisN); + val this_name = ProofContext.full_name goal_ctxt (Binding.name Auto_Bind.thisN); val result_ctxt = state diff -r 2b5b0f9e271c -r 1464ddca182b src/Pure/Isar/proof_context.ML --- a/src/Pure/Isar/proof_context.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/Pure/Isar/proof_context.ML Wed Nov 04 11:40:59 2009 +0100 @@ -112,9 +112,9 @@ val add_assms_i: Assumption.export -> (Thm.binding * (term * term list) list) list -> Proof.context -> (string * thm list) list * Proof.context - val add_cases: bool -> (string * RuleCases.T option) list -> Proof.context -> Proof.context - val apply_case: RuleCases.T -> Proof.context -> (string * term list) list * Proof.context - val get_case: Proof.context -> string -> string option list -> RuleCases.T + val add_cases: bool -> (string * Rule_Cases.T option) list -> Proof.context -> Proof.context + val apply_case: Rule_Cases.T -> Proof.context -> (string * term list) list * Proof.context + val get_case: Proof.context -> string -> string option list -> Rule_Cases.T val notation: bool -> Syntax.mode -> (term * mixfix) list -> Proof.context -> Proof.context val target_notation: bool -> Syntax.mode -> (term * mixfix) list -> morphism -> Context.generic -> Context.generic @@ -166,10 +166,10 @@ Ctxt of {mode: mode, (*inner syntax mode*) naming: Name_Space.naming, (*local naming conventions*) - syntax: LocalSyntax.T, (*local syntax*) + syntax: Local_Syntax.T, (*local syntax*) consts: Consts.T * Consts.T, (*local/global consts*) facts: Facts.T, (*local facts*) - cases: (string * (RuleCases.T * bool)) list}; (*named case contexts*) + cases: (string * (Rule_Cases.T * bool)) list}; (*named case contexts*) fun make_ctxt (mode, naming, syntax, consts, facts, cases) = Ctxt {mode = mode, naming = naming, syntax = syntax, @@ -181,7 +181,7 @@ ( type T = ctxt; fun init thy = - make_ctxt (mode_default, local_naming, LocalSyntax.init thy, + make_ctxt (mode_default, local_naming, Local_Syntax.init thy, (Sign.consts_of thy, Sign.consts_of thy), Facts.empty, []); ); @@ -230,9 +230,9 @@ val full_name = Name_Space.full_name o naming_of; val syntax_of = #syntax o rep_context; -val syn_of = LocalSyntax.syn_of o syntax_of; -val set_syntax_mode = map_syntax o LocalSyntax.set_mode; -val restore_syntax_mode = map_syntax o LocalSyntax.restore_mode o syntax_of; +val syn_of = Local_Syntax.syn_of o syntax_of; +val set_syntax_mode = map_syntax o Local_Syntax.set_mode; +val restore_syntax_mode = map_syntax o Local_Syntax.restore_mode o syntax_of; val consts_of = #1 o #consts o rep_context; val const_syntax_name = Consts.syntax_name o consts_of; @@ -247,7 +247,7 @@ (* theory transfer *) fun transfer_syntax thy = - map_syntax (LocalSyntax.rebuild thy) #> + map_syntax (Local_Syntax.rebuild thy) #> map_consts (fn consts as (local_consts, global_consts) => let val thy_consts = Sign.consts_of thy in if Consts.eq_consts (thy_consts, global_consts) then consts @@ -710,8 +710,8 @@ in t |> Sign.extern_term (Consts.extern_early consts) thy - |> LocalSyntax.extern_term syntax - |> Syntax.standard_unparse_term (Consts.extern consts) ctxt (LocalSyntax.syn_of syntax) + |> Local_Syntax.extern_term syntax + |> Syntax.standard_unparse_term (Consts.extern consts) ctxt (Local_Syntax.syn_of syntax) (not (PureThy.old_appl_syntax thy)) end; @@ -773,8 +773,8 @@ fun auto_bind f ts ctxt = ctxt |> bind_terms (map drop_schematic (f (theory_of ctxt) ts)); -val auto_bind_goal = auto_bind AutoBind.goal; -val auto_bind_facts = auto_bind AutoBind.facts; +val auto_bind_goal = auto_bind Auto_Bind.goal; +val auto_bind_facts = auto_bind Auto_Bind.facts; (* match_bind(_i) *) @@ -932,7 +932,7 @@ let val pos = Binding.pos_of b; val name = full_name ctxt b; - val _ = ContextPosition.report_visible ctxt (Markup.local_fact_decl name) pos; + val _ = Context_Position.report_visible ctxt (Markup.local_fact_decl name) pos; val facts = PureThy.name_thmss false pos name raw_facts; fun app (th, attrs) x = @@ -945,9 +945,9 @@ fun put_thms do_props thms ctxt = ctxt |> map_naming (K local_naming) - |> ContextPosition.set_visible false + |> Context_Position.set_visible false |> update_thms do_props (apfst Binding.name thms) - |> ContextPosition.restore_visible ctxt + |> Context_Position.restore_visible ctxt |> restore_naming ctxt; end; @@ -1029,7 +1029,7 @@ fun notation add mode args ctxt = ctxt |> map_syntax - (LocalSyntax.update_modesyntax (theory_of ctxt) add mode (map_filter (const_syntax ctxt) args)); + (Local_Syntax.update_modesyntax (theory_of ctxt) add mode (map_filter (const_syntax ctxt) args)); fun target_notation add mode args phi = let val args' = filter (fn (t, _) => Type.similar_types (t, Morphism.term phi t)) args; @@ -1083,9 +1083,9 @@ val ctxt'' = ctxt' |> fold_map declare_var (map2 (fn x' => fn (_, T, mx) => (x', T, mx)) xs' vars) - |-> (map_syntax o LocalSyntax.add_syntax (theory_of ctxt) o map prep_mixfix); + |-> (map_syntax o Local_Syntax.add_syntax (theory_of ctxt) o map prep_mixfix); val _ = (vars ~~ xs') |> List.app (fn ((b, _, _), x') => - ContextPosition.report_visible ctxt (Markup.fixed_decl x') (Binding.pos_of b)); + Context_Position.report_visible ctxt (Markup.fixed_decl x') (Binding.pos_of b)); in (xs', ctxt'') end; end; @@ -1150,13 +1150,13 @@ fun replace (opt_x :: xs) ((y, T) :: ys) = (the_default y opt_x, T) :: replace xs ys | replace [] ys = ys | replace (_ :: _) [] = error ("Too many parameters for case " ^ quote name); - val RuleCases.Case {fixes, assumes, binds, cases} = c; + val Rule_Cases.Case {fixes, assumes, binds, cases} = c; val fixes' = replace fxs fixes; val binds' = map drop_schematic binds; in if null (fold (Term.add_tvarsT o snd) fixes []) andalso null (fold (fold Term.add_vars o snd) assumes []) then - RuleCases.Case {fixes = fixes', assumes = assumes, binds = binds', cases = cases} + Rule_Cases.Case {fixes = fixes', assumes = assumes, binds = binds', cases = cases} else error ("Illegal schematic variable(s) in case " ^ quote name) end; @@ -1172,9 +1172,9 @@ fun case_result c ctxt = let - val RuleCases.Case {fixes, ...} = c; + val Rule_Cases.Case {fixes, ...} = c; val (ts, ctxt') = ctxt |> fold_map fix fixes; - val RuleCases.Case {assumes, binds, cases, ...} = RuleCases.apply ts c; + val Rule_Cases.Case {assumes, binds, cases, ...} = Rule_Cases.apply ts c; in ctxt' |> bind_terms (map drop_schematic binds) @@ -1292,7 +1292,7 @@ fun pretty_cases ctxt = let fun add_case (_, (_, false)) = I - | add_case (name, (c as RuleCases.Case {fixes, ...}, true)) = + | add_case (name, (c as Rule_Cases.Case {fixes, ...}, true)) = cons (name, (fixes, case_result c ctxt)); val cases = fold add_case (cases_of ctxt) []; in @@ -1316,7 +1316,7 @@ val prt_term = Syntax.pretty_term ctxt; (*structures*) - val structs = LocalSyntax.structs_of (syntax_of ctxt); + val structs = Local_Syntax.structs_of (syntax_of ctxt); val prt_structs = if null structs then [] else [Pretty.block (Pretty.str "structures:" :: Pretty.brk 1 :: Pretty.commas (map Pretty.str structs))]; diff -r 2b5b0f9e271c -r 1464ddca182b src/Pure/Isar/proof_display.ML --- a/src/Pure/Isar/proof_display.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/Pure/Isar/proof_display.ML Wed Nov 04 11:40:59 2009 +0100 @@ -21,7 +21,7 @@ val pretty_consts: Proof.context -> (string * typ -> bool) -> (string * typ) list -> Pretty.T end -structure ProofDisplay: PROOF_DISPLAY = +structure Proof_Display: PROOF_DISPLAY = struct (* toplevel pretty printing *) diff -r 2b5b0f9e271c -r 1464ddca182b src/Pure/Isar/proof_node.ML --- a/src/Pure/Isar/proof_node.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/Pure/Isar/proof_node.ML Wed Nov 04 11:40:59 2009 +0100 @@ -15,36 +15,36 @@ val apply: (Proof.state -> Proof.state) -> T -> T end; -structure ProofNode: PROOF_NODE = +structure Proof_Node: PROOF_NODE = struct (* datatype *) -datatype T = ProofNode of +datatype T = Proof_Node of (Proof.state * (*first result*) Proof.state Seq.seq) * (*alternative results*) int; (*linear proof position*) -fun init st = ProofNode ((st, Seq.empty), 0); +fun init st = Proof_Node ((st, Seq.empty), 0); -fun current (ProofNode ((st, _), _)) = st; -fun position (ProofNode (_, n)) = n; +fun current (Proof_Node ((st, _), _)) = st; +fun position (Proof_Node (_, n)) = n; (* backtracking *) -fun back (ProofNode ((_, stq), n)) = +fun back (Proof_Node ((_, stq), n)) = (case Seq.pull stq of NONE => error "back: no alternatives" - | SOME res => ProofNode (res, n)); + | SOME res => Proof_Node (res, n)); (* apply transformer *) -fun applys f (ProofNode ((st, _), n)) = +fun applys f (Proof_Node ((st, _), n)) = (case Seq.pull (f st) of NONE => error "empty result sequence -- proof command failed" - | SOME res => ProofNode (res, n + 1)); + | SOME res => Proof_Node (res, n + 1)); fun apply f = applys (Seq.single o f); diff -r 2b5b0f9e271c -r 1464ddca182b src/Pure/Isar/rule_cases.ML --- a/src/Pure/Isar/rule_cases.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/Pure/Isar/rule_cases.ML Wed Nov 04 11:40:59 2009 +0100 @@ -37,6 +37,8 @@ val name: string list -> thm -> thm val case_names: string list -> attribute val case_conclusion: string * string list -> attribute + val is_inner_rule: thm -> bool + val inner_rule: attribute val save: thm -> thm -> thm val get: thm -> (string * string list) list * int val rename_params: string list list -> thm -> thm @@ -45,7 +47,7 @@ val strict_mutual_rule: Proof.context -> thm list -> int list * thm end; -structure RuleCases: RULE_CASES = +structure Rule_Cases: RULE_CASES = struct (** cases **) @@ -90,7 +92,7 @@ fun extract_case is_open thy (case_outline, raw_prop) name concls = let - val rename = if is_open then I else (apfst (Name.internal o Name.clean)); + val rename = if is_open then I else apfst (Name.internal o Name.clean); val props = Logic.dest_conjunctions (Drule.norm_hhf thy raw_prop); val len = length props; @@ -212,7 +214,7 @@ val consumes_tagN = "consumes"; fun lookup_consumes th = - (case AList.lookup (op =) (Thm.get_tags th) (consumes_tagN) of + (case AList.lookup (op =) (Thm.get_tags th) consumes_tagN of NONE => NONE | SOME s => (case Lexicon.read_nat s of SOME n => SOME n @@ -223,14 +225,13 @@ fun put_consumes NONE th = th | put_consumes (SOME n) th = th |> Thm.untag_rule consumes_tagN - |> Thm.tag_rule - (consumes_tagN, Library.string_of_int (if n < 0 then Thm.nprems_of th + n else n)); + |> Thm.tag_rule (consumes_tagN, string_of_int (if n < 0 then Thm.nprems_of th + n else n)); fun add_consumes k th = put_consumes (SOME (k + get_consumes th)) th; val save_consumes = put_consumes o lookup_consumes; -fun consumes n x = Thm.rule_attribute (K (put_consumes (SOME n))) x; +fun consumes n = Thm.rule_attribute (K (put_consumes (SOME n))); fun consumes_default n x = if is_some (lookup_consumes (#2 x)) then x else consumes n x; @@ -282,7 +283,24 @@ else NONE) in fold add_case_concl concls end; -fun case_conclusion concl = Thm.rule_attribute (fn _ => add_case_concl concl); +fun case_conclusion concl = Thm.rule_attribute (K (add_case_concl concl)); + + + +(** inner rule **) + +val inner_rule_tagN = "inner_rule"; + +fun is_inner_rule th = + AList.defined (op =) (Thm.get_tags th) inner_rule_tagN; + +fun put_inner_rule inner = + Thm.untag_rule inner_rule_tagN + #> inner ? Thm.tag_rule (inner_rule_tagN, ""); + +val save_inner_rule = put_inner_rule o is_inner_rule; + +val inner_rule = Thm.rule_attribute (K (put_inner_rule true)); @@ -290,7 +308,11 @@ (* access hints *) -fun save th = save_consumes th #> save_case_names th #> save_case_concls th; +fun save th = + save_consumes th #> + save_case_names th #> + save_case_concls th #> + save_inner_rule th; fun get th = let @@ -357,5 +379,5 @@ end; -structure BasicRuleCases: BASIC_RULE_CASES = RuleCases; -open BasicRuleCases; +structure Basic_Rule_Cases: BASIC_RULE_CASES = Rule_Cases; +open Basic_Rule_Cases; diff -r 2b5b0f9e271c -r 1464ddca182b src/Pure/Isar/rule_insts.ML --- a/src/Pure/Isar/rule_insts.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/Pure/Isar/rule_insts.ML Wed Nov 04 11:40:59 2009 +0100 @@ -177,7 +177,7 @@ else T.assign (SOME (T.Term (the (AList.lookup (op =) term_insts xi)))) arg); in - Drule.instantiate insts thm |> RuleCases.save thm + Drule.instantiate insts thm |> Rule_Cases.save thm end; fun read_instantiate_mixed' ctxt (args, concl_args) thm = @@ -268,7 +268,7 @@ (* Separate type and term insts *) fun has_type_var ((x, _), _) = (case Symbol.explode x of "'" :: _ => true | _ => false); - val Tinsts = List.filter has_type_var insts; + val Tinsts = filter has_type_var insts; val tinsts = filter_out has_type_var insts; (* Tactic *) diff -r 2b5b0f9e271c -r 1464ddca182b src/Pure/Isar/spec_rules.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Pure/Isar/spec_rules.ML Wed Nov 04 11:40:59 2009 +0100 @@ -0,0 +1,49 @@ +(* Title: Pure/Isar/spec_rules.ML + Author: Makarius + +Rules that characterize functional/relational specifications. NB: In +the face of arbitrary morphisms, the original shape of specifications +may get transformed almost arbitrarily. +*) + +signature SPEC_RULES = +sig + datatype kind = Functional | Relational | Co_Relational + val dest: Proof.context -> (kind * (term * thm list) list) list + val dest_global: theory -> (kind * (term * thm list) list) list + val add: kind * (term * thm list) list -> local_theory -> local_theory + val add_global: kind * (term * thm list) list -> theory -> theory +end; + +structure Spec_Rules: SPEC_RULES = +struct + +(* maintain rules *) + +datatype kind = Functional | Relational | Co_Relational; + +structure Rules = GenericDataFun +( + type T = (kind * (term * thm list) list) Item_Net.T; + val empty : T = + Item_Net.init + (fn ((k1, spec1), (k2, spec2)) => k1 = k2 andalso + eq_list (fn ((t1, ths1), (t2, ths2)) => + t1 aconv t2 andalso eq_list Thm.eq_thm_prop (ths1, ths2)) (spec1, spec2)) + (map #1 o #2); + val extend = I; + fun merge _ = Item_Net.merge; +); + +val dest = Item_Net.content o Rules.get o Context.Proof; +val dest_global = Item_Net.content o Rules.get o Context.Theory; + +fun add (kind, specs) = LocalTheory.declaration + (fn phi => + let val specs' = map (fn (t, ths) => (Morphism.term phi t, Morphism.fact phi ths)) specs; + in Rules.map (Item_Net.update (kind, specs')) end); + +fun add_global entry = + Context.theory_map (Rules.map (Item_Net.update entry)); + +end; diff -r 2b5b0f9e271c -r 1464ddca182b src/Pure/Isar/specification.ML --- a/src/Pure/Isar/specification.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/Pure/Isar/specification.ML Wed Nov 04 11:40:59 2009 +0100 @@ -67,7 +67,7 @@ (* diagnostics *) fun print_consts _ _ [] = () - | print_consts ctxt pred cs = Pretty.writeln (ProofDisplay.pretty_consts ctxt pred cs); + | print_consts ctxt pred cs = Pretty.writeln (Proof_Display.pretty_consts ctxt pred cs); (* prepare specification *) @@ -267,7 +267,7 @@ ((name, map attrib atts), bs |> map (fn (b, more_atts) => (prep_fact lthy b, map attrib more_atts)))); val (res, lthy') = lthy |> LocalTheory.notes kind facts; - val _ = ProofDisplay.print_results true lthy' ((kind, ""), res); + val _ = Proof_Display.print_results true lthy' ((kind, ""), res); in (res, lthy') end; val theorems = gen_theorems (K I) (K I); @@ -298,7 +298,7 @@ val raw_propp = obtains |> map (fn (_, (_, props)) => map (rpair []) props); val (propp, elems_ctxt) = prep_stmt (elems @ constraints) raw_propp ctxt; - val thesis = ObjectLogic.fixed_judgment (ProofContext.theory_of ctxt) AutoBind.thesisN; + val thesis = ObjectLogic.fixed_judgment (ProofContext.theory_of ctxt) Auto_Bind.thesisN; fun assume_case ((name, (vars, _)), asms) ctxt' = let @@ -313,17 +313,17 @@ ctxt' |> (snd o ProofContext.add_fixes (map (fn b => (b, NONE, NoSyn)) bs)); ctxt' |> Variable.auto_fixes asm |> ProofContext.add_assms_i Assumption.assume_export - [((name, [ContextRules.intro_query NONE]), [(asm, [])])] + [((name, [Context_Rules.intro_query NONE]), [(asm, [])])] |>> (fn [(_, [th])] => th) end; val atts = map (Attrib.internal o K) - [RuleCases.consumes (~ (length obtains)), RuleCases.case_names case_names]; + [Rule_Cases.consumes (~ (length obtains)), Rule_Cases.case_names case_names]; val prems = Assumption.local_prems_of elems_ctxt ctxt; val stmt = [((Binding.empty, atts), [(thesis, [])])]; val (facts, goal_ctxt) = elems_ctxt - |> (snd o ProofContext.add_fixes [(Binding.name AutoBind.thesisN, NONE, NoSyn)]) + |> (snd o ProofContext.add_fixes [(Binding.name Auto_Bind.thesisN, NONE, NoSyn)]) |> fold_map assume_case (obtains ~~ propp) |-> (fn ths => ProofContext.note_thmss Thm.assumptionK [((Binding.name Obtain.thatN, []), [(ths, [])])] #> #2 #> pair ths); @@ -357,19 +357,19 @@ |> LocalTheory.notes kind (map2 (fn (a, _) => fn ths => (a, [(ths, [])])) stmt results') |> (fn (res, lthy') => if Binding.is_empty name andalso null atts then - (ProofDisplay.print_results true lthy' ((kind, ""), res); lthy') + (Proof_Display.print_results true lthy' ((kind, ""), res); lthy') else let val ([(res_name, _)], lthy'') = lthy' |> LocalTheory.notes kind [((name, atts), [(maps #2 res, [])])]; - val _ = ProofDisplay.print_results true lthy' ((kind, res_name), res); + val _ = Proof_Display.print_results true lthy' ((kind, res_name), res); in lthy'' end) |> after_qed results' end; in goal_ctxt |> ProofContext.note_thmss Thm.assumptionK - [((Binding.name AutoBind.assmsN, []), [(prems, [])])] + [((Binding.name Auto_Bind.assmsN, []), [(prems, [])])] |> snd |> Proof.theorem_i before_qed after_qed' (map snd stmt) |> (case facts of NONE => I | SOME ths => Proof.refine_insert ths) diff -r 2b5b0f9e271c -r 1464ddca182b src/Pure/Isar/toplevel.ML --- a/src/Pure/Isar/toplevel.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/Pure/Isar/toplevel.ML Wed Nov 04 11:40:59 2009 +0100 @@ -72,7 +72,7 @@ val proof': (bool -> Proof.state -> Proof.state) -> transition -> transition val proofs: (Proof.state -> Proof.state Seq.seq) -> transition -> transition val proof: (Proof.state -> Proof.state) -> transition -> transition - val actual_proof: (ProofNode.T -> ProofNode.T) -> transition -> transition + val actual_proof: (Proof_Node.T -> Proof_Node.T) -> transition -> transition val skip_proof: (int -> int) -> transition -> transition val skip_proof_to_theory: (int -> bool) -> transition -> transition val get_id: transition -> string option @@ -121,7 +121,7 @@ datatype node = Theory of generic_theory * Proof.context option (*theory with presentation context*) | - Proof of ProofNode.T * ((Proof.context -> generic_theory) * generic_theory) + Proof of Proof_Node.T * ((Proof.context -> generic_theory) * generic_theory) (*proof node, finish, original theory*) | SkipProof of int * (generic_theory * generic_theory); (*proof depth, resulting theory, original theory*) @@ -130,7 +130,7 @@ val proof_node = fn Proof (prf, _) => SOME prf | _ => NONE; fun cases_node f _ (Theory (gthy, _)) = f gthy - | cases_node _ g (Proof (prf, _)) = g (ProofNode.current prf) + | cases_node _ g (Proof (prf, _)) = g (Proof_Node.current prf) | cases_node f _ (SkipProof (_, (gthy, _))) = f gthy; val context_node = cases_node Context.proof_of Proof.context_of; @@ -148,7 +148,7 @@ fun level (State (NONE, _)) = 0 | level (State (SOME (Theory _, _), _)) = 0 - | level (State (SOME (Proof (prf, _), _), _)) = Proof.level (ProofNode.current prf) + | level (State (SOME (Proof (prf, _), _), _)) = Proof.level (Proof_Node.current prf) | level (State (SOME (SkipProof (d, _), _), _)) = d + 1; (*different notion of proof depth!*) fun str_of_state (State (NONE, _)) = "at top level" @@ -184,7 +184,7 @@ fun proof_position_of state = (case node_of state of - Proof (prf, _) => ProofNode.position prf + Proof (prf, _) => Proof_Node.position prf | _ => raise UNDEF); val enter_proof_body = node_case (Proof.init o Context.proof_of) Proof.enter_forward; @@ -207,7 +207,7 @@ NONE => [] | SOME (Theory (gthy, _)) => if prf_only then [] else pretty_context gthy | SOME (Proof (prf, _)) => - Proof.pretty_state (ProofNode.position prf) (ProofNode.current prf) + Proof.pretty_state (Proof_Node.position prf) (Proof_Node.current prf) | SOME (SkipProof (d, _)) => [Pretty.str ("skipped proof: depth " ^ string_of_int d)]) |> Pretty.markup_chunks Markup.state |> Pretty.writeln; @@ -450,7 +450,7 @@ fun end_proof f = transaction (fn int => (fn Proof (prf, (finish, _)) => - let val state = ProofNode.current prf in + let val state = Proof_Node.current prf in if can (Proof.assert_bottom true) state then let val ctxt' = f int state; @@ -475,7 +475,7 @@ else (); if skip andalso not schematic then SkipProof (0, (finish gthy (Proof.global_skip_proof int prf), gthy)) - else Proof (ProofNode.init prf, (finish gthy, gthy)) + else Proof (Proof_Node.init prf, (finish gthy, gthy)) end | _ => raise UNDEF)); @@ -499,12 +499,12 @@ | _ => raise UNDEF)); val present_proof = present_transaction (fn _ => - (fn Proof (prf, x) => Proof (ProofNode.apply I prf, x) + (fn Proof (prf, x) => Proof (Proof_Node.apply I prf, x) | skip as SkipProof _ => skip | _ => raise UNDEF)); fun proofs' f = transaction (fn int => - (fn Proof (prf, x) => Proof (ProofNode.applys (f int) prf, x) + (fn Proof (prf, x) => Proof (Proof_Node.applys (f int) prf, x) | skip as SkipProof _ => skip | _ => raise UNDEF)); @@ -654,7 +654,7 @@ Future.fork_pri ~1 (fn () => let val (states, result_state) = (case st' of State (SOME (Proof (_, (_, orig_gthy)), exit), prev) - => State (SOME (Proof (ProofNode.init prf, (finish, orig_gthy)), exit), prev)) + => State (SOME (Proof (Proof_Node.init prf, (finish, orig_gthy)), exit), prev)) |> fold_map command_result body_trs ||> command (end_tr |> set_print false); in (states, presentation_context_of result_state) end)) diff -r 2b5b0f9e271c -r 1464ddca182b src/Pure/Proof/extraction.ML --- a/src/Pure/Proof/extraction.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/Pure/Proof/extraction.ML Wed Nov 04 11:40:59 2009 +0100 @@ -77,12 +77,12 @@ val empty_rules : rules = {next = 0, rs = [], net = Net.empty}; -fun add_rule (r as (_, (lhs, _)), {next, rs, net} : rules) = +fun add_rule (r as (_, (lhs, _))) ({next, rs, net} : rules) = {next = next - 1, rs = r :: rs, net = Net.insert_term (K false) (Envir.eta_contract lhs, (next, r)) net}; fun merge_rules ({next, rs = rs1, net} : rules) ({rs = rs2, ...} : rules) = - List.foldr add_rule {next = next, rs = rs1, net = net} (subtract (op =) rs1 rs2); + fold_rev add_rule (subtract (op =) rs1 rs2) {next = next, rs = rs1, net = net}; fun condrew thy rules procs = let @@ -231,7 +231,7 @@ defs, expand, prep} = ExtractionData.get thy; in ExtractionData.put - {realizes_eqns = List.foldr add_rule realizes_eqns (map (prep_eq thy) eqns), + {realizes_eqns = fold_rev add_rule (map (prep_eq thy) eqns) realizes_eqns, typeof_eqns = typeof_eqns, types = types, realizers = realizers, defs = defs, expand = expand, prep = prep} thy end @@ -249,7 +249,7 @@ in ExtractionData.put {realizes_eqns = realizes_eqns, realizers = realizers, - typeof_eqns = List.foldr add_rule typeof_eqns eqns', + typeof_eqns = fold_rev add_rule eqns' typeof_eqns, types = types, defs = defs, expand = expand, prep = prep} thy end @@ -300,7 +300,7 @@ val rtypes = map fst types; val eqns = Net.merge (K false) (#net realizes_eqns, #net typeof_eqns); val thy' = add_syntax thy; - val rd = ProofSyntax.read_proof thy' false + val rd = Proof_Syntax.read_proof thy' false; in fn (thm, (vs, s1, s2)) => let val name = Thm.get_name thm; @@ -359,8 +359,8 @@ in (ExtractionData.put (if is_def then {realizes_eqns = realizes_eqns, - typeof_eqns = add_rule (([], - Logic.dest_equals (prop_of (Drule.abs_def thm))), typeof_eqns), + typeof_eqns = add_rule ([], + Logic.dest_equals (prop_of (Drule.abs_def thm))) typeof_eqns, types = types, realizers = realizers, defs = insert Thm.eq_thm thm defs, expand = expand, prep = prep} @@ -458,7 +458,7 @@ val vars = vars_of prop; val n = Int.min (length vars, length ts); - fun add_args ((Var ((a, i), _), t), (vs', tye)) = + fun add_args (Var ((a, i), _), t) (vs', tye) = if member (op =) rvs a then let val T = etype_of thy' vs Ts t in if T = nullT then (vs', tye) @@ -466,7 +466,7 @@ end else (vs', tye) - in List.foldr add_args ([], []) (Library.take (n, vars) ~~ Library.take (n, ts)) end; + in fold_rev add_args (Library.take (n, vars) ~~ Library.take (n, ts)) ([], []) end; fun find (vs: string list) = Option.map snd o find_first (curry (eq_set (op =)) vs o fst); fun find' (s: string) = map_filter (fn (s', x) => if s = s' then SOME x else NONE); @@ -651,7 +651,7 @@ val nt = Envir.beta_norm t; val args = filter_out (fn v => member (op =) rtypes (tname_of (body_type (fastype_of v)))) (vfs_of prop); - val args' = List.filter (fn v => Logic.occs (v, nt)) args; + val args' = filter (fn v => Logic.occs (v, nt)) args; val t' = mkabs nt args'; val T = fastype_of t'; val cname = extr_name s vs'; diff -r 2b5b0f9e271c -r 1464ddca182b src/Pure/Proof/proof_syntax.ML --- a/src/Pure/Proof/proof_syntax.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/Pure/Proof/proof_syntax.ML Wed Nov 04 11:40:59 2009 +0100 @@ -19,7 +19,7 @@ val pretty_proof_of: Proof.context -> bool -> thm -> Pretty.T end; -structure ProofSyntax : PROOF_SYNTAX = +structure Proof_Syntax : PROOF_SYNTAX = struct open Proofterm; diff -r 2b5b0f9e271c -r 1464ddca182b src/Pure/ProofGeneral/proof_general_pgip.ML --- a/src/Pure/ProofGeneral/proof_general_pgip.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/Pure/ProofGeneral/proof_general_pgip.ML Wed Nov 04 11:40:59 2009 +0100 @@ -658,7 +658,7 @@ fun strings_of_thm (thy, name) = map (Display.string_of_thm_global thy) (PureThy.get_thms thy name) - val string_of_thy = Pretty.string_of o ProofDisplay.pretty_full_theory false + val string_of_thy = Pretty.string_of o Proof_Display.pretty_full_theory false in case (thyname, objtype) of (_, ObjTheory) => idvalue [string_of_thy (ThyInfo.get_theory name)] diff -r 2b5b0f9e271c -r 1464ddca182b src/Pure/ROOT.ML --- a/src/Pure/ROOT.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/Pure/ROOT.ML Wed Nov 04 11:40:59 2009 +0100 @@ -212,6 +212,7 @@ (*specifications*) use "Isar/spec_parse.ML"; +use "Isar/spec_rules.ML"; use "Isar/specification.ML"; use "Isar/constdefs.ML"; diff -r 2b5b0f9e271c -r 1464ddca182b src/Pure/Syntax/parser.ML --- a/src/Pure/Syntax/parser.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/Pure/Syntax/parser.ML Wed Nov 04 11:40:59 2009 +0100 @@ -615,11 +615,11 @@ (*Get all rhss with precedence >= minPrec*) -fun getRHS minPrec = List.filter (fn (_, _, prec:int) => prec >= minPrec); +fun getRHS minPrec = filter (fn (_, _, prec:int) => prec >= minPrec); (*Get all rhss with precedence >= minPrec and < maxPrec*) fun getRHS' minPrec maxPrec = - List.filter (fn (_, _, prec:int) => prec >= minPrec andalso prec < maxPrec); + filter (fn (_, _, prec:int) => prec >= minPrec andalso prec < maxPrec); (*Make states using a list of rhss*) fun mkStates i minPrec lhsID rhss = @@ -655,19 +655,19 @@ in update (used, []) end; fun getS A maxPrec Si = - List.filter + filter (fn (_, _, _, Nonterminal (B, prec) :: _, _, _) => A = B andalso prec <= maxPrec | _ => false) Si; fun getS' A maxPrec minPrec Si = - List.filter + filter (fn (_, _, _, Nonterminal (B, prec) :: _, _, _) => A = B andalso prec > minPrec andalso prec <= maxPrec | _ => false) Si; fun getStates Estate i ii A maxPrec = - List.filter + filter (fn (_, _, _, Nonterminal (B, prec) :: _, _, _) => A = B andalso prec <= maxPrec | _ => false) diff -r 2b5b0f9e271c -r 1464ddca182b src/Pure/Syntax/simple_syntax.ML --- a/src/Pure/Syntax/simple_syntax.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/Pure/Syntax/simple_syntax.ML Wed Nov 04 11:40:59 2009 +0100 @@ -11,7 +11,7 @@ val read_prop: string -> term end; -structure SimpleSyntax: SIMPLE_SYNTAX = +structure Simple_Syntax: SIMPLE_SYNTAX = struct (* scanning tokens *) diff -r 2b5b0f9e271c -r 1464ddca182b src/Pure/Syntax/syn_ext.ML --- a/src/Pure/Syntax/syn_ext.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/Pure/Syntax/syn_ext.ML Wed Nov 04 11:40:59 2009 +0100 @@ -218,7 +218,7 @@ val read_symbs = map_filter I o the o Scan.read Symbol.stopper scan_symbs; fun unique_index xsymbs = - if length (List.filter is_index xsymbs) <= 1 then xsymbs + if length (filter is_index xsymbs) <= 1 then xsymbs else error "Duplicate index arguments (\\)"; in @@ -226,7 +226,7 @@ val read_mfix = unique_index o read_symbs o Symbol.explode; fun mfix_delims sy = fold_rev (fn Delim s => cons s | _ => I) (read_mfix sy) []; -val mfix_args = length o List.filter is_argument o read_mfix; +val mfix_args = length o filter is_argument o read_mfix; val escape_mfix = implode o map (fn s => if is_meta s then "'" ^ s else s) o Symbol.explode; @@ -276,7 +276,7 @@ val raw_symbs = read_mfix sy handle ERROR msg => err_in_mfix msg mfix; - val args = List.filter (fn Argument _ => true | _ => false) raw_symbs; + val args = filter (fn Argument _ => true | _ => false) raw_symbs; val (const', typ', parse_rules) = if not (exists is_index args) then (const, typ, []) else @@ -312,7 +312,7 @@ val xprod' = if Lexicon.is_terminal lhs' then err_in_mfix ("Illegal lhs: " ^ lhs') mfix else if const <> "" then xprod - else if length (List.filter is_argument symbs') <> 1 then + else if length (filter is_argument symbs') <> 1 then err_in_mfix "Copy production must have exactly one argument" mfix else if exists is_terminal symbs' then xprod else XProd (lhs', map rem_pri symbs', "", chain_pri); diff -r 2b5b0f9e271c -r 1464ddca182b src/Pure/Thy/thm_deps.ML --- a/src/Pure/Thy/thm_deps.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/Pure/Thy/thm_deps.ML Wed Nov 04 11:40:59 2009 +0100 @@ -10,7 +10,7 @@ val unused_thms: theory list * theory list -> (string * thm) list end; -structure ThmDeps: THM_DEPS = +structure Thm_Deps: THM_DEPS = struct (* thm_deps *) diff -r 2b5b0f9e271c -r 1464ddca182b src/Pure/Thy/thy_output.ML --- a/src/Pure/Thy/thy_output.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/Pure/Thy/thy_output.ML Wed Nov 04 11:40:59 2009 +0100 @@ -478,7 +478,7 @@ val ctxt' = Variable.auto_fixes eq ctxt; in ProofContext.pretty_term_abbrev ctxt' eq end; -fun pretty_prf full ctxt = ProofSyntax.pretty_proof_of ctxt full; +fun pretty_prf full ctxt = Proof_Syntax.pretty_proof_of ctxt full; fun pretty_theory ctxt name = (Theory.requires (ProofContext.theory_of ctxt) name "presentation"; Pretty.str name); diff -r 2b5b0f9e271c -r 1464ddca182b src/Pure/Tools/find_consts.ML --- a/src/Pure/Tools/find_consts.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/Pure/Tools/find_consts.ML Wed Nov 04 11:40:59 2009 +0100 @@ -11,11 +11,10 @@ Strict of string | Loose of string | Name of string - val find_consts : Proof.context -> (bool * criterion) list -> unit end; -structure FindConsts : FIND_CONSTS = +structure Find_Consts : FIND_CONSTS = struct (* search criteria *) @@ -162,7 +161,7 @@ val _ = OuterSyntax.improper_command "find_consts" "search constants by type pattern" K.diag - (Scan.repeat (((Scan.option P.minus >> is_none) -- criterion)) + (Scan.repeat ((Scan.option P.minus >> is_none) -- criterion) >> (Toplevel.no_timing oo find_consts_cmd)); end; diff -r 2b5b0f9e271c -r 1464ddca182b src/Pure/Tools/find_theorems.ML --- a/src/Pure/Tools/find_theorems.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/Pure/Tools/find_theorems.ML Wed Nov 04 11:40:59 2009 +0100 @@ -18,7 +18,7 @@ (bool * string criterion) list -> unit end; -structure FindTheorems: FIND_THEOREMS = +structure Find_Theorems: FIND_THEOREMS = struct (** search criteria **) @@ -28,24 +28,22 @@ Pattern of 'term; fun apply_dummies tm = - strip_abs tm - |> fst - |> map (Term.dummy_pattern o snd) - |> betapplys o pair tm - |> (fn x => Term.replace_dummy_patterns x 1) - |> fst; + let + val (xs, _) = Term.strip_abs tm; + val tm' = Term.betapplys (tm, map (Term.dummy_pattern o #2) xs); + in #1 (Term.replace_dummy_patterns tm' 1) end; fun parse_pattern ctxt nm = let - val nm' = case Syntax.parse_term ctxt nm of Const (n, _) => n | _ => nm; val consts = ProofContext.consts_of ctxt; + val nm' = + (case Syntax.parse_term ctxt nm of + Const (c, _) => c + | _ => Consts.intern consts nm); in - nm' - |> Consts.intern consts - |> Consts.the_abbreviation consts - |> snd - |> apply_dummies - handle TYPE _ => ProofContext.read_term_pattern ctxt nm + (case try (Consts.the_abbreviation consts) nm' of + SOME (_, rhs) => apply_dummies rhs + | NONE => ProofContext.read_term_pattern ctxt nm) end; fun read_criterion _ (Name name) = Name name @@ -377,9 +375,15 @@ (* print_theorems *) fun all_facts_of ctxt = - maps Facts.selections - (Facts.dest_static [] (PureThy.facts_of (ProofContext.theory_of ctxt)) @ - Facts.dest_static [] (ProofContext.facts_of ctxt)); + let + fun visible_facts facts = + Facts.dest_static [] facts + |> filter_out (Facts.is_concealed facts o #1); + in + maps Facts.selections + (visible_facts (PureThy.facts_of (ProofContext.theory_of ctxt)) @ + visible_facts (ProofContext.facts_of ctxt)) + end; val limit = Unsynchronized.ref 40; diff -r 2b5b0f9e271c -r 1464ddca182b src/Pure/Tools/named_thms.ML --- a/src/Pure/Tools/named_thms.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/Pure/Tools/named_thms.ML Wed Nov 04 11:40:59 2009 +0100 @@ -1,7 +1,8 @@ (* Title: Pure/Tools/named_thms.ML Author: Makarius -Named collections of theorems in canonical order. +Named collections of theorems in canonical order. Based on naive data +structures -- not scalable! *) signature NAMED_THMS = diff -r 2b5b0f9e271c -r 1464ddca182b src/Pure/Tools/xml_syntax.ML --- a/src/Pure/Tools/xml_syntax.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/Pure/Tools/xml_syntax.ML Wed Nov 04 11:40:59 2009 +0100 @@ -17,7 +17,7 @@ val proof_of_xml: XML.tree -> Proofterm.proof end; -structure XMLSyntax : XML_SYNTAX = +structure XML_Syntax : XML_SYNTAX = struct (**** XML output ****) diff -r 2b5b0f9e271c -r 1464ddca182b src/Pure/axclass.ML --- a/src/Pure/axclass.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/Pure/axclass.ML Wed Nov 04 11:40:59 2009 +0100 @@ -311,7 +311,7 @@ (Binding.name (Thm.def_name c'), Logic.mk_equals (Const (c, T'), const')) #>> Thm.varifyT #-> (fn thm => add_inst_param (c, tyco) (c'', thm) - #> PureThy.add_thm ((Binding.conceal (Binding.name c'), thm), [Thm.kind_internal]) + #> PureThy.add_thm ((Binding.conceal (Binding.name c'), thm), []) #> snd #> Sign.restore_naming thy #> pair (Const (c, T)))) diff -r 2b5b0f9e271c -r 1464ddca182b src/Pure/codegen.ML --- a/src/Pure/codegen.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/Pure/codegen.ML Wed Nov 04 11:40:59 2009 +0100 @@ -137,7 +137,7 @@ | args_of (Ignore :: ms) (_ :: xs) = args_of ms xs | args_of (_ :: ms) xs = args_of ms xs; -fun num_args_of x = length (List.filter is_arg x); +fun num_args_of x = length (filter is_arg x); (**** theory data ****) diff -r 2b5b0f9e271c -r 1464ddca182b src/Pure/conjunction.ML --- a/src/Pure/conjunction.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/Pure/conjunction.ML Wed Nov 04 11:40:59 2009 +0100 @@ -30,7 +30,7 @@ (** abstract syntax **) fun certify t = Thm.cterm_of (Context.the_theory (Context.the_thread_data ())) t; -val read_prop = certify o SimpleSyntax.read_prop; +val read_prop = certify o Simple_Syntax.read_prop; val true_prop = certify Logic.true_prop; val conjunction = certify Logic.conjunction; diff -r 2b5b0f9e271c -r 1464ddca182b src/Pure/consts.ML --- a/src/Pure/consts.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/Pure/consts.ML Wed Nov 04 11:40:59 2009 +0100 @@ -71,10 +71,10 @@ (* reverted abbrevs *) val empty_abbrevs = - Item_Net.init (fn ((t, u), (t', u')) => t aconv t' andalso u aconv u') #1; + Item_Net.init (fn ((t, u), (t', u')) => t aconv t' andalso u aconv u') (single o #1); -fun insert_abbrevs mode abbrs = - Symtab.map_default (mode, empty_abbrevs) (Item_Net.insert abbrs); +fun update_abbrevs mode abbrs = + Symtab.map_default (mode, empty_abbrevs) (Item_Net.update abbrs); fun retrieve_abbrevs (Consts {rev_abbrevs, ...}) modes = let val nets = map_filter (Symtab.lookup rev_abbrevs) modes @@ -290,7 +290,7 @@ val (_, decls') = decls |> Name_Space.define true naming (b, (decl, SOME abbr)); val rev_abbrevs' = rev_abbrevs - |> insert_abbrevs mode (rev_abbrev lhs rhs); + |> update_abbrevs mode (rev_abbrev lhs rhs); in (decls', constraints, rev_abbrevs') end) |> pair (lhs, rhs) end; @@ -299,7 +299,7 @@ let val (T, rhs) = the_abbreviation consts c; val rev_abbrevs' = rev_abbrevs - |> insert_abbrevs mode (rev_abbrev (Const (c, T)) rhs); + |> update_abbrevs mode (rev_abbrev (Const (c, T)) rhs); in (decls, constraints, rev_abbrevs') end); end; diff -r 2b5b0f9e271c -r 1464ddca182b src/Pure/context_position.ML --- a/src/Pure/context_position.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/Pure/context_position.ML Wed Nov 04 11:40:59 2009 +0100 @@ -12,7 +12,7 @@ val report_visible: Proof.context -> Markup.T -> Position.T -> unit end; -structure ContextPosition: CONTEXT_POSITION = +structure Context_Position: CONTEXT_POSITION = struct structure Data = ProofDataFun diff -r 2b5b0f9e271c -r 1464ddca182b src/Pure/drule.ML --- a/src/Pure/drule.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/Pure/drule.ML Wed Nov 04 11:40:59 2009 +0100 @@ -450,7 +450,7 @@ (*** Meta-Rewriting Rules ***) -val read_prop = certify o SimpleSyntax.read_prop; +val read_prop = certify o Simple_Syntax.read_prop; fun get_def thy = Thm.axiom thy o Name_Space.intern (Theory.axiom_space thy) o Thm.def_name; diff -r 2b5b0f9e271c -r 1464ddca182b src/Pure/interpretation.ML --- a/src/Pure/interpretation.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/Pure/interpretation.ML Wed Nov 04 11:40:59 2009 +0100 @@ -13,7 +13,7 @@ val init: theory -> theory end; -functor InterpretationFun(type T val eq: T * T -> bool): INTERPRETATION = +functor Interpretation(type T val eq: T * T -> bool): INTERPRETATION = struct type T = T; diff -r 2b5b0f9e271c -r 1464ddca182b src/Pure/item_net.ML --- a/src/Pure/item_net.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/Pure/item_net.ML Wed Nov 04 11:40:59 2009 +0100 @@ -8,12 +8,12 @@ signature ITEM_NET = sig type 'a T + val init: ('a * 'a -> bool) -> ('a -> term list) -> 'a T val content: 'a T -> 'a list val retrieve: 'a T -> term -> 'a list - val init: ('a * 'a -> bool) -> ('a -> term) -> 'a T val merge: 'a T * 'a T -> 'a T - val delete: 'a -> 'a T -> 'a T - val insert: 'a -> 'a T -> 'a T + val remove: 'a -> 'a T -> 'a T + val update: 'a -> 'a T -> 'a T end; structure Item_Net: ITEM_NET = @@ -24,7 +24,7 @@ datatype 'a T = Items of { eq: 'a * 'a -> bool, - index: 'a -> term, + index: 'a -> term list, content: 'a list, next: int, net: (int * 'a) Net.net}; @@ -32,27 +32,31 @@ fun mk_items eq index content next net = Items {eq = eq, index = index, content = content, next = next, net = net}; +fun init eq index = mk_items eq index [] ~1 Net.empty; + fun content (Items {content, ...}) = content; fun retrieve (Items {net, ...}) = order_list o Net.unify_term net; -(* build net *) +(* standard operations *) -fun init eq index = mk_items eq index [] ~1 Net.empty; +fun member (Items {eq, index, net, ...}) x = + exists (fn t => exists (fn (_, y) => eq (x, y)) (Net.unify_term net t)) (index x); -fun add x (Items {eq, index, content, next, net}) = +fun cons x (Items {eq, index, content, next, net}) = mk_items eq index (x :: content) (next - 1) - (Net.insert_term (K false) (index x, (next, x)) net); + (fold (fn t => Net.insert_term_safe (eq o pairself #2) (t, (next, x))) (index x) net); + -fun merge (Items {eq, index, content = content1, ...}, Items {content = content2, ...}) = - let val content = Library.merge eq (content1, content2) - in fold_rev add content (init eq index) end; +fun merge (items1, Items {content = content2, ...}) = + fold_rev (fn y => if member items1 y then I else cons y) content2 items1; -fun delete x (items as Items {eq, index, content, next, net}) = - if not (member eq content x) then items - else mk_items eq index (remove eq x content) next - (Net.delete_term (eq o pairself #2) (index x, (0, x)) net); +fun remove x (items as Items {eq, index, content, next, net}) = + if member items x then + mk_items eq index (Library.remove eq x content) next + (fold (fn t => Net.delete_term_safe (eq o pairself #2) (t, (0, x))) (index x) net) + else items; -fun insert x items = add x (delete x items); (*ensure that added entry gets precedence*) +fun update x items = cons x (remove x items); end; diff -r 2b5b0f9e271c -r 1464ddca182b src/Pure/meta_simplifier.ML diff -r 2b5b0f9e271c -r 1464ddca182b src/Pure/more_thm.ML --- a/src/Pure/more_thm.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/Pure/more_thm.ML Wed Nov 04 11:40:59 2009 +0100 @@ -91,13 +91,9 @@ val lemmaK: string val corollaryK: string val internalK: string - val has_kind: thm -> bool val get_kind: thm -> string val kind_rule: string -> thm -> thm val kind: string -> attribute - val kind_internal: attribute - val has_internal: Properties.property list -> bool - val is_internal: thm -> bool end; structure Thm: THM = @@ -250,8 +246,8 @@ val del_thm = remove eq_thm_prop; val merge_thms = merge eq_thm_prop; -val intro_rules = Item_Net.init eq_thm_prop Thm.concl_of; -val elim_rules = Item_Net.init eq_thm_prop Thm.major_prem_of; +val intro_rules = Item_Net.init eq_thm_prop (single o Thm.concl_of); +val elim_rules = Item_Net.init eq_thm_prop (single o Thm.major_prem_of); @@ -425,16 +421,10 @@ val corollaryK = "corollary"; val internalK = Markup.internalK; -fun the_kind thm = the (Properties.get (Thm.get_tags thm) Markup.kindN); - -val has_kind = can the_kind; -val get_kind = the_default "" o try the_kind; +fun get_kind thm = the_default "" (Properties.get (Thm.get_tags thm) Markup.kindN); fun kind_rule k = tag_rule (Markup.kindN, k) o untag_rule Markup.kindN; fun kind k x = if k = "" then x else rule_attribute (K (kind_rule k)) x; -fun kind_internal x = kind internalK x; -fun has_internal tags = exists (fn tg => tg = (Markup.kindN, internalK)) tags; -val is_internal = has_internal o Thm.get_tags; open Thm; diff -r 2b5b0f9e271c -r 1464ddca182b src/Pure/net.ML --- a/src/Pure/net.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/Pure/net.ML Wed Nov 04 11:40:59 2009 +0100 @@ -22,9 +22,13 @@ exception INSERT val insert: ('a * 'a -> bool) -> key list * 'a -> 'a net -> 'a net val insert_term: ('a * 'a -> bool) -> term * 'a -> 'a net -> 'a net + val insert_safe: ('a * 'a -> bool) -> key list * 'a -> 'a net -> 'a net + val insert_term_safe: ('a * 'a -> bool) -> term * 'a -> 'a net -> 'a net exception DELETE val delete: ('b * 'a -> bool) -> key list * 'b -> 'a net -> 'a net val delete_term: ('b * 'a -> bool) -> term * 'b -> 'a net -> 'a net + val delete_safe: ('b * 'a -> bool) -> key list * 'b -> 'a net -> 'a net + val delete_term_safe: ('b * 'a -> bool) -> term * 'b -> 'a net -> 'a net val lookup: 'a net -> key list -> 'a list val match_term: 'a net -> term -> 'a list val unify_term: 'a net -> term -> 'a list @@ -99,8 +103,10 @@ in Net{comb=comb, var=var, atoms=atoms'} end in ins1 (keys,net) end; +fun insert_term eq (t, x) = insert eq (key_of_term t, x); + fun insert_safe eq entry net = insert eq entry net handle INSERT => net; -fun insert_term eq (t, x) = insert eq (key_of_term t, x); +fun insert_term_safe eq entry net = insert_term eq entry net handle INSERT => net; (*** Deletion from a discrimination net ***) @@ -137,6 +143,9 @@ fun delete_term eq (t, x) = delete eq (key_of_term t, x); +fun delete_safe eq entry net = delete eq entry net handle DELETE => net; +fun delete_term_safe eq entry net = delete_term eq entry net handle DELETE => net; + (*** Retrieval functions for discrimination nets ***) diff -r 2b5b0f9e271c -r 1464ddca182b src/Pure/primitive_defs.ML --- a/src/Pure/primitive_defs.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/Pure/primitive_defs.ML Wed Nov 04 11:40:59 2009 +0100 @@ -12,7 +12,7 @@ val mk_defpair: term * term -> string * term end; -structure PrimitiveDefs: PRIMITIVE_DEFS = +structure Primitive_Defs: PRIMITIVE_DEFS = struct fun term_kind (Const _) = "existing constant " diff -r 2b5b0f9e271c -r 1464ddca182b src/Pure/proofterm.ML --- a/src/Pure/proofterm.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/Pure/proofterm.ML Wed Nov 04 11:40:59 2009 +0100 @@ -1024,7 +1024,7 @@ (** see pattern.ML **) -fun flt (i: int) = List.filter (fn n => n < i); +fun flt (i: int) = filter (fn n => n < i); fun fomatch Ts tymatch j = let diff -r 2b5b0f9e271c -r 1464ddca182b src/Pure/pure_setup.ML --- a/src/Pure/pure_setup.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/Pure/pure_setup.ML Wed Nov 04 11:40:59 2009 +0100 @@ -22,13 +22,13 @@ toplevel_pp ["Task_Queue", "group"] "Pretty.str o Task_Queue.str_of_group"; toplevel_pp ["Position", "T"] "Pretty.position"; toplevel_pp ["Binding", "binding"] "Pretty.str o quote o Binding.str_of"; -toplevel_pp ["Thm", "thm"] "ProofDisplay.pp_thm"; -toplevel_pp ["Thm", "cterm"] "ProofDisplay.pp_cterm"; -toplevel_pp ["Thm", "ctyp"] "ProofDisplay.pp_ctyp"; -toplevel_pp ["typ"] "ProofDisplay.pp_typ Pure.thy"; +toplevel_pp ["Thm", "thm"] "Proof_Display.pp_thm"; +toplevel_pp ["Thm", "cterm"] "Proof_Display.pp_cterm"; +toplevel_pp ["Thm", "ctyp"] "Proof_Display.pp_ctyp"; +toplevel_pp ["typ"] "Proof_Display.pp_typ Pure.thy"; toplevel_pp ["Context", "theory"] "Context.pretty_thy"; toplevel_pp ["Context", "theory_ref"] "Context.pretty_thy o Theory.deref"; -toplevel_pp ["Context", "Proof", "context"] "ProofDisplay.pp_context"; +toplevel_pp ["Context", "Proof", "context"] "Proof_Display.pp_context"; toplevel_pp ["Syntax", "ast"] "Syntax.pretty_ast"; toplevel_pp ["Path", "T"] "Pretty.str o quote o Path.implode"; toplevel_pp ["File", "ident"] "Pretty.str o quote o File.rep_ident"; diff -r 2b5b0f9e271c -r 1464ddca182b src/Pure/pure_thy.ML --- a/src/Pure/pure_thy.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/Pure/pure_thy.ML Wed Nov 04 11:40:59 2009 +0100 @@ -227,8 +227,8 @@ (*** Pure theory syntax and logical content ***) -val typ = SimpleSyntax.read_typ; -val prop = SimpleSyntax.read_prop; +val typ = Simple_Syntax.read_typ; +val prop = Simple_Syntax.read_prop; val typeT = Syntax.typeT; val spropT = Syntax.spropT; @@ -268,7 +268,8 @@ (* main content *) val _ = Context.>> (Context.map_theory - (OldApplSyntax.init #> + (Sign.map_naming (Name_Space.set_theory_name Context.PureN) #> + OldApplSyntax.init #> Sign.add_types [(Binding.name "fun", 2, NoSyn), (Binding.name "prop", 0, NoSyn), diff -r 2b5b0f9e271c -r 1464ddca182b src/Pure/search.ML --- a/src/Pure/search.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/Pure/search.ML Wed Nov 04 11:40:59 2009 +0100 @@ -21,6 +21,7 @@ val iter_deepen_limit: int Unsynchronized.ref val THEN_ITER_DEEPEN: tactic -> (thm -> bool) -> (int -> tactic) -> tactic val ITER_DEEPEN: (thm -> bool) -> (int -> tactic) -> tactic + val trace_DEEPEN: bool Unsynchronized.ref val DEEPEN: int * int -> (int -> int -> tactic) -> int -> int -> tactic val trace_BEST_FIRST: bool Unsynchronized.ref val THEN_BEST_FIRST: tactic -> (thm -> bool) * (thm -> int) -> tactic -> tactic @@ -174,14 +175,19 @@ (*Simple iterative deepening tactical. It merely "deepens" any search tactic using increment "inc" up to limit "lim". *) +val trace_DEEPEN = Unsynchronized.ref false; + fun DEEPEN (inc,lim) tacf m i = - let fun dpn m st = - st |> (if has_fewer_prems i st then no_tac - else if m>lim then - (warning "Search depth limit exceeded: giving up"; - no_tac) - else (priority ("Search depth = " ^ string_of_int m); - tacf m i ORELSE dpn (m+inc))) + let + fun dpn m st = + st |> + (if has_fewer_prems i st then no_tac + else if m>lim then + (if !trace_DEEPEN then tracing "Search depth limit exceeded: giving up" else (); + no_tac) + else + (if !trace_DEEPEN then tracing ("Search depth = " ^ string_of_int m) else (); + tacf m i ORELSE dpn (m+inc))) in dpn m end; diff -r 2b5b0f9e271c -r 1464ddca182b src/Pure/sign.ML --- a/src/Pure/sign.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/Pure/sign.ML Wed Nov 04 11:40:59 2009 +0100 @@ -392,7 +392,7 @@ let val ((lhs, rhs), _) = tm |> no_vars (Syntax.pp ctxt) |> Logic.strip_imp_concl - |> PrimitiveDefs.dest_def ctxt Term.is_Const (K false) (K false) + |> Primitive_Defs.dest_def ctxt Term.is_Const (K false) (K false) in (Term.dest_Const (Term.head_of lhs), rhs) end handle TERM (msg, _) => error msg; diff -r 2b5b0f9e271c -r 1464ddca182b src/Pure/unify.ML --- a/src/Pure/unify.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/Pure/unify.ML Wed Nov 04 11:40:59 2009 +0100 @@ -473,7 +473,7 @@ if i j < i-lev) banned)) + else Bound (i - length (filter (fn j => j < i-lev) banned)) | change lev (Abs (a,T,t)) = Abs (a, T, change(lev+1) t) | change lev (t$u) = change lev t $ change lev u | change lev t = t diff -r 2b5b0f9e271c -r 1464ddca182b src/Tools/Code/code_haskell.ML --- a/src/Tools/Code/code_haskell.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/Tools/Code/code_haskell.ML Wed Nov 04 11:40:59 2009 +0100 @@ -360,14 +360,10 @@ fun serialize_module1 (module_name', (deps, (stmts, _))) = let val stmt_names = map fst stmts; - val deps' = subtract (op =) stmt_names deps + val qualified = is_none module_name; + val imports = subtract (op =) stmt_names deps |> distinct (op =) - |> map_filter (try deresolver); - val qualified = is_none module_name andalso - map deresolver stmt_names @ deps' - |> map Long_Name.base_name - |> has_duplicates (op =); - val imports = deps' + |> map_filter (try deresolver) |> map Long_Name.qualifier |> distinct (op =); fun pr_import_include (name, _) = str ("import qualified " ^ name ^ ";"); diff -r 2b5b0f9e271c -r 1464ddca182b src/Tools/Code/code_thingol.ML --- a/src/Tools/Code/code_thingol.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/Tools/Code/code_thingol.ML Wed Nov 04 11:40:59 2009 +0100 @@ -261,7 +261,8 @@ | NONE => (case Code.get_datatype_of_constr thy c of SOME dtco => Codegen.thyname_of_type thy dtco | NONE => Codegen.thyname_of_const thy c); - fun purify_base "op &" = "and" + fun purify_base "==>" = "follows" + | purify_base "op &" = "and" | purify_base "op |" = "or" | purify_base "op -->" = "implies" | purify_base "op :" = "member" diff -r 2b5b0f9e271c -r 1464ddca182b src/Tools/IsaPlanner/rw_inst.ML --- a/src/Tools/IsaPlanner/rw_inst.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/Tools/IsaPlanner/rw_inst.ML Wed Nov 04 11:40:59 2009 +0100 @@ -169,7 +169,7 @@ OldTerm.add_term_tfrees (t,tfrees))) ([],[]) ts; val unfixed_tvars = - List.filter (fn (ix,s) => not (member (op =) ignore_ixs ix)) tvars; + filter (fn (ix,s) => not (member (op =) ignore_ixs ix)) tvars; val (fixtyinsts, _) = List.foldr new_tfree ([], map fst tfrees) unfixed_tvars in (fixtyinsts, tfrees) end; diff -r 2b5b0f9e271c -r 1464ddca182b src/Tools/auto_solve.ML --- a/src/Tools/auto_solve.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/Tools/auto_solve.ML Wed Nov 04 11:40:59 2009 +0100 @@ -5,7 +5,7 @@ existing theorem. Duplicate lemmas can be detected in this way. The implementation is based in part on Berghofer and Haftmann's -quickcheck.ML. It relies critically on the FindTheorems solves +quickcheck.ML. It relies critically on the Find_Theorems solves feature. *) @@ -45,8 +45,8 @@ let val ctxt = Proof.context_of state; - val crits = [(true, FindTheorems.Solves)]; - fun find g = snd (FindTheorems.find_theorems ctxt (SOME g) (SOME (! limit)) false crits); + val crits = [(true, Find_Theorems.Solves)]; + fun find g = snd (Find_Theorems.find_theorems ctxt (SOME g) (SOME (! limit)) false crits); fun prt_result (goal, results) = let @@ -57,7 +57,7 @@ in Pretty.big_list (msg ^ " could be solved directly with:") - (map (FindTheorems.pretty_thm ctxt) results) + (map (Find_Theorems.pretty_thm ctxt) results) end; fun seek_against_goal () = diff -r 2b5b0f9e271c -r 1464ddca182b src/Tools/induct.ML --- a/src/Tools/induct.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/Tools/induct.ML Wed Nov 04 11:40:59 2009 +0100 @@ -113,9 +113,10 @@ type rules = (string * thm) Item_Net.T; -val init_rules = - Item_Net.init (fn ((s1: string, th1), (s2, th2)) => s1 = s2 andalso - Thm.eq_thm_prop (th1, th2)); +fun init_rules index : rules = + Item_Net.init + (fn ((s1, th1), (s2, th2)) => s1 = s2 andalso Thm.eq_thm_prop (th1, th2)) + (single o index); fun filter_rules (rs: rules) th = filter (fn (_, th') => Thm.eq_thm_prop (th, th')) (Item_Net.content rs); @@ -203,21 +204,21 @@ let val (x, thm) = g arg in (InductData.map (f (name, thm)) x, thm) end; fun del_att which = Thm.declaration_attribute (fn th => InductData.map (which (pairself (fn rs => - fold Item_Net.delete (filter_rules rs th) rs)))); + fold Item_Net.remove (filter_rules rs th) rs)))); fun map1 f (x, y, z) = (f x, y, z); fun map2 f (x, y, z) = (x, f y, z); fun map3 f (x, y, z) = (x, y, f z); -fun add_casesT rule x = map1 (apfst (Item_Net.insert rule)) x; -fun add_casesP rule x = map1 (apsnd (Item_Net.insert rule)) x; -fun add_inductT rule x = map2 (apfst (Item_Net.insert rule)) x; -fun add_inductP rule x = map2 (apsnd (Item_Net.insert rule)) x; -fun add_coinductT rule x = map3 (apfst (Item_Net.insert rule)) x; -fun add_coinductP rule x = map3 (apsnd (Item_Net.insert rule)) x; +fun add_casesT rule x = map1 (apfst (Item_Net.update rule)) x; +fun add_casesP rule x = map1 (apsnd (Item_Net.update rule)) x; +fun add_inductT rule x = map2 (apfst (Item_Net.update rule)) x; +fun add_inductP rule x = map2 (apsnd (Item_Net.update rule)) x; +fun add_coinductT rule x = map3 (apfst (Item_Net.update rule)) x; +fun add_coinductP rule x = map3 (apsnd (Item_Net.update rule)) x; -val consumes0 = RuleCases.consumes_default 0; -val consumes1 = RuleCases.consumes_default 1; +val consumes0 = Rule_Cases.consumes_default 0; +val consumes1 = Rule_Cases.consumes_default 1; in @@ -344,10 +345,10 @@ val thy = ProofContext.theory_of ctxt; fun inst_rule r = - if null insts then `RuleCases.get r + if null insts then `Rule_Cases.get r else (align_left "Rule has fewer premises than arguments given" (Thm.prems_of r) insts |> maps (prep_inst ctxt align_left I) - |> Drule.cterm_instantiate) r |> pair (RuleCases.get r); + |> Drule.cterm_instantiate) r |> pair (Rule_Cases.get r); val ruleq = (case opt_rule of @@ -359,9 +360,9 @@ in fn i => fn st => ruleq - |> Seq.maps (RuleCases.consume [] facts) + |> Seq.maps (Rule_Cases.consume [] facts) |> Seq.maps (fn ((cases, (_, more_facts)), rule) => - CASES (RuleCases.make_common false (thy, Thm.prop_of rule) cases) + CASES (Rule_Cases.make_common false (thy, Thm.prop_of rule) cases) (Method.insert_tac more_facts i THEN Tactic.rtac rule i) st) end; @@ -483,7 +484,7 @@ in Logic.list_implies (map rename_asm As, C) end; val cp' = cterm_fun rename_prop (Thm.cprop_of thm); val thm' = Thm.equal_elim (Thm.reflexive cp') thm; - in [RuleCases.save thm thm'] end + in [Rule_Cases.save thm thm'] end | special_rename_params _ _ ths = ths; @@ -570,7 +571,7 @@ ((fn [] => NONE | ts => List.last ts) #> (fn NONE => TVar (("'a", 0), []) | SOME t => Term.fastype_of t) #> find_inductT ctxt)) [[]] - |> filter_out (forall Thm.is_internal); + |> filter_out (forall Rule_Cases.is_inner_rule); fun get_inductP ctxt (fact :: _) = map single (find_inductP ctxt (Thm.concl_of fact)) | get_inductP _ _ = []; @@ -583,29 +584,29 @@ val atomized_defs = map (map (Conv.fconv_rule ObjectLogic.atomize)) defs; fun inst_rule (concls, r) = - (if null insts then `RuleCases.get r + (if null insts then `Rule_Cases.get r else (align_left "Rule has fewer conclusions than arguments given" (map Logic.strip_imp_concl (Logic.dest_conjunctions (Thm.concl_of r))) insts |> maps (prep_inst ctxt align_right (atomize_term thy)) - |> Drule.cterm_instantiate) r |> pair (RuleCases.get r)) + |> Drule.cterm_instantiate) r |> pair (Rule_Cases.get r)) |> (fn ((cases, consumes), th) => (((cases, concls), consumes), th)); val ruleq = (case opt_rule of - SOME rs => Seq.single (inst_rule (RuleCases.strict_mutual_rule ctxt rs)) + SOME rs => Seq.single (inst_rule (Rule_Cases.strict_mutual_rule ctxt rs)) | NONE => (get_inductP ctxt facts @ map (special_rename_params defs_ctxt insts) (get_inductT ctxt insts)) - |> map_filter (RuleCases.mutual_rule ctxt) + |> map_filter (Rule_Cases.mutual_rule ctxt) |> tap (trace_rules ctxt inductN o map #2) |> Seq.of_list |> Seq.maps (Seq.try inst_rule)); fun rule_cases rule = - RuleCases.make_nested false (Thm.prop_of rule) (rulified_term rule); + Rule_Cases.make_nested false (Thm.prop_of rule) (rulified_term rule); in (fn i => fn st => ruleq - |> Seq.maps (RuleCases.consume (flat defs) facts) + |> Seq.maps (Rule_Cases.consume (flat defs) facts) |> Seq.maps (fn (((cases, concls), (more_consumes, more_facts)), rule) => (PRECISE_CONJUNCTS (length concls) (ALLGOALS (fn j => (CONJUNCTS (ALLGOALS @@ -643,7 +644,7 @@ fun get_coinductP ctxt goal = find_coinductP ctxt (Logic.strip_assums_concl goal); fun main_prop_of th = - if RuleCases.get_consumes th > 0 then Thm.major_prem_of th else Thm.concl_of th; + if Rule_Cases.get_consumes th > 0 then Thm.major_prem_of th else Thm.concl_of th; in @@ -652,9 +653,9 @@ val thy = ProofContext.theory_of ctxt; fun inst_rule r = - if null inst then `RuleCases.get r + if null inst then `Rule_Cases.get r else Drule.cterm_instantiate (prep_inst ctxt align_right I (main_prop_of r, inst)) r - |> pair (RuleCases.get r); + |> pair (Rule_Cases.get r); fun ruleq goal = (case opt_rule of @@ -666,12 +667,12 @@ in SUBGOAL_CASES (fn (goal, i) => fn st => ruleq goal - |> Seq.maps (RuleCases.consume [] facts) + |> Seq.maps (Rule_Cases.consume [] facts) |> Seq.maps (fn ((cases, (_, more_facts)), rule) => guess_instance ctxt rule i st |> Seq.map (rule_instance ctxt (burrow_options (Variable.polymorphic ctxt) taking)) |> Seq.maps (fn rule' => - CASES (RuleCases.make_common false (thy, Thm.prop_of rule') cases) + CASES (Rule_Cases.make_common false (thy, Thm.prop_of rule') cases) (Method.insert_tac more_facts i THEN Tactic.rtac rule' i) st))) end; diff -r 2b5b0f9e271c -r 1464ddca182b src/Tools/induct_tacs.ML --- a/src/Tools/induct_tacs.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/Tools/induct_tacs.ML Wed Nov 04 11:40:59 2009 +0100 @@ -86,9 +86,9 @@ val argss = map (map (Option.map induct_var)) varss; val rule = (case opt_rules of - SOME rules => #2 (RuleCases.strict_mutual_rule ctxt rules) + SOME rules => #2 (Rule_Cases.strict_mutual_rule ctxt rules) | NONE => - (case map_filter (RuleCases.mutual_rule ctxt) (Induct.get_inductT ctxt argss) of + (case map_filter (Rule_Cases.mutual_rule ctxt) (Induct.get_inductT ctxt argss) of (_, rule) :: _ => rule | [] => raise THM ("No induction rules", 0, []))); diff -r 2b5b0f9e271c -r 1464ddca182b src/Tools/intuitionistic.ML --- a/src/Tools/intuitionistic.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/Tools/intuitionistic.ML Wed Nov 04 11:40:59 2009 +0100 @@ -25,18 +25,18 @@ fun REMDUPS tac = tac THEN_ALL_NEW remdups_tac; -val bires_tac = Tactic.biresolution_from_nets_tac ContextRules.orderlist; +val bires_tac = Tactic.biresolution_from_nets_tac Context_Rules.orderlist; fun safe_step_tac ctxt = - ContextRules.Swrap ctxt + Context_Rules.Swrap ctxt (eq_assume_tac ORELSE' - bires_tac true (ContextRules.netpair_bang ctxt)); + bires_tac true (Context_Rules.netpair_bang ctxt)); fun unsafe_step_tac ctxt = - ContextRules.wrap ctxt + Context_Rules.wrap ctxt (assume_tac APPEND' - bires_tac false (ContextRules.netpair_bang ctxt) APPEND' - bires_tac false (ContextRules.netpair ctxt)); + bires_tac false (Context_Rules.netpair_bang ctxt) APPEND' + bires_tac false (Context_Rules.netpair ctxt)); fun step_tac ctxt i = REPEAT_DETERM1 (REMDUPS (safe_step_tac ctxt) i) ORELSE @@ -75,13 +75,13 @@ >> (pair (I: Proof.context -> Proof.context) o att); val modifiers = - [modifier destN Args.bang_colon Args.bang ContextRules.dest_bang, - modifier destN Args.colon (Scan.succeed ()) ContextRules.dest, - modifier elimN Args.bang_colon Args.bang ContextRules.elim_bang, - modifier elimN Args.colon (Scan.succeed ()) ContextRules.elim, - modifier introN Args.bang_colon Args.bang ContextRules.intro_bang, - modifier introN Args.colon (Scan.succeed ()) ContextRules.intro, - Args.del -- Args.colon >> K (I, ContextRules.rule_del)]; + [modifier destN Args.bang_colon Args.bang Context_Rules.dest_bang, + modifier destN Args.colon (Scan.succeed ()) Context_Rules.dest, + modifier elimN Args.bang_colon Args.bang Context_Rules.elim_bang, + modifier elimN Args.colon (Scan.succeed ()) Context_Rules.elim, + modifier introN Args.bang_colon Args.bang Context_Rules.intro_bang, + modifier introN Args.colon (Scan.succeed ()) Context_Rules.intro, + Args.del -- Args.colon >> K (I, Context_Rules.rule_del)]; in diff -r 2b5b0f9e271c -r 1464ddca182b src/Tools/project_rule.ML --- a/src/Tools/project_rule.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/Tools/project_rule.ML Wed Nov 04 11:40:59 2009 +0100 @@ -47,8 +47,8 @@ Thm.permute_prems 0 (~ k) #> singleton (Variable.export ctxt' ctxt) #> Drule.zero_var_indexes - #> RuleCases.save raw_rule - #> RuleCases.add_consumes k); + #> Rule_Cases.save raw_rule + #> Rule_Cases.add_consumes k); in map result is end; fun project ctxt i th = hd (projects ctxt [i] th); diff -r 2b5b0f9e271c -r 1464ddca182b src/ZF/Tools/datatype_package.ML --- a/src/ZF/Tools/datatype_package.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/ZF/Tools/datatype_package.ML Wed Nov 04 11:40:59 2009 +0100 @@ -108,7 +108,7 @@ let val ncon = length con_ty_list (*number of constructors*) fun mk_def (((id,T,syn), name, args, prems), kcon) = (*kcon is index of constructor*) - PrimitiveDefs.mk_defpair (list_comb (Const (full_name name, T), args), + Primitive_Defs.mk_defpair (list_comb (Const (full_name name, T), args), mk_inject npart kpart (mk_inject ncon kcon (mk_tuple args))) in ListPair.map mk_def (con_ty_list, 1 upto ncon) end; @@ -129,7 +129,7 @@ Non-identifiers (e.g. infixes) get a name of the form f_op_nnn. **) (*The function variable for a single constructor*) - fun add_case (((_, T, _), name, args, _), (opno, cases)) = + fun add_case ((_, T, _), name, args, _) (opno, cases) = if Syntax.is_identifier name then (opno, (Free (case_varname ^ "_" ^ name, T), args) :: cases) else @@ -138,12 +138,12 @@ (*Treatment of a list of constructors, for one part Result adds a list of terms, each a function variable with arguments*) - fun add_case_list (con_ty_list, (opno, case_lists)) = - let val (opno', case_list) = List.foldr add_case (opno, []) con_ty_list + fun add_case_list con_ty_list (opno, case_lists) = + let val (opno', case_list) = fold_rev add_case con_ty_list (opno, []) in (opno', case_list :: case_lists) end; (*Treatment of all parts*) - val (_, case_lists) = List.foldr add_case_list (1,[]) con_ty_lists; + val (_, case_lists) = fold_rev add_case_list con_ty_lists (1, []); (*extract the types of all the variables*) val case_typ = maps (map (#2 o #1)) con_ty_lists ---> @{typ "i => i"}; @@ -157,7 +157,7 @@ val case_const = Const (case_name, case_typ); val case_tm = list_comb (case_const, case_args); - val case_def = PrimitiveDefs.mk_defpair + val case_def = Primitive_Defs.mk_defpair (case_tm, Balanced_Tree.make (fn (t1, t2) => Su.elim $ t1 $ t2) (map call_case case_lists)); @@ -215,7 +215,7 @@ val rec_ty_lists = (map (map rec_ty_elem) con_ty_lists); (*Treatment of all parts*) - val (_, recursor_lists) = List.foldr add_case_list (1,[]) rec_ty_lists; + val (_, recursor_lists) = fold_rev add_case_list rec_ty_lists (1, []); (*extract the types of all the variables*) val recursor_typ = maps (map (#2 o #1)) rec_ty_lists ---> @{typ "i => i"}; @@ -232,7 +232,7 @@ val recursor_cases = map call_recursor (flat case_lists ~~ flat recursor_lists); val recursor_def = - PrimitiveDefs.mk_defpair + Primitive_Defs.mk_defpair (recursor_tm, @{const Univ.Vrecursor} $ absfree ("rec", @{typ i}, list_comb (case_const, recursor_cases))); diff -r 2b5b0f9e271c -r 1464ddca182b src/ZF/Tools/inductive_package.ML --- a/src/ZF/Tools/inductive_package.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/ZF/Tools/inductive_package.ML Wed Nov 04 11:40:59 2009 +0100 @@ -66,7 +66,7 @@ val intr_specs = map (apfst (apfst Binding.name_of)) raw_intr_specs; val (intr_names, intr_tms) = split_list (map fst intr_specs); - val case_names = RuleCases.case_names intr_names; + val case_names = Rule_Cases.case_names intr_names; (*recT and rec_params should agree for all mutually recursive components*) val rec_hds = map head_of rec_tms; @@ -156,7 +156,7 @@ val big_rec_tm = list_comb(Const(big_rec_name,recT), rec_params); (*The individual sets must already be declared*) - val axpairs = map PrimitiveDefs.mk_defpair + val axpairs = map Primitive_Defs.mk_defpair ((big_rec_tm, fp_rhs) :: (case parts of [_] => [] (*no mutual recursion*) diff -r 2b5b0f9e271c -r 1464ddca182b src/ZF/arith_data.ML --- a/src/ZF/arith_data.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/ZF/arith_data.ML Wed Nov 04 11:40:59 2009 +0100 @@ -68,7 +68,7 @@ fun prove_conv name tacs ctxt prems (t,u) = if t aconv u then NONE else - let val prems' = List.filter (not o is_eq_thm) prems + let val prems' = filter_out is_eq_thm prems val goal = Logic.list_implies (map (#prop o Thm.rep_thm) prems', FOLogic.mk_Trueprop (mk_eq_iff (t, u))); in SOME (prems' MRS Goal.prove ctxt [] [] goal (K (EVERY tacs))) diff -r 2b5b0f9e271c -r 1464ddca182b src/ZF/ind_syntax.ML --- a/src/ZF/ind_syntax.ML Thu Oct 29 16:23:57 2009 +0100 +++ b/src/ZF/ind_syntax.ML Wed Nov 04 11:40:59 2009 +0100 @@ -96,8 +96,8 @@ fun union_params (rec_tm, cs) = let val (_,args) = strip_comb rec_tm fun is_ind arg = (type_of arg = iT) - in case List.filter is_ind (args @ cs) of - [] => @{const 0} + in case filter is_ind (args @ cs) of + [] => @{const 0} | u_args => Balanced_Tree.make mk_Un u_args end;