# HG changeset patch # User boehmes # Date 1256026290 -7200 # Node ID 39f73a59e8554cd64f0275e9edfa3b012fb6613a # Parent b0ff69f0a24859a512f4a154f3b3488c74b4a1aa added proof reconstructon for Z3, added certificates for simpler re-checking of proofs (no need to invoke external solvers), added examples and certificates for all examples, removed Unsynchronized.ref (in smt_normalize.ML) diff -r b0ff69f0a248 -r 39f73a59e855 CONTRIBUTORS --- a/CONTRIBUTORS Tue Oct 20 08:10:47 2009 +0200 +++ b/CONTRIBUTORS Tue Oct 20 10:11:30 2009 +0200 @@ -7,10 +7,13 @@ Contributions to this Isabelle version -------------------------------------- -* Oktober 2009: Florian Haftmann, TUM +* October 2009: Sascha Boehme, TUM + Extension of SMT method: proof-reconstruction for the SMT solver Z3 + +* October 2009: Florian Haftmann, TUM Refinement of parts of the HOL datatype package -* Oktober 2009: Florian Haftmann, TUM +* October 2009: Florian Haftmann, TUM Generic term styles for term antiquotations * September 2009: Thomas Sewell, NICTA diff -r b0ff69f0a248 -r 39f73a59e855 NEWS --- a/NEWS Tue Oct 20 08:10:47 2009 +0200 +++ b/NEWS Tue Oct 20 10:11:30 2009 +0200 @@ -46,7 +46,9 @@ arithmetic, and fixed-size bitvectors; there is also basic support for higher-order features (esp. lambda abstractions). It is an incomplete decision procedure based on external SMT -solvers using the oracle mechanism. +solvers using the oracle mechanism; for the SMT solver Z3, +this method is proof-producing. Certificates are provided to +avoid calling the external solvers solely for re-checking proofs. * Reorganization of number theory: * former session NumberTheory now named Old_Number_Theory diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/IsaMakefile --- a/src/HOL/IsaMakefile Tue Oct 20 08:10:47 2009 +0200 +++ b/src/HOL/IsaMakefile Tue Oct 20 10:11:30 2009 +0200 @@ -6,7 +6,7 @@ default: HOL generate: HOL-Generate-HOL HOL-Generate-HOLLight -images: HOL HOL-Base HOL-Plain HOL-Main HOL-Algebra HOL-Nominal HOL-NSA HOL-Word TLA HOL4 +images: HOL HOL-Base HOL-Plain HOL-Main HOL-Algebra HOL-Nominal HOL-NSA HOL-SMT HOL-Word TLA HOL4 #Note: keep targets sorted (except for HOL-Library and HOL-ex) test: \ @@ -40,7 +40,7 @@ HOL-Prolog \ HOL-SET-Protocol \ HOL-SizeChange \ - HOL-SMT \ + HOL-SMT-Examples \ HOL-Statespace \ HOL-Subst \ TLA-Buffer \ @@ -1146,15 +1146,222 @@ HOL-SMT: HOL-Word $(LOG)/HOL-SMT.gz -$(LOG)/HOL-SMT.gz: $(OUT)/HOL-Word SMT/SMT_Definitions.thy SMT/SMT.thy \ +$(LOG)/HOL-SMT.gz: $(OUT)/HOL-Word 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/smt_builtin.ML \ SMT/Tools/smtlib_interface.ML SMT/Tools/smt_solver.ML \ SMT/Tools/cvc3_solver.ML SMT/Tools/yices_solver.ML \ - SMT/Tools/z3_interface.ML SMT/Tools/z3_solver.ML SMT/Tools/z3_model.ML + SMT/Tools/z3_proof_terms.ML SMT/Tools/z3_proof_rules.ML \ + SMT/Tools/z3_proof.ML SMT/Tools/z3_model.ML SMT/Tools/z3_interface.ML \ + SMT/Tools/z3_solver.ML @cd SMT; $(ISABELLE_TOOL) usedir -b -g true $(OUT)/HOL-Word HOL-SMT +## HOL-SMT-Examples + +HOL-SMT-Examples: HOL-SMT $(LOG)/HOL-SMT-Examples.gz + +$(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 \ + SMT/Examples/cert/z3_arith_quant_02.proof \ + SMT/Examples/cert/z3_arith_quant_03 \ + SMT/Examples/cert/z3_arith_quant_03.proof \ + SMT/Examples/cert/z3_arith_quant_04 \ + SMT/Examples/cert/z3_arith_quant_04.proof \ + SMT/Examples/cert/z3_arith_quant_05 \ + SMT/Examples/cert/z3_arith_quant_05.proof \ + SMT/Examples/cert/z3_arith_quant_06 \ + SMT/Examples/cert/z3_arith_quant_06.proof \ + SMT/Examples/cert/z3_arith_quant_07 \ + SMT/Examples/cert/z3_arith_quant_07.proof \ + SMT/Examples/cert/z3_arith_quant_08 \ + SMT/Examples/cert/z3_arith_quant_08.proof \ + SMT/Examples/cert/z3_arith_quant_09 \ + SMT/Examples/cert/z3_arith_quant_09.proof \ + SMT/Examples/cert/z3_arith_quant_10 \ + SMT/Examples/cert/z3_arith_quant_10.proof \ + SMT/Examples/cert/z3_arith_quant_11 \ + SMT/Examples/cert/z3_arith_quant_11.proof \ + SMT/Examples/cert/z3_arith_quant_12 \ + SMT/Examples/cert/z3_arith_quant_12.proof \ + SMT/Examples/cert/z3_arith_quant_13 \ + SMT/Examples/cert/z3_arith_quant_13.proof \ + SMT/Examples/cert/z3_arith_quant_14 \ + SMT/Examples/cert/z3_arith_quant_14.proof \ + SMT/Examples/cert/z3_arith_quant_15 \ + SMT/Examples/cert/z3_arith_quant_15.proof \ + SMT/Examples/cert/z3_arith_quant_16 \ + SMT/Examples/cert/z3_arith_quant_16.proof \ + SMT/Examples/cert/z3_arith_quant_17 \ + SMT/Examples/cert/z3_arith_quant_17.proof \ + SMT/Examples/cert/z3_arith_quant_18 \ + SMT/Examples/cert/z3_arith_quant_18.proof \ + SMT/Examples/cert/z3_bv_01 \ + SMT/Examples/cert/z3_bv_01.proof \ + SMT/Examples/cert/z3_bv_02 \ + SMT/Examples/cert/z3_bv_02.proof \ + SMT/Examples/cert/z3_bv_arith_01 \ + SMT/Examples/cert/z3_bv_arith_01.proof \ + SMT/Examples/cert/z3_bv_arith_02 \ + SMT/Examples/cert/z3_bv_arith_02.proof \ + SMT/Examples/cert/z3_bv_arith_03 \ + SMT/Examples/cert/z3_bv_arith_03.proof \ + SMT/Examples/cert/z3_bv_arith_04 \ + SMT/Examples/cert/z3_bv_arith_04.proof \ + SMT/Examples/cert/z3_bv_arith_05 \ + SMT/Examples/cert/z3_bv_arith_05.proof \ + SMT/Examples/cert/z3_bv_arith_06 \ + SMT/Examples/cert/z3_bv_arith_06.proof \ + SMT/Examples/cert/z3_bv_arith_07 \ + SMT/Examples/cert/z3_bv_arith_07.proof \ + SMT/Examples/cert/z3_bv_arith_08 \ + SMT/Examples/cert/z3_bv_arith_08.proof \ + SMT/Examples/cert/z3_bv_arith_09 \ + SMT/Examples/cert/z3_bv_arith_09.proof \ + SMT/Examples/cert/z3_bv_arith_10 \ + SMT/Examples/cert/z3_bv_arith_10.proof \ + SMT/Examples/cert/z3_bv_bit_01 \ + SMT/Examples/cert/z3_bv_bit_01.proof \ + SMT/Examples/cert/z3_bv_bit_02 \ + SMT/Examples/cert/z3_bv_bit_02.proof \ + SMT/Examples/cert/z3_bv_bit_03 \ + SMT/Examples/cert/z3_bv_bit_03.proof \ + SMT/Examples/cert/z3_bv_bit_04 \ + SMT/Examples/cert/z3_bv_bit_04.proof \ + SMT/Examples/cert/z3_bv_bit_05 \ + SMT/Examples/cert/z3_bv_bit_05.proof \ + SMT/Examples/cert/z3_bv_bit_06 \ + SMT/Examples/cert/z3_bv_bit_06.proof \ + SMT/Examples/cert/z3_bv_bit_07 \ + SMT/Examples/cert/z3_bv_bit_07.proof \ + SMT/Examples/cert/z3_bv_bit_08 \ + SMT/Examples/cert/z3_bv_bit_08.proof \ + SMT/Examples/cert/z3_bv_bit_09 \ + SMT/Examples/cert/z3_bv_bit_09.proof \ + SMT/Examples/cert/z3_bv_bit_10 \ + SMT/Examples/cert/z3_bv_bit_10.proof \ + SMT/Examples/cert/z3_bv_bit_11 \ + SMT/Examples/cert/z3_bv_bit_11.proof \ + SMT/Examples/cert/z3_bv_bit_12 \ + SMT/Examples/cert/z3_bv_bit_12.proof \ + SMT/Examples/cert/z3_bv_bit_13 \ + SMT/Examples/cert/z3_bv_bit_13.proof \ + SMT/Examples/cert/z3_bv_bit_14 \ + SMT/Examples/cert/z3_bv_bit_14.proof \ + SMT/Examples/cert/z3_bv_bit_15 \ + SMT/Examples/cert/z3_bv_bit_15.proof \ + SMT/Examples/cert/z3_fol_01 \ + SMT/Examples/cert/z3_fol_01.proof \ + SMT/Examples/cert/z3_fol_02 \ + SMT/Examples/cert/z3_fol_02.proof \ + SMT/Examples/cert/z3_fol_03 \ + SMT/Examples/cert/z3_fol_03.proof \ + SMT/Examples/cert/z3_fol_04 \ + SMT/Examples/cert/z3_fol_04.proof \ + SMT/Examples/cert/z3_hol_01 \ + SMT/Examples/cert/z3_hol_01.proof \ + SMT/Examples/cert/z3_hol_02 \ + SMT/Examples/cert/z3_hol_02.proof \ + SMT/Examples/cert/z3_hol_03 \ + SMT/Examples/cert/z3_hol_03.proof \ + SMT/Examples/cert/z3_hol_04 \ + SMT/Examples/cert/z3_hol_04.proof \ + SMT/Examples/cert/z3_hol_05 \ + SMT/Examples/cert/z3_hol_05.proof \ + SMT/Examples/cert/z3_hol_06 \ + SMT/Examples/cert/z3_hol_06.proof \ + SMT/Examples/cert/z3_hol_07 \ + SMT/Examples/cert/z3_hol_07.proof \ + SMT/Examples/cert/z3_hol_08 \ + SMT/Examples/cert/z3_hol_08.proof \ + SMT/Examples/cert/z3_linarith_01 \ + SMT/Examples/cert/z3_linarith_01.proof \ + SMT/Examples/cert/z3_linarith_02 \ + SMT/Examples/cert/z3_linarith_02.proof \ + SMT/Examples/cert/z3_linarith_03 \ + SMT/Examples/cert/z3_linarith_03.proof \ + SMT/Examples/cert/z3_linarith_04 \ + SMT/Examples/cert/z3_linarith_04.proof \ + SMT/Examples/cert/z3_linarith_05 \ + SMT/Examples/cert/z3_linarith_05.proof \ + SMT/Examples/cert/z3_linarith_06 \ + SMT/Examples/cert/z3_linarith_06.proof \ + SMT/Examples/cert/z3_linarith_07 \ + SMT/Examples/cert/z3_linarith_07.proof \ + SMT/Examples/cert/z3_linarith_08 \ + SMT/Examples/cert/z3_linarith_08.proof \ + SMT/Examples/cert/z3_linarith_09 \ + SMT/Examples/cert/z3_linarith_09.proof \ + SMT/Examples/cert/z3_linarith_10 \ + SMT/Examples/cert/z3_linarith_10.proof \ + SMT/Examples/cert/z3_linarith_11 \ + SMT/Examples/cert/z3_linarith_11.proof \ + SMT/Examples/cert/z3_linarith_12 \ + SMT/Examples/cert/z3_linarith_12.proof \ + SMT/Examples/cert/z3_linarith_13 \ + SMT/Examples/cert/z3_linarith_13.proof \ + SMT/Examples/cert/z3_linarith_14 \ + SMT/Examples/cert/z3_linarith_14.proof \ + SMT/Examples/cert/z3_linarith_15 \ + SMT/Examples/cert/z3_linarith_15.proof \ + SMT/Examples/cert/z3_linarith_16 \ + SMT/Examples/cert/z3_linarith_16.proof \ + SMT/Examples/cert/z3_mono_01 \ + SMT/Examples/cert/z3_mono_01.proof \ + SMT/Examples/cert/z3_mono_02 \ + SMT/Examples/cert/z3_mono_02.proof \ + SMT/Examples/cert/z3_nat_arith_01 \ + SMT/Examples/cert/z3_nat_arith_01.proof \ + SMT/Examples/cert/z3_nat_arith_02 \ + SMT/Examples/cert/z3_nat_arith_02.proof \ + SMT/Examples/cert/z3_nat_arith_03 \ + SMT/Examples/cert/z3_nat_arith_03.proof \ + SMT/Examples/cert/z3_nat_arith_04 \ + SMT/Examples/cert/z3_nat_arith_04.proof \ + SMT/Examples/cert/z3_nat_arith_05 \ + SMT/Examples/cert/z3_nat_arith_05.proof \ + SMT/Examples/cert/z3_nat_arith_06 \ + SMT/Examples/cert/z3_nat_arith_06.proof \ + SMT/Examples/cert/z3_nat_arith_07 \ + SMT/Examples/cert/z3_nat_arith_07.proof \ + SMT/Examples/cert/z3_nlarith_01 \ + SMT/Examples/cert/z3_nlarith_01.proof \ + SMT/Examples/cert/z3_nlarith_02 \ + SMT/Examples/cert/z3_nlarith_02.proof \ + SMT/Examples/cert/z3_nlarith_03 \ + SMT/Examples/cert/z3_nlarith_03.proof \ + SMT/Examples/cert/z3_nlarith_04 \ + SMT/Examples/cert/z3_nlarith_04.proof \ + SMT/Examples/cert/z3_pair_01 \ + SMT/Examples/cert/z3_pair_01.proof \ + SMT/Examples/cert/z3_pair_02 \ + SMT/Examples/cert/z3_pair_02.proof \ + SMT/Examples/cert/z3_prop_01 \ + SMT/Examples/cert/z3_prop_01.proof \ + SMT/Examples/cert/z3_prop_02 \ + SMT/Examples/cert/z3_prop_02.proof \ + SMT/Examples/cert/z3_prop_03 \ + SMT/Examples/cert/z3_prop_03.proof \ + SMT/Examples/cert/z3_prop_04 \ + SMT/Examples/cert/z3_prop_04.proof \ + SMT/Examples/cert/z3_prop_05 \ + SMT/Examples/cert/z3_prop_05.proof \ + SMT/Examples/cert/z3_prop_06 \ + SMT/Examples/cert/z3_prop_06.proof \ + SMT/Examples/cert/z3_prop_07 \ + SMT/Examples/cert/z3_prop_07.proof \ + SMT/Examples/cert/z3_prop_08 \ + SMT/Examples/cert/z3_prop_08.proof \ + SMT/Examples/cert/z3_prop_09 \ + SMT/Examples/cert/z3_prop_09.proof \ + SMT/Examples/cert/z3_prop_10 \ + SMT/Examples/cert/z3_prop_10.proof + @cd SMT; $(ISABELLE_TOOL) usedir $(OUT)/HOL-SMT Examples + + ## clean clean: @@ -1177,4 +1384,6 @@ $(LOG)/HOL-Unix.gz $(OUT)/HOL-Word $(LOG)/HOL-Word.gz \ $(LOG)/HOL-Word-Examples.gz $(OUT)/HOL-NSA \ $(LOG)/HOL-NSA.gz $(LOG)/HOL-NSA-Examples.gz \ - $(LOG)/HOL-Mirabelle.gz $(LOG)/HOL-SMT.gz + $(LOG)/HOL-Mirabelle.gz $(LOG)/HOL-SMT.gz \ + $(LOG)/HOL-SMT-Examples.gz + diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/ROOT.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/ROOT.ML Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,1 @@ +use_thy "SMT_Examples"; diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/SMT_Examples.thy --- a/src/HOL/SMT/Examples/SMT_Examples.thy Tue Oct 20 08:10:47 2009 +0200 +++ b/src/HOL/SMT/Examples/SMT_Examples.thy Tue Oct 20 10:11:30 2009 +0200 @@ -5,35 +5,361 @@ header {* Examples for the 'smt' tactic. *} theory SMT_Examples -imports "../SMT" +imports SMT begin -declare [[smt_solver=z3, z3_proofs=false]] -declare [[smt_trace=false]] +declare [[smt_solver=z3, z3_proofs=true]] section {* Propositional and first-order logic *} -lemma "True" by smt -lemma "p \ \p" by smt -lemma "(p \ True) = p" by smt -lemma "(p \ q) \ \p \ q" by smt -lemma "(a \ b) \ (c \ d) \ (a \ b) \ (c \ d)" by smt -lemma "P=P=P=P=P=P=P=P=P=P" by smt +lemma "True" + using [[smt_cert="~/isabelle/SMT/Examples/cert/z3_prop_01"]] + by smt + +lemma "p \ \p" + using [[smt_cert="~/isabelle/SMT/Examples/cert/z3_prop_02"]] + by smt + +lemma "(p \ True) = p" + using [[smt_cert="~/isabelle/SMT/Examples/cert/z3_prop_03"]] + by smt + +lemma "(p \ q) \ \p \ q" + using [[smt_cert="~/isabelle/SMT/Examples/cert/z3_prop_04"]] + by smt + +lemma "(a \ b) \ (c \ d) \ (a \ b) \ (c \ d)" + using [[smt_cert="~/isabelle/SMT/Examples/cert/z3_prop_05"]] + using [[z3_proofs=false]] (* no Z3 proof *) + by smt + +lemma "(p1 \ p2) \ p3 \ (p1 \ (p3 \ p2) \ (p1 \ p3)) \ p1" + using [[smt_cert="~/isabelle/SMT/Examples/cert/z3_prop_06"]] + by smt + +lemma "P=P=P=P=P=P=P=P=P=P" + using [[smt_cert="~/isabelle/SMT/Examples/cert/z3_prop_07"]] + by smt + +lemma + assumes "a | b | c | d" + and "e | f | (a & d)" + and "~(a | (c & ~c)) | b" + and "~(b & (x | ~x)) | c" + and "~(d | False) | c" + and "~(c | (~p & (p | (q & ~q))))" + shows False + using [[smt_cert="~/isabelle/SMT/Examples/cert/z3_prop_08"]] + using assms by smt axiomatization symm_f :: "'a \ 'a \ 'a" where symm_f: "symm_f x y = symm_f y x" -lemma "a = a \ symm_f a b = symm_f b a" by (smt add: symm_f) +lemma "a = a \ symm_f a b = symm_f b a" + using [[smt_cert="~/isabelle/SMT/Examples/cert/z3_prop_09"]] + by (smt add: symm_f) + +(* +Taken from ~~/src/HOL/ex/SAT_Examples.thy. +Translated from TPTP problem library: PUZ015-2.006.dimacs +*) +lemma + assumes "~x0" + and "~x30" + and "~x29" + and "~x59" + and "x1 | x31 | x0" + and "x2 | x32 | x1" + and "x3 | x33 | x2" + and "x4 | x34 | x3" + and "x35 | x4" + and "x5 | x36 | x30" + and "x6 | x37 | x5 | x31" + and "x7 | x38 | x6 | x32" + and "x8 | x39 | x7 | x33" + and "x9 | x40 | x8 | x34" + and "x41 | x9 | x35" + and "x10 | x42 | x36" + and "x11 | x43 | x10 | x37" + and "x12 | x44 | x11 | x38" + and "x13 | x45 | x12 | x39" + and "x14 | x46 | x13 | x40" + and "x47 | x14 | x41" + and "x15 | x48 | x42" + and "x16 | x49 | x15 | x43" + and "x17 | x50 | x16 | x44" + and "x18 | x51 | x17 | x45" + and "x19 | x52 | x18 | x46" + and "x53 | x19 | x47" + and "x20 | x54 | x48" + and "x21 | x55 | x20 | x49" + and "x22 | x56 | x21 | x50" + and "x23 | x57 | x22 | x51" + and "x24 | x58 | x23 | x52" + and "x59 | x24 | x53" + and "x25 | x54" + and "x26 | x25 | x55" + and "x27 | x26 | x56" + and "x28 | x27 | x57" + and "x29 | x28 | x58" + and "~x1 | ~x31" + and "~x1 | ~x0" + and "~x31 | ~x0" + and "~x2 | ~x32" + and "~x2 | ~x1" + and "~x32 | ~x1" + and "~x3 | ~x33" + and "~x3 | ~x2" + and "~x33 | ~x2" + and "~x4 | ~x34" + and "~x4 | ~x3" + and "~x34 | ~x3" + and "~x35 | ~x4" + and "~x5 | ~x36" + and "~x5 | ~x30" + and "~x36 | ~x30" + and "~x6 | ~x37" + and "~x6 | ~x5" + and "~x6 | ~x31" + and "~x37 | ~x5" + and "~x37 | ~x31" + and "~x5 | ~x31" + and "~x7 | ~x38" + and "~x7 | ~x6" + and "~x7 | ~x32" + and "~x38 | ~x6" + and "~x38 | ~x32" + and "~x6 | ~x32" + and "~x8 | ~x39" + and "~x8 | ~x7" + and "~x8 | ~x33" + and "~x39 | ~x7" + and "~x39 | ~x33" + and "~x7 | ~x33" + and "~x9 | ~x40" + and "~x9 | ~x8" + and "~x9 | ~x34" + and "~x40 | ~x8" + and "~x40 | ~x34" + and "~x8 | ~x34" + and "~x41 | ~x9" + and "~x41 | ~x35" + and "~x9 | ~x35" + and "~x10 | ~x42" + and "~x10 | ~x36" + and "~x42 | ~x36" + and "~x11 | ~x43" + and "~x11 | ~x10" + and "~x11 | ~x37" + and "~x43 | ~x10" + and "~x43 | ~x37" + and "~x10 | ~x37" + and "~x12 | ~x44" + and "~x12 | ~x11" + and "~x12 | ~x38" + and "~x44 | ~x11" + and "~x44 | ~x38" + and "~x11 | ~x38" + and "~x13 | ~x45" + and "~x13 | ~x12" + and "~x13 | ~x39" + and "~x45 | ~x12" + and "~x45 | ~x39" + and "~x12 | ~x39" + and "~x14 | ~x46" + and "~x14 | ~x13" + and "~x14 | ~x40" + and "~x46 | ~x13" + and "~x46 | ~x40" + and "~x13 | ~x40" + and "~x47 | ~x14" + and "~x47 | ~x41" + and "~x14 | ~x41" + and "~x15 | ~x48" + and "~x15 | ~x42" + and "~x48 | ~x42" + and "~x16 | ~x49" + and "~x16 | ~x15" + and "~x16 | ~x43" + and "~x49 | ~x15" + and "~x49 | ~x43" + and "~x15 | ~x43" + and "~x17 | ~x50" + and "~x17 | ~x16" + and "~x17 | ~x44" + and "~x50 | ~x16" + and "~x50 | ~x44" + and "~x16 | ~x44" + and "~x18 | ~x51" + and "~x18 | ~x17" + and "~x18 | ~x45" + and "~x51 | ~x17" + and "~x51 | ~x45" + and "~x17 | ~x45" + and "~x19 | ~x52" + and "~x19 | ~x18" + and "~x19 | ~x46" + and "~x52 | ~x18" + and "~x52 | ~x46" + and "~x18 | ~x46" + and "~x53 | ~x19" + and "~x53 | ~x47" + and "~x19 | ~x47" + and "~x20 | ~x54" + and "~x20 | ~x48" + and "~x54 | ~x48" + and "~x21 | ~x55" + and "~x21 | ~x20" + and "~x21 | ~x49" + and "~x55 | ~x20" + and "~x55 | ~x49" + and "~x20 | ~x49" + and "~x22 | ~x56" + and "~x22 | ~x21" + and "~x22 | ~x50" + and "~x56 | ~x21" + and "~x56 | ~x50" + and "~x21 | ~x50" + and "~x23 | ~x57" + and "~x23 | ~x22" + and "~x23 | ~x51" + and "~x57 | ~x22" + and "~x57 | ~x51" + and "~x22 | ~x51" + and "~x24 | ~x58" + and "~x24 | ~x23" + and "~x24 | ~x52" + and "~x58 | ~x23" + and "~x58 | ~x52" + and "~x23 | ~x52" + and "~x59 | ~x24" + and "~x59 | ~x53" + and "~x24 | ~x53" + and "~x25 | ~x54" + and "~x26 | ~x25" + and "~x26 | ~x55" + and "~x25 | ~x55" + and "~x27 | ~x26" + and "~x27 | ~x56" + and "~x26 | ~x56" + and "~x28 | ~x27" + and "~x28 | ~x57" + and "~x27 | ~x57" + and "~x29 | ~x28" + and "~x29 | ~x58" + and "~x28 | ~x58" + shows False + using assms + using [[smt_cert="~/isabelle/SMT/Examples/cert/z3_prop_10"]] + by smt + +lemma "\x::int. P x \ (\y::int. P x \ P y)" + using [[smt_cert="~/isabelle/SMT/Examples/cert/z3_fol_01"]] + by smt + +lemma + assumes "(\x y. P x y = x)" + shows "(\y. P x y) = P x c" + using assms + using [[smt_cert="~/isabelle/SMT/Examples/cert/z3_fol_02"]] + by smt + +lemma + assumes "(\x y. P x y = x)" + and "(\x. \y. P x y) = (\x. P x c)" + shows "(EX y. P x y) = P x c" + using assms + using [[smt_cert="~/isabelle/SMT/Examples/cert/z3_fol_03"]] + by smt + +lemma + assumes "if P x then \(\y. P y) else (\y. \P y)" + shows "P x \ P y" + using assms + using [[smt_cert="~/isabelle/SMT/Examples/cert/z3_fol_04"]] + by smt -section {* Linear arithmetic *} +section {* Arithmetic *} + +subsection {* Linear arithmetic over integers and reals *} + +lemma "(3::int) = 3" + using [[smt_cert="~/isabelle/SMT/Examples/cert/z3_linarith_01"]] + by smt + +lemma "(3::real) = 3" + using [[smt_cert="~/isabelle/SMT/Examples/cert/z3_linarith_02"]] + by smt + +lemma "(3 :: int) + 1 = 4" + using [[smt_cert="~/isabelle/SMT/Examples/cert/z3_linarith_03"]] + by smt + +lemma "x + (y + z) = y + (z + (x::int))" + using [[smt_cert="~/isabelle/SMT/Examples/cert/z3_linarith_04"]] + by smt + +lemma "max (3::int) 8 > 5" + using [[smt_cert="~/isabelle/SMT/Examples/cert/z3_linarith_05"]] + by smt + +lemma "abs (x :: real) + abs y \ abs (x + y)" + using [[smt_cert="~/isabelle/SMT/Examples/cert/z3_linarith_06"]] + by smt + +lemma "P ((2::int) < 3) = P True" + using [[smt_cert="~/isabelle/SMT/Examples/cert/z3_linarith_07"]] + by smt + +lemma "x + 3 \ 4 \ x < (1::int)" + using [[smt_cert="~/isabelle/SMT/Examples/cert/z3_linarith_08"]] + by smt -lemma "(3::int) = 3" by smt -lemma "(3::real) = 3" by smt -lemma "(3 :: int) + 1 = 4" by smt -lemma "max (3::int) 8 > 5" by smt -lemma "abs (x :: real) + abs y \ abs (x + y)" by smt -lemma "let x = (2 :: int) in x + x \ 5" by smt +lemma + assumes "x \ (3::int)" and "y = x + 4" + shows "y - x > 0" + using assms + using [[smt_cert="~/isabelle/SMT/Examples/cert/z3_linarith_09"]] + by smt + +lemma "let x = (2 :: int) in x + x \ 5" + using [[smt_cert="~/isabelle/SMT/Examples/cert/z3_linarith_10"]] + by smt + +lemma + fixes x :: real + assumes "3 * x + 7 * a < 4" and "3 < 2 * x" + shows "a < 0" + using assms + using [[smt_cert="~/isabelle/SMT/Examples/cert/z3_linarith_11"]] + by smt + +lemma "(0 \ y + -1 * x \ \ 0 \ x \ 0 \ (x::int)) = (\ False)" + using [[smt_cert="~/isabelle/SMT/Examples/cert/z3_linarith_12"]] + by smt + +lemma "distinct [x < (3::int), 3 \ x]" + using [[smt_cert="~/isabelle/SMT/Examples/cert/z3_linarith_13"]] + by smt + +lemma + assumes "a > (0::int)" + shows "distinct [a, a * 2, a - a]" + using assms + using [[smt_cert="~/isabelle/SMT/Examples/cert/z3_linarith_14"]] + by smt + +lemma " + (n < m & m < n') | (n < m & m = n') | (n < n' & n' < m) | + (n = n' & n' < m) | (n = m & m < n') | + (n' < m & m < n) | (n' < m & m = n) | + (n' < n & n < m) | (n' = n & n < m) | (n' = m & m < n) | + (m < n & n < n') | (m < n & n' = n) | (m < n' & n' < n) | + (m = n & n < n') | (m = n' & n' < n) | + (n' = m & m = (n::int))" + using [[smt_cert="~/isabelle/SMT/Examples/cert/z3_linarith_15"]] + by smt text{* The following example was taken from HOL/ex/PresburgerEx.thy, where it says: @@ -47,107 +373,333 @@ Warning: it takes (in 2006) over 4.2 minutes! There, it is proved by "arith". SMT is able to prove this within a fraction -of one second. +of one second. With proof reconstruction, it takes about 13 seconds on a Core2 +processor. *} lemma "\ x3 = abs x2 - x1; x4 = abs x3 - x2; x5 = abs x4 - x3; x6 = abs x5 - x4; x7 = abs x6 - x5; x8 = abs x7 - x6; x9 = abs x8 - x7; x10 = abs x9 - x8; x11 = abs x10 - x9 \ \ x1 = x10 & x2 = (x11::int)" + using [[smt_cert="~/isabelle/SMT/Examples/cert/z3_linarith_16"]] + by smt + + +subsection {* Linear arithmetic with quantifiers *} + +lemma "~ (\x::int. False)" + using [[smt_cert="~/isabelle/SMT/Examples/cert/z3_arith_quant_01"]] + by smt + +lemma "~ (\x::real. False)" + using [[smt_cert="~/isabelle/SMT/Examples/cert/z3_arith_quant_02"]] + by smt + +lemma "\x::int. 0 < x" + using [[smt_cert="~/isabelle/SMT/Examples/cert/z3_arith_quant_03"]] + using [[z3_proofs=false]] (* no Z3 proof *) + by smt + +lemma "\x::real. 0 < x" + using [[smt_cert="~/isabelle/SMT/Examples/cert/z3_arith_quant_04"]] + using [[z3_proofs=false]] (* no Z3 proof *) + by smt + +lemma "\x::int. \y. y > x" + using [[smt_cert="~/isabelle/SMT/Examples/cert/z3_arith_quant_05"]] + using [[z3_proofs=false]] (* no Z3 proof *) + by smt + +lemma "\x y::int. (x = 0 \ y = 1) \ x \ y" + using [[smt_cert="~/isabelle/SMT/Examples/cert/z3_arith_quant_06"]] + by smt + +lemma "\x::int. \y. x < y \ y < 0 \ y >= 0" + using [[smt_cert="~/isabelle/SMT/Examples/cert/z3_arith_quant_07"]] + by smt + +lemma "\x y::int. x < y \ (2 * x + 1) < (2 * y)" + using [[smt_cert="~/isabelle/SMT/Examples/cert/z3_arith_quant_08"]] + by smt + +lemma "\x y::int. (2 * x + 1) \ (2 * y)" + using [[smt_cert="~/isabelle/SMT/Examples/cert/z3_arith_quant_09"]] + by smt + +lemma "\x y::int. x + y > 2 \ x + y = 2 \ x + y < 2" + using [[smt_cert="~/isabelle/SMT/Examples/cert/z3_arith_quant_10"]] by smt -lemma "\x::int. 0 < x" by smt -lemma "\x::real. 0 < x" by smt -lemma "\x y::int. x < y \ (2 * x + 1) < (2 * y)" by smt -lemma "\x y::int. (2 * x + 1) \ (2 * y)" by smt -lemma "~ (\x y z::int. 4 * x + -6 * y = (1::int))" by smt -lemma "~ (\x::int. False)" by smt +lemma "\x::int. if x > 0 then x + 1 > 0 else 1 > x" + using [[smt_cert="~/isabelle/SMT/Examples/cert/z3_arith_quant_11"]] + by smt + +lemma "if (ALL x::int. x < 0 \ x > 0) then False else True" + using [[smt_cert="~/isabelle/SMT/Examples/cert/z3_arith_quant_12"]] + by smt + +lemma "(if (ALL x::int. x < 0 \ x > 0) then -1 else 3) > (0::int)" + using [[smt_cert="~/isabelle/SMT/Examples/cert/z3_arith_quant_13"]] + by smt + +lemma "~ (\x y z::int. 4 * x + -6 * y = (1::int))" + using [[smt_cert="~/isabelle/SMT/Examples/cert/z3_arith_quant_14"]] + by smt + +lemma "\x::int. \x y. 0 < x \ 0 < y \ (0::int) < x + y" + using [[smt_cert="~/isabelle/SMT/Examples/cert/z3_arith_quant_15"]] + by smt + +lemma "\u::int. \(x::int) y::real. 0 < x \ 0 < y \ -1 < x" + using [[smt_cert="~/isabelle/SMT/Examples/cert/z3_arith_quant_16"]] + by smt + +lemma "\x::int. (\y. y \ x \ y > 0) \ x > 0" + using [[smt_cert="~/isabelle/SMT/Examples/cert/z3_arith_quant_17"]] + by smt + +lemma "\x::int. trigger [pat x] (x < a \ 2 * x < 2 * a)" + using [[smt_cert="~/isabelle/SMT/Examples/cert/z3_arith_quant_18"]] + by smt -section {* Non-linear arithmetic *} +subsection {* Non-linear arithmetic over integers and reals *} + +lemma "a > (0::int) \ a*b > 0 \ b > 0" + using [[smt_cert="~/isabelle/SMT/Examples/cert/z3_nlarith_01"]] + using [[z3_proofs=false]] -- {* Isabelle's arithmetic decision procedures + are too weak to automatically prove @{thm zero_less_mult_pos}. *} + by smt -lemma "((x::int) * (1 + y) - x * (1 - y)) = (2 * x * y)" by smt +lemma "(a::int) * (x + 1 + y) = a * x + a * (y + 1)" + using [[smt_cert="~/isabelle/SMT/Examples/cert/z3_nlarith_02"]] + by smt + +lemma "((x::real) * (1 + y) - x * (1 - y)) = (2 * x * y)" + using [[smt_cert="~/isabelle/SMT/Examples/cert/z3_nlarith_03"]] + by smt + lemma "(U::int) + (1 + p) * (b + e) + p * d = U + (2 * (1 + p) * (b + e) + (1 + p) * d + d * p) - (1 + p) * (b + d + e)" + using [[smt_cert="~/isabelle/SMT/Examples/cert/z3_nlarith_04"]] by smt -section {* Linear arithmetic for natural numbers *} +subsection {* Linear arithmetic for natural numbers *} + +lemma "2 * (x::nat) ~= 1" + using [[smt_cert="~/isabelle/SMT/Examples/cert/z3_nat_arith_01"]] + by smt -lemma "a < 3 \ (7::nat) > 2 * a" by smt -lemma "let x = (1::nat) + y in x - y > 0 * x" by smt +lemma "a < 3 \ (7::nat) > 2 * a" + using [[smt_cert="~/isabelle/SMT/Examples/cert/z3_nat_arith_02"]] + by smt + +lemma "let x = (1::nat) + y in x - y > 0 * x" + using [[smt_cert="~/isabelle/SMT/Examples/cert/z3_nat_arith_03"]] + by smt + lemma "let x = (1::nat) + y in let P = (if x > 0 then True else False) in False \ P = (x - 1 = y) \ (\P \ False)" + using [[smt_cert="~/isabelle/SMT/Examples/cert/z3_nat_arith_04"]] by smt +lemma "distinct [a + (1::nat), a * 2 + 3, a - a]" + using [[smt_cert="~/isabelle/SMT/Examples/cert/z3_nat_arith_05"]] + by smt + +lemma "int (nat \x::int\) = \x\" + using [[smt_cert="~/isabelle/SMT/Examples/cert/z3_nat_arith_06"]] + by smt + +definition prime_nat :: "nat \ bool" where + "prime_nat p = (1 < p \ (\m. m dvd p --> m = 1 \ m = p))" +lemma "prime_nat (4*m + 1) \ m \ (1::nat)" + using [[smt_cert="~/isabelle/SMT/Examples/cert/z3_nat_arith_07"]] + by (smt add: prime_nat_def) + section {* Bitvectors *} locale bv begin -declare [[smt_solver=z3]] +text {* +The following examples only work for Z3, and only without proof reconstruction. +*} + +declare [[smt_solver=z3, z3_proofs=false]] + + +subsection {* Bitvector arithmetic *} + +lemma "(27 :: 4 word) = -5" + using [[smt_cert="~/isabelle/SMT/Examples/cert/z3_bv_arith_01"]] + by smt + +lemma "(27 :: 4 word) = 11" + using [[smt_cert="~/isabelle/SMT/Examples/cert/z3_bv_arith_02"]] + by smt + +lemma "23 < (27::8 word)" + using [[smt_cert="~/isabelle/SMT/Examples/cert/z3_bv_arith_03"]] + by smt + +lemma "27 + 11 = (6::5 word)" + using [[smt_cert="~/isabelle/SMT/Examples/cert/z3_bv_arith_04"]] + by smt -lemma "(27 :: 4 word) = -5" by smt -lemma "(27 :: 4 word) = 11" by smt -lemma "23 < (27::8 word)" by smt -lemma "27 + 11 = (6::5 word)" by smt -lemma "7 * 3 = (21::8 word)" by smt -lemma "11 - 27 = (-16::8 word)" by smt -lemma "- -11 = (11::5 word)" by smt -lemma "-40 + 1 = (-39::7 word)" by smt -lemma "a + 2 * b + c - b = (b + c) + (a :: 32 word)" by smt +lemma "7 * 3 = (21::8 word)" + using [[smt_cert="~/isabelle/SMT/Examples/cert/z3_bv_arith_05"]] + by smt +lemma "11 - 27 = (-16::8 word)" + using [[smt_cert="~/isabelle/SMT/Examples/cert/z3_bv_arith_06"]] + by smt + +lemma "- -11 = (11::5 word)" + using [[smt_cert="~/isabelle/SMT/Examples/cert/z3_bv_arith_07"]] + by smt + +lemma "-40 + 1 = (-39::7 word)" + using [[smt_cert="~/isabelle/SMT/Examples/cert/z3_bv_arith_08"]] + by smt -lemma "0b110 AND 0b101 = (0b100 :: 32 word)" by smt -lemma "0b110 OR 0b011 = (0b111 :: 8 word)" by smt -lemma "0xF0 XOR 0xFF = (0x0F :: 8 word)" by smt -lemma "NOT (0xF0 :: 16 word) = 0xFF0F" by smt +lemma "a + 2 * b + c - b = (b + c) + (a :: 32 word)" + using [[smt_cert="~/isabelle/SMT/Examples/cert/z3_bv_arith_09"]] + by smt + +lemma "x = (5 :: 4 word) \ 4 * x = 4" + using [[smt_cert="~/isabelle/SMT/Examples/cert/z3_bv_arith_10"]] + by smt + + +subsection {* Bit-level logic *} -lemma "word_cat (27::4 word) (27::8 word) = (2843::12 word)" by smt -lemma "word_cat (0b0011::4 word) (0b1111::6word) = (0b0011001111 :: 10 word)" +lemma "0b110 AND 0b101 = (0b100 :: 32 word)" + using [[smt_cert="~/isabelle/SMT/Examples/cert/z3_bv_bit_01"]] + by smt + +lemma "0b110 OR 0b011 = (0b111 :: 8 word)" + using [[smt_cert="~/isabelle/SMT/Examples/cert/z3_bv_bit_02"]] + by smt + +lemma "0xF0 XOR 0xFF = (0x0F :: 8 word)" + using [[smt_cert="~/isabelle/SMT/Examples/cert/z3_bv_bit_03"]] by smt -lemma "slice 1 (0b10110 :: 4 word) = (0b11 :: 2 word)" by smt +lemma "NOT (0xF0 :: 16 word) = 0xFF0F" + using [[smt_cert="~/isabelle/SMT/Examples/cert/z3_bv_bit_04"]] + by smt + +lemma "word_cat (27::4 word) (27::8 word) = (2843::12 word)" + using [[smt_cert="~/isabelle/SMT/Examples/cert/z3_bv_bit_05"]] + by smt + +lemma "word_cat (0b0011::4 word) (0b1111::6word) = (0b0011001111 :: 10 word)" + using [[smt_cert="~/isabelle/SMT/Examples/cert/z3_bv_bit_06"]] + by smt -lemma "ucast (0b1010 :: 4 word) = (0b1010 :: 10 word)" by smt -lemma "scast (0b1010 :: 4 word) = (0b111010 :: 6 word)" by smt +lemma "slice 1 (0b10110 :: 4 word) = (0b11 :: 2 word)" + using [[smt_cert="~/isabelle/SMT/Examples/cert/z3_bv_bit_07"]] + by smt + +lemma "ucast (0b1010 :: 4 word) = (0b1010 :: 10 word)" + using [[smt_cert="~/isabelle/SMT/Examples/cert/z3_bv_bit_08"]] + by smt + +lemma "scast (0b1010 :: 4 word) = (0b111010 :: 6 word)" + using [[smt_cert="~/isabelle/SMT/Examples/cert/z3_bv_bit_09"]] + by smt -lemma "bv_lshr 0b10011 2 = (0b100::8 word)" by smt -lemma "bv_ashr 0b10011 2 = (0b100::8 word)" by smt +lemma "bv_lshr 0b10011 2 = (0b100::8 word)" + using [[smt_cert="~/isabelle/SMT/Examples/cert/z3_bv_bit_10"]] + by smt + +lemma "bv_ashr 0b10011 2 = (0b100::8 word)" + using [[smt_cert="~/isabelle/SMT/Examples/cert/z3_bv_bit_11"]] + by smt + +lemma "word_rotr 2 0b0110 = (0b1001::4 word)" + using [[smt_cert="~/isabelle/SMT/Examples/cert/z3_bv_bit_12"]] + by smt -lemma "word_rotr 2 0b0110 = (0b1001::4 word)" by smt -lemma "word_rotl 1 0b1110 = (0b1101::4 word)" by smt +lemma "word_rotl 1 0b1110 = (0b1101::4 word)" + using [[smt_cert="~/isabelle/SMT/Examples/cert/z3_bv_bit_13"]] + by smt -lemma "(x AND 0xff00) OR (x AND 0x00ff) = (x::16 word)" by smt +lemma "(x AND 0xff00) OR (x AND 0x00ff) = (x::16 word)" + using [[smt_cert="~/isabelle/SMT/Examples/cert/z3_bv_bit_14"]] + by smt -lemma "w < 256 \ (w :: 16 word) AND 0x00FF = w" by smt +lemma "w < 256 \ (w :: 16 word) AND 0x00FF = w" + using [[smt_cert="~/isabelle/SMT/Examples/cert/z3_bv_bit_15"]] + by smt end +lemma + assumes "bv2int 0 = 0" + and "bv2int 1 = 1" + and "bv2int 2 = 2" + and "bv2int 3 = 3" + and "\x::2 word. bv2int x > 0" + shows "\i::int. i < 0 \ (\x::2 word. bv2int x > i)" + using assms + using [[smt_solver=z3]] + using [[smt_cert="~/isabelle/SMT/Examples/cert/z3_bv_01"]] + by smt + +lemma "P (0 \ (a :: 4 word)) = P True" + using [[smt_solver=z3, z3_proofs=false]] + using [[smt_cert="~/isabelle/SMT/Examples/cert/z3_bv_02"]] + by smt + section {* Pairs *} -lemma "fst (x, y) = a \ x = a" by smt -lemma "p1 = (x, y) \ p2 = (y, x) \ fst p1 = snd p2" by smt +lemma "fst (x, y) = a \ x = a" + using [[smt_cert="~/isabelle/SMT/Examples/cert/z3_pair_01"]] + by smt + +lemma "p1 = (x, y) \ p2 = (y, x) \ fst p1 = snd p2" + using [[smt_cert="~/isabelle/SMT/Examples/cert/z3_pair_02"]] + by smt section {* Higher-order problems and recursion *} -lemma "(f g x = (g x \ True)) \ (f g x = True) \ (g x = True)" by smt -lemma "P ((2::int) < 3) = P True" by smt -lemma "P ((2::int) < 3) = (P True :: bool)" by smt -lemma "P (0 \ (a :: 4 word)) = P True" using [[smt_solver=z3]] by smt -lemma "id 3 = 3 \ id True = True" by (smt add: id_def) -lemma "i \ i1 \ i \ i2 \ ((f (i1 := v1)) (i2 := v2)) i = f i" by smt -lemma "map (\i::nat. i + 1) [0, 1] = [1, 2]" by (smt add: map.simps) -lemma "(ALL x. P x) | ~ All P" by smt +lemma "i \ i1 \ i \ i2 \ (f (i1 := v1, i2 := v2)) i = f i" + using [[smt_cert="~/isabelle/SMT/Examples/cert/z3_hol_01"]] + by smt + +lemma "(f g x = (g x \ True)) \ (f g x = True) \ (g x = True)" + using [[smt_cert="~/isabelle/SMT/Examples/cert/z3_hol_02"]] + by smt + +lemma "id 3 = 3 \ id True = True" + using [[smt_cert="~/isabelle/SMT/Examples/cert/z3_hol_03"]] + by (smt add: id_def) + +lemma "i \ i1 \ i \ i2 \ ((f (i1 := v1)) (i2 := v2)) i = f i" + using [[smt_cert="~/isabelle/SMT/Examples/cert/z3_hol_04"]] + by smt + +lemma "map (\i::nat. i + 1) [0, 1] = [1, 2]" + using [[smt_cert="~/isabelle/SMT/Examples/cert/z3_hol_05"]] + by (smt add: map.simps) + +lemma "(ALL x. P x) | ~ All P" + using [[smt_cert="~/isabelle/SMT/Examples/cert/z3_hol_06"]] + by smt fun dec_10 :: "nat \ nat" where "dec_10 n = (if n < 10 then n else dec_10 (n - 10))" -lemma "dec_10 (4 * dec_10 4) = 6" by (smt add: dec_10.simps) +lemma "dec_10 (4 * dec_10 4) = 6" + using [[smt_cert="~/isabelle/SMT/Examples/cert/z3_hol_07"]] + by (smt add: dec_10.simps) axiomatization eval_dioph :: "int list \ nat list \ int" @@ -163,7 +715,7 @@ (eval_dioph ks (map (\x. x mod 2) xs) mod 2 = l mod 2 \ eval_dioph ks (map (\x. x div 2) xs) = (l - eval_dioph ks (map (\x. x mod 2) xs)) div 2)" - using [[smt_solver=z3]] + using [[smt_cert="~/isabelle/SMT/Examples/cert/z3_hol_08"]] by (smt add: eval_dioph_mod[where n=2] eval_dioph_div_mult[where n=2]) @@ -171,13 +723,17 @@ definition P :: "'a \ bool" where "P x = True" lemma poly_P: "P x \ (P [x] \ \P[x])" by (simp add: P_def) -lemma "P (1::int)" by (smt add: poly_P) +lemma "P (1::int)" + using [[smt_cert="~/isabelle/SMT/Examples/cert/z3_mono_01"]] + by (smt add: poly_P) consts g :: "'a \ nat" axioms g1: "g (Some x) = g [x]" g2: "g None = g []" g3: "g xs = length xs" -lemma "g (Some (3::int)) = g (Some True)" by (smt add: g1 g2 g3 list.size) +lemma "g (Some (3::int)) = g (Some True)" + using [[smt_cert="~/isabelle/SMT/Examples/cert/z3_mono_02"]] + by (smt add: g1 g2 g3 list.size) end diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_arith_quant_01 --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_arith_quant_01 Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,4 @@ +(benchmark Isabelle +:assumption (not (not (exists (?x1 Int) false))) +:formula true +) diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_arith_quant_01.proof --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_arith_quant_01.proof Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,24 @@ +#2 := false +#4 := (exists (vars (?x1 int)) false) +#5 := (not #4) +#6 := (not #5) +#37 := (iff #6 false) +#1 := true +#32 := (not true) +#35 := (iff #32 false) +#36 := [rewrite]: #35 +#33 := (iff #6 #32) +#30 := (iff #5 true) +#25 := (not false) +#28 := (iff #25 true) +#29 := [rewrite]: #28 +#26 := (iff #5 #25) +#23 := (iff #4 false) +#24 := [elim-unused]: #23 +#27 := [monotonicity #24]: #26 +#31 := [trans #27 #29]: #30 +#34 := [monotonicity #31]: #33 +#38 := [trans #34 #36]: #37 +#22 := [asserted]: #6 +[mp #22 #38]: false +unsat diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_arith_quant_02 --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_arith_quant_02 Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,4 @@ +(benchmark Isabelle +:assumption (not (not (exists (?x1 Real) false))) +:formula true +) diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_arith_quant_02.proof --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_arith_quant_02.proof Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,24 @@ +#2 := false +#4 := (exists (vars (?x1 real)) false) +#5 := (not #4) +#6 := (not #5) +#37 := (iff #6 false) +#1 := true +#32 := (not true) +#35 := (iff #32 false) +#36 := [rewrite]: #35 +#33 := (iff #6 #32) +#30 := (iff #5 true) +#25 := (not false) +#28 := (iff #25 true) +#29 := [rewrite]: #28 +#26 := (iff #5 #25) +#23 := (iff #4 false) +#24 := [elim-unused]: #23 +#27 := [monotonicity #24]: #26 +#31 := [trans #27 #29]: #30 +#34 := [monotonicity #31]: #33 +#38 := [trans #34 #36]: #37 +#22 := [asserted]: #6 +[mp #22 #38]: false +unsat diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_arith_quant_03 --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_arith_quant_03 Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,4 @@ +(benchmark Isabelle +:assumption (not (exists (?x1 Int) (< 0 ?x1))) +:formula true +) diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_arith_quant_03.proof --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_arith_quant_03.proof Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,1 @@ +unsat diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_arith_quant_04 --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_arith_quant_04 Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,4 @@ +(benchmark Isabelle +:assumption (not (exists (?x1 Real) (< 0.0 ?x1))) +:formula true +) diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_arith_quant_04.proof --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_arith_quant_04.proof Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,1 @@ +unsat diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_arith_quant_05 --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_arith_quant_05 Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,4 @@ +(benchmark Isabelle +:assumption (not (forall (?x1 Int) (exists (?x2 Int) (< ?x1 ?x2)))) +:formula true +) diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_arith_quant_05.proof --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_arith_quant_05.proof Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,1 @@ +unsat diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_arith_quant_06 --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_arith_quant_06 Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,4 @@ +(benchmark Isabelle +:assumption (not (forall (?x1 Int) (?x2 Int) (implies (and (= ?x1 0) (= ?x2 1)) (not (= ?x1 ?x2))))) +:formula true +) diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_arith_quant_06.proof --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_arith_quant_06.proof Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,73 @@ +#2 := false +#5 := 0::int +#8 := 1::int +#143 := (= 1::int 0::int) +#145 := (iff #143 false) +#146 := [rewrite]: #145 +decl ?x1!1 :: int +#47 := ?x1!1 +#51 := (= ?x1!1 0::int) +decl ?x2!0 :: int +#46 := ?x2!0 +#50 := (= ?x2!0 1::int) +#63 := (and #50 #51) +#69 := (= ?x2!0 ?x1!1) +#72 := (not #69) +#66 := (not #63) +#75 := (or #66 #72) +#78 := (not #75) +#48 := (= ?x1!1 ?x2!0) +#49 := (not #48) +#52 := (and #51 #50) +#53 := (not #52) +#54 := (or #53 #49) +#55 := (not #54) +#79 := (iff #55 #78) +#76 := (iff #54 #75) +#73 := (iff #49 #72) +#70 := (iff #48 #69) +#71 := [rewrite]: #70 +#74 := [monotonicity #71]: #73 +#67 := (iff #53 #66) +#64 := (iff #52 #63) +#65 := [rewrite]: #64 +#68 := [monotonicity #65]: #67 +#77 := [monotonicity #68 #74]: #76 +#80 := [monotonicity #77]: #79 +#7 := (:var 0 int) +#4 := (:var 1 int) +#11 := (= #4 #7) +#12 := (not #11) +#9 := (= #7 1::int) +#6 := (= #4 0::int) +#10 := (and #6 #9) +#32 := (not #10) +#33 := (or #32 #12) +#36 := (forall (vars (?x1 int) (?x2 int)) #33) +#39 := (not #36) +#56 := (~ #39 #55) +#57 := [sk]: #56 +#13 := (implies #10 #12) +#14 := (forall (vars (?x1 int) (?x2 int)) #13) +#15 := (not #14) +#40 := (iff #15 #39) +#37 := (iff #14 #36) +#34 := (iff #13 #33) +#35 := [rewrite]: #34 +#38 := [quant-intro #35]: #37 +#41 := [monotonicity #38]: #40 +#31 := [asserted]: #15 +#44 := [mp #31 #41]: #39 +#60 := [mp~ #44 #57]: #55 +#61 := [mp #60 #80]: #78 +#62 := [not-or-elim #61]: #63 +#82 := [and-elim #62]: #51 +#141 := (= 1::int ?x1!1) +#83 := [not-or-elim #61]: #69 +#139 := (= 1::int ?x2!0) +#81 := [and-elim #62]: #50 +#140 := [symm #81]: #139 +#142 := [trans #140 #83]: #141 +#144 := [trans #142 #82]: #143 +[mp #144 #146]: false +unsat diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_arith_quant_07 --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_arith_quant_07 Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,4 @@ +(benchmark Isabelle +:assumption (not (exists (?x1 Int) (forall (?x2 Int) (implies (< ?x1 ?x2) (or (< ?x2 0) (<= 0 ?x2)))))) +:formula true +) diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_arith_quant_07.proof --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_arith_quant_07.proof Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,82 @@ +#2 := false +#5 := (:var 0 int) +#7 := 0::int +#9 := (<= 0::int #5) +#8 := (< #5 0::int) +#10 := (or #8 #9) +#4 := (:var 1 int) +#6 := (< #4 #5) +#11 := (implies #6 #10) +#12 := (forall (vars (?x2 int)) #11) +#13 := (exists (vars (?x1 int)) #12) +#14 := (not #13) +#95 := (iff #14 false) +#31 := (not #6) +#32 := (or #31 #10) +#35 := (forall (vars (?x2 int)) #32) +#38 := (exists (vars (?x1 int)) #35) +#41 := (not #38) +#93 := (iff #41 false) +#1 := true +#88 := (not true) +#91 := (iff #88 false) +#92 := [rewrite]: #91 +#89 := (iff #41 #88) +#86 := (iff #38 true) +#81 := (exists (vars (?x1 int)) true) +#84 := (iff #81 true) +#85 := [elim-unused]: #84 +#82 := (iff #38 #81) +#79 := (iff #35 true) +#74 := (forall (vars (?x2 int)) true) +#77 := (iff #74 true) +#78 := [elim-unused]: #77 +#75 := (iff #35 #74) +#72 := (iff #32 true) +#46 := (>= #5 0::int) +#44 := (not #46) +#64 := (or #44 #46) +#50 := -1::int +#53 := (* -1::int #5) +#54 := (+ #4 #53) +#52 := (>= #54 0::int) +#67 := (or #52 #64) +#70 := (iff #67 true) +#71 := [rewrite]: #70 +#68 := (iff #32 #67) +#65 := (iff #10 #64) +#48 := (iff #9 #46) +#49 := [rewrite]: #48 +#45 := (iff #8 #44) +#47 := [rewrite]: #45 +#66 := [monotonicity #47 #49]: #65 +#62 := (iff #31 #52) +#51 := (not #52) +#57 := (not #51) +#60 := (iff #57 #52) +#61 := [rewrite]: #60 +#58 := (iff #31 #57) +#55 := (iff #6 #51) +#56 := [rewrite]: #55 +#59 := [monotonicity #56]: #58 +#63 := [trans #59 #61]: #62 +#69 := [monotonicity #63 #66]: #68 +#73 := [trans #69 #71]: #72 +#76 := [quant-intro #73]: #75 +#80 := [trans #76 #78]: #79 +#83 := [quant-intro #80]: #82 +#87 := [trans #83 #85]: #86 +#90 := [monotonicity #87]: #89 +#94 := [trans #90 #92]: #93 +#42 := (iff #14 #41) +#39 := (iff #13 #38) +#36 := (iff #12 #35) +#33 := (iff #11 #32) +#34 := [rewrite]: #33 +#37 := [quant-intro #34]: #36 +#40 := [quant-intro #37]: #39 +#43 := [monotonicity #40]: #42 +#96 := [trans #43 #94]: #95 +#30 := [asserted]: #14 +[mp #30 #96]: false +unsat diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_arith_quant_08 --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_arith_quant_08 Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,4 @@ +(benchmark Isabelle +:assumption (not (forall (?x1 Int) (?x2 Int) (implies (< ?x1 ?x2) (< (+ (* 2 ?x1) 1) (* 2 ?x2))))) +:formula true +) diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_arith_quant_08.proof --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_arith_quant_08.proof Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,128 @@ +#2 := false +#9 := 1::int +decl ?x1!1 :: int +#91 := ?x1!1 +#68 := -2::int +#129 := (* -2::int ?x1!1) +decl ?x2!0 :: int +#90 := ?x2!0 +#7 := 2::int +#128 := (* 2::int ?x2!0) +#130 := (+ #128 #129) +#131 := (<= #130 1::int) +#136 := (not #131) +#55 := 0::int +#53 := -1::int +#115 := (* -1::int ?x1!1) +#116 := (+ ?x2!0 #115) +#117 := (<= #116 0::int) +#139 := (or #117 #136) +#142 := (not #139) +#92 := (* -2::int ?x2!0) +#93 := (* 2::int ?x1!1) +#94 := (+ #93 #92) +#95 := (>= #94 -1::int) +#96 := (not #95) +#97 := (* -1::int ?x2!0) +#98 := (+ ?x1!1 #97) +#99 := (>= #98 0::int) +#100 := (or #99 #96) +#101 := (not #100) +#143 := (iff #101 #142) +#140 := (iff #100 #139) +#137 := (iff #96 #136) +#134 := (iff #95 #131) +#122 := (+ #92 #93) +#125 := (>= #122 -1::int) +#132 := (iff #125 #131) +#133 := [rewrite]: #132 +#126 := (iff #95 #125) +#123 := (= #94 #122) +#124 := [rewrite]: #123 +#127 := [monotonicity #124]: #126 +#135 := [trans #127 #133]: #134 +#138 := [monotonicity #135]: #137 +#120 := (iff #99 #117) +#109 := (+ #97 ?x1!1) +#112 := (>= #109 0::int) +#118 := (iff #112 #117) +#119 := [rewrite]: #118 +#113 := (iff #99 #112) +#110 := (= #98 #109) +#111 := [rewrite]: #110 +#114 := [monotonicity #111]: #113 +#121 := [trans #114 #119]: #120 +#141 := [monotonicity #121 #138]: #140 +#144 := [monotonicity #141]: #143 +#5 := (:var 0 int) +#71 := (* -2::int #5) +#4 := (:var 1 int) +#8 := (* 2::int #4) +#72 := (+ #8 #71) +#70 := (>= #72 -1::int) +#69 := (not #70) +#57 := (* -1::int #5) +#58 := (+ #4 #57) +#56 := (>= #58 0::int) +#75 := (or #56 #69) +#78 := (forall (vars (?x1 int) (?x2 int)) #75) +#81 := (not #78) +#102 := (~ #81 #101) +#103 := [sk]: #102 +#11 := (* 2::int #5) +#10 := (+ #8 1::int) +#12 := (< #10 #11) +#6 := (< #4 #5) +#13 := (implies #6 #12) +#14 := (forall (vars (?x1 int) (?x2 int)) #13) +#15 := (not #14) +#84 := (iff #15 #81) +#32 := (+ 1::int #8) +#35 := (< #32 #11) +#41 := (not #6) +#42 := (or #41 #35) +#47 := (forall (vars (?x1 int) (?x2 int)) #42) +#50 := (not #47) +#82 := (iff #50 #81) +#79 := (iff #47 #78) +#76 := (iff #42 #75) +#73 := (iff #35 #69) +#74 := [rewrite]: #73 +#66 := (iff #41 #56) +#54 := (not #56) +#61 := (not #54) +#64 := (iff #61 #56) +#65 := [rewrite]: #64 +#62 := (iff #41 #61) +#59 := (iff #6 #54) +#60 := [rewrite]: #59 +#63 := [monotonicity #60]: #62 +#67 := [trans #63 #65]: #66 +#77 := [monotonicity #67 #74]: #76 +#80 := [quant-intro #77]: #79 +#83 := [monotonicity #80]: #82 +#51 := (iff #15 #50) +#48 := (iff #14 #47) +#45 := (iff #13 #42) +#38 := (implies #6 #35) +#43 := (iff #38 #42) +#44 := [rewrite]: #43 +#39 := (iff #13 #38) +#36 := (iff #12 #35) +#33 := (= #10 #32) +#34 := [rewrite]: #33 +#37 := [monotonicity #34]: #36 +#40 := [monotonicity #37]: #39 +#46 := [trans #40 #44]: #45 +#49 := [quant-intro #46]: #48 +#52 := [monotonicity #49]: #51 +#85 := [trans #52 #83]: #84 +#31 := [asserted]: #15 +#86 := [mp #31 #85]: #81 +#106 := [mp~ #86 #103]: #101 +#107 := [mp #106 #144]: #142 +#146 := [not-or-elim #107]: #131 +#108 := (not #117) +#145 := [not-or-elim #107]: #108 +[th-lemma #145 #146]: false +unsat diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_arith_quant_09 --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_arith_quant_09 Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,4 @@ +(benchmark Isabelle +:assumption (not (forall (?x1 Int) (?x2 Int) (not (= (+ (* 2 ?x1) 1) (* 2 ?x2))))) +:formula true +) diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_arith_quant_09.proof --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_arith_quant_09.proof Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,100 @@ +#2 := false +#7 := 1::int +decl ?x1!1 :: int +#74 := ?x1!1 +#51 := -2::int +#96 := (* -2::int ?x1!1) +decl ?x2!0 :: int +#73 := ?x2!0 +#4 := 2::int +#95 := (* 2::int ?x2!0) +#97 := (+ #95 #96) +#166 := (<= #97 1::int) +#94 := (= #97 1::int) +#53 := -1::int +#75 := (* -2::int ?x2!0) +#76 := (* 2::int ?x1!1) +#77 := (+ #76 #75) +#78 := (= #77 -1::int) +#79 := (not #78) +#80 := (not #79) +#110 := (iff #80 #94) +#102 := (not #94) +#105 := (not #102) +#108 := (iff #105 #94) +#109 := [rewrite]: #108 +#106 := (iff #80 #105) +#103 := (iff #79 #102) +#100 := (iff #78 #94) +#88 := (+ #75 #76) +#91 := (= #88 -1::int) +#98 := (iff #91 #94) +#99 := [rewrite]: #98 +#92 := (iff #78 #91) +#89 := (= #77 #88) +#90 := [rewrite]: #89 +#93 := [monotonicity #90]: #92 +#101 := [trans #93 #99]: #100 +#104 := [monotonicity #101]: #103 +#107 := [monotonicity #104]: #106 +#111 := [trans #107 #109]: #110 +#9 := (:var 0 int) +#55 := (* -2::int #9) +#5 := (:var 1 int) +#6 := (* 2::int #5) +#56 := (+ #6 #55) +#54 := (= #56 -1::int) +#58 := (not #54) +#61 := (forall (vars (?x1 int) (?x2 int)) #58) +#64 := (not #61) +#81 := (~ #64 #80) +#82 := [sk]: #81 +#10 := (* 2::int #9) +#8 := (+ #6 1::int) +#11 := (= #8 #10) +#12 := (not #11) +#13 := (forall (vars (?x1 int) (?x2 int)) #12) +#14 := (not #13) +#67 := (iff #14 #64) +#31 := (+ 1::int #6) +#37 := (= #10 #31) +#42 := (not #37) +#45 := (forall (vars (?x1 int) (?x2 int)) #42) +#48 := (not #45) +#65 := (iff #48 #64) +#62 := (iff #45 #61) +#59 := (iff #42 #58) +#52 := (iff #37 #54) +#57 := [rewrite]: #52 +#60 := [monotonicity #57]: #59 +#63 := [quant-intro #60]: #62 +#66 := [monotonicity #63]: #65 +#49 := (iff #14 #48) +#46 := (iff #13 #45) +#43 := (iff #12 #42) +#40 := (iff #11 #37) +#34 := (= #31 #10) +#38 := (iff #34 #37) +#39 := [rewrite]: #38 +#35 := (iff #11 #34) +#32 := (= #8 #31) +#33 := [rewrite]: #32 +#36 := [monotonicity #33]: #35 +#41 := [trans #36 #39]: #40 +#44 := [monotonicity #41]: #43 +#47 := [quant-intro #44]: #46 +#50 := [monotonicity #47]: #49 +#68 := [trans #50 #66]: #67 +#30 := [asserted]: #14 +#69 := [mp #30 #68]: #64 +#85 := [mp~ #69 #82]: #80 +#86 := [mp #85 #111]: #94 +#168 := (or #102 #166) +#169 := [th-lemma]: #168 +#170 := [unit-resolution #169 #86]: #166 +#167 := (>= #97 1::int) +#171 := (or #102 #167) +#172 := [th-lemma]: #171 +#173 := [unit-resolution #172 #86]: #167 +[th-lemma #173 #170]: false +unsat diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_arith_quant_10 --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_arith_quant_10 Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,4 @@ +(benchmark Isabelle +:assumption (not (forall (?x1 Int) (?x2 Int) (or (< 2 (+ ?x1 ?x2)) (or (= (+ ?x1 ?x2) 2) (< (+ ?x1 ?x2) 2))))) +:formula true +) diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_arith_quant_10.proof --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_arith_quant_10.proof Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,111 @@ +#2 := false +#4 := 2::int +decl ?x1!1 :: int +#85 := ?x1!1 +decl ?x2!0 :: int +#84 := ?x2!0 +#101 := (+ ?x2!0 ?x1!1) +#107 := (>= #101 2::int) +#113 := (<= #101 2::int) +#116 := (not #113) +#110 := (not #107) +#104 := (= #101 2::int) +#119 := (or #104 #110 #116) +#122 := (not #119) +#86 := (+ ?x1!1 ?x2!0) +#87 := (<= #86 2::int) +#88 := (not #87) +#89 := (>= #86 2::int) +#90 := (not #89) +#91 := (= #86 2::int) +#92 := (or #91 #90 #88) +#93 := (not #92) +#123 := (iff #93 #122) +#120 := (iff #92 #119) +#117 := (iff #88 #116) +#114 := (iff #87 #113) +#102 := (= #86 #101) +#103 := [rewrite]: #102 +#115 := [monotonicity #103]: #114 +#118 := [monotonicity #115]: #117 +#111 := (iff #90 #110) +#108 := (iff #89 #107) +#109 := [monotonicity #103]: #108 +#112 := [monotonicity #109]: #111 +#105 := (iff #91 #104) +#106 := [monotonicity #103]: #105 +#121 := [monotonicity #106 #112 #118]: #120 +#124 := [monotonicity #121]: #123 +#6 := (:var 0 int) +#5 := (:var 1 int) +#7 := (+ #5 #6) +#56 := (<= #7 2::int) +#58 := (not #56) +#54 := (>= #7 2::int) +#51 := (not #54) +#9 := (= #7 2::int) +#67 := (or #9 #51 #58) +#72 := (forall (vars (?x1 int) (?x2 int)) #67) +#75 := (not #72) +#94 := (~ #75 #93) +#95 := [sk]: #94 +#10 := (< #7 2::int) +#11 := (or #9 #10) +#8 := (< 2::int #7) +#12 := (or #8 #11) +#13 := (forall (vars (?x1 int) (?x2 int)) #12) +#14 := (not #13) +#78 := (iff #14 #75) +#31 := (= 2::int #7) +#37 := (or #10 #31) +#42 := (or #8 #37) +#45 := (forall (vars (?x1 int) (?x2 int)) #42) +#48 := (not #45) +#76 := (iff #48 #75) +#73 := (iff #45 #72) +#70 := (iff #42 #67) +#61 := (or #51 #9) +#64 := (or #58 #61) +#68 := (iff #64 #67) +#69 := [rewrite]: #68 +#65 := (iff #42 #64) +#62 := (iff #37 #61) +#55 := (iff #31 #9) +#57 := [rewrite]: #55 +#53 := (iff #10 #51) +#52 := [rewrite]: #53 +#63 := [monotonicity #52 #57]: #62 +#59 := (iff #8 #58) +#60 := [rewrite]: #59 +#66 := [monotonicity #60 #63]: #65 +#71 := [trans #66 #69]: #70 +#74 := [quant-intro #71]: #73 +#77 := [monotonicity #74]: #76 +#49 := (iff #14 #48) +#46 := (iff #13 #45) +#43 := (iff #12 #42) +#40 := (iff #11 #37) +#34 := (or #31 #10) +#38 := (iff #34 #37) +#39 := [rewrite]: #38 +#35 := (iff #11 #34) +#32 := (iff #9 #31) +#33 := [rewrite]: #32 +#36 := [monotonicity #33]: #35 +#41 := [trans #36 #39]: #40 +#44 := [monotonicity #41]: #43 +#47 := [quant-intro #44]: #46 +#50 := [monotonicity #47]: #49 +#79 := [trans #50 #77]: #78 +#30 := [asserted]: #14 +#80 := [mp #30 #79]: #75 +#98 := [mp~ #80 #95]: #93 +#99 := [mp #98 #124]: #122 +#126 := [not-or-elim #99]: #107 +#100 := (not #104) +#125 := [not-or-elim #99]: #100 +#127 := [not-or-elim #99]: #113 +#183 := (or #104 #116 #110) +#184 := [th-lemma]: #183 +[unit-resolution #184 #127 #125 #126]: false +unsat diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_arith_quant_11 --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_arith_quant_11 Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,4 @@ +(benchmark Isabelle +:assumption (not (forall (?x1 Int) (if_then_else (< 0 ?x1) (< 0 (+ ?x1 1)) (< ?x1 1)))) +:formula true +) diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_arith_quant_11.proof --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_arith_quant_11.proof Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,89 @@ +#2 := false +#4 := 0::int +decl ?x1!0 :: int +#78 := ?x1!0 +#83 := (<= ?x1!0 0::int) +#146 := (not #83) +#155 := [hypothesis]: #83 +#7 := 1::int +#81 := (>= ?x1!0 1::int) +#82 := (not #81) +#156 := (or #82 #146) +#157 := [th-lemma]: #156 +#158 := [unit-resolution #157 #155]: #82 +#159 := (or #146 #81) +#49 := -1::int +#79 := (<= ?x1!0 -1::int) +#80 := (not #79) +#84 := (ite #83 #82 #80) +#85 := (not #84) +#5 := (:var 0 int) +#50 := (<= #5 -1::int) +#51 := (not #50) +#55 := (>= #5 1::int) +#54 := (not #55) +#45 := (<= #5 0::int) +#61 := (ite #45 #54 #51) +#66 := (forall (vars (?x1 int)) #61) +#69 := (not #66) +#86 := (~ #69 #85) +#87 := [sk]: #86 +#10 := (< #5 1::int) +#8 := (+ #5 1::int) +#9 := (< 0::int #8) +#6 := (< 0::int #5) +#11 := (ite #6 #9 #10) +#12 := (forall (vars (?x1 int)) #11) +#13 := (not #12) +#72 := (iff #13 #69) +#30 := (+ 1::int #5) +#33 := (< 0::int #30) +#36 := (ite #6 #33 #10) +#39 := (forall (vars (?x1 int)) #36) +#42 := (not #39) +#70 := (iff #42 #69) +#67 := (iff #39 #66) +#64 := (iff #36 #61) +#46 := (not #45) +#58 := (ite #46 #51 #54) +#62 := (iff #58 #61) +#63 := [rewrite]: #62 +#59 := (iff #36 #58) +#56 := (iff #10 #54) +#57 := [rewrite]: #56 +#52 := (iff #33 #51) +#53 := [rewrite]: #52 +#47 := (iff #6 #46) +#48 := [rewrite]: #47 +#60 := [monotonicity #48 #53 #57]: #59 +#65 := [trans #60 #63]: #64 +#68 := [quant-intro #65]: #67 +#71 := [monotonicity #68]: #70 +#43 := (iff #13 #42) +#40 := (iff #12 #39) +#37 := (iff #11 #36) +#34 := (iff #9 #33) +#31 := (= #8 #30) +#32 := [rewrite]: #31 +#35 := [monotonicity #32]: #34 +#38 := [monotonicity #35]: #37 +#41 := [quant-intro #38]: #40 +#44 := [monotonicity #41]: #43 +#73 := [trans #44 #71]: #72 +#29 := [asserted]: #13 +#74 := [mp #29 #73]: #69 +#90 := [mp~ #74 #87]: #85 +#151 := (or #84 #146 #81) +#152 := [def-axiom]: #151 +#160 := [unit-resolution #152 #90]: #159 +#161 := [unit-resolution #160 #158 #155]: false +#162 := [lemma #161]: #146 +#163 := (or #80 #83) +#164 := [th-lemma]: #163 +#165 := [unit-resolution #164 #162]: #80 +#166 := (or #83 #79) +#153 := (or #84 #83 #79) +#154 := [def-axiom]: #153 +#167 := [unit-resolution #154 #90]: #166 +[unit-resolution #167 #165 #162]: false +unsat diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_arith_quant_12 --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_arith_quant_12 Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,4 @@ +(benchmark Isabelle +:assumption (not (if_then_else (forall (?x1 Int) (or (< ?x1 0) (< 0 ?x1))) false true)) +:formula true +) diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_arith_quant_12.proof --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_arith_quant_12.proof Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,83 @@ +#2 := false +#5 := 0::int +#4 := (:var 0 int) +#42 := (<= #4 0::int) +#43 := (not #42) +#40 := (>= #4 0::int) +#38 := (not #40) +#46 := (or #38 #43) +#49 := (forall (vars (?x1 int)) #46) +#524 := (not #49) +#118 := (<= 0::int 0::int) +#205 := (not #118) +#119 := (>= 0::int 0::int) +#206 := (not #119) +#120 := (or #206 #205) +#183 := (or #524 #120) +#172 := (iff #183 #524) +#525 := (or #524 false) +#168 := (iff #525 #524) +#510 := [rewrite]: #168 +#184 := (iff #183 #525) +#528 := (iff #120 false) +#197 := (or false false) +#532 := (iff #197 false) +#533 := [rewrite]: #532 +#530 := (iff #120 #197) +#523 := (iff #205 false) +#1 := true +#209 := (not true) +#211 := (iff #209 false) +#208 := [rewrite]: #211 +#185 := (iff #205 #209) +#527 := (iff #118 true) +#529 := [rewrite]: #527 +#316 := [monotonicity #529]: #185 +#196 := [trans #316 #208]: #523 +#212 := (iff #206 false) +#210 := (iff #206 #209) +#207 := (iff #119 true) +#198 := [rewrite]: #207 +#138 := [monotonicity #198]: #210 +#191 := [trans #138 #208]: #212 +#531 := [monotonicity #191 #196]: #530 +#534 := [trans #531 #533]: #528 +#526 := [monotonicity #534]: #184 +#173 := [trans #526 #510]: #172 +#188 := [quant-inst]: #183 +#174 := [mp #188 #173]: #524 +#60 := (~ #49 #49) +#58 := (~ #46 #46) +#59 := [refl]: #58 +#61 := [nnf-pos #59]: #60 +#7 := (< 0::int #4) +#6 := (< #4 0::int) +#8 := (or #6 #7) +#9 := (forall (vars (?x1 int)) #8) +#10 := (ite #9 false true) +#11 := (not #10) +#52 := (iff #11 #49) +#50 := (iff #9 #49) +#47 := (iff #8 #46) +#44 := (iff #7 #43) +#45 := [rewrite]: #44 +#39 := (iff #6 #38) +#41 := [rewrite]: #39 +#48 := [monotonicity #41 #45]: #47 +#51 := [quant-intro #48]: #50 +#36 := (iff #11 #9) +#28 := (not #9) +#31 := (not #28) +#34 := (iff #31 #9) +#35 := [rewrite]: #34 +#32 := (iff #11 #31) +#29 := (iff #10 #28) +#30 := [rewrite]: #29 +#33 := [monotonicity #30]: #32 +#37 := [trans #33 #35]: #36 +#53 := [trans #37 #51]: #52 +#27 := [asserted]: #11 +#54 := [mp #27 #53]: #49 +#62 := [mp~ #54 #61]: #49 +[unit-resolution #62 #174]: false +unsat diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_arith_quant_13 --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_arith_quant_13 Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,4 @@ +(benchmark Isabelle +:assumption (not (< 0 (ite (forall (?x1 Int) (or (< ?x1 0) (< 0 ?x1))) (~ 1) 3))) +:formula true +) diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_arith_quant_13.proof --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_arith_quant_13.proof Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,180 @@ +#2 := false +#4 := 0::int +#5 := (:var 0 int) +#48 := (<= #5 0::int) +#49 := (not #48) +#45 := (>= #5 0::int) +#44 := (not #45) +#52 := (or #44 #49) +#55 := (forall (vars (?x1 int)) #52) +#86 := (not #55) +#604 := (<= 0::int 0::int) +#264 := (not #604) +#269 := (>= 0::int 0::int) +#605 := (not #269) +#265 := (or #605 #264) +#588 := (or #86 #265) +#584 := (iff #588 #86) +#311 := (or #86 false) +#314 := (iff #311 #86) +#208 := [rewrite]: #314 +#312 := (iff #588 #311) +#599 := (iff #265 false) +#598 := (or false false) +#241 := (iff #598 false) +#601 := [rewrite]: #241 +#600 := (iff #265 #598) +#597 := (iff #264 false) +#1 := true +#590 := (not true) +#255 := (iff #590 false) +#256 := [rewrite]: #255 +#596 := (iff #264 #590) +#594 := (iff #604 true) +#595 := [rewrite]: #594 +#591 := [monotonicity #595]: #596 +#235 := [trans #591 #256]: #597 +#592 := (iff #605 false) +#253 := (iff #605 #590) +#606 := (iff #269 true) +#249 := [rewrite]: #606 +#254 := [monotonicity #249]: #253 +#593 := [trans #254 #256]: #592 +#240 := [monotonicity #593 #235]: #600 +#602 := [trans #240 #601]: #599 +#313 := [monotonicity #602]: #312 +#585 := [trans #313 #208]: #584 +#589 := [quant-inst]: #588 +#307 := [mp #589 #585]: #86 +decl z3name!0 :: bool +#83 := z3name!0 +#12 := 3::int +#32 := -1::int +#92 := (ite z3name!0 -1::int 3::int) +#290 := (= #92 3::int) +#610 := (not #290) +#607 := (>= #92 3::int) +#609 := (not #607) +#95 := (<= #92 0::int) +#58 := (ite #55 -1::int 3::int) +#64 := (<= #58 0::int) +#96 := (~ #64 #95) +#93 := (= #58 #92) +#90 := (~ #55 z3name!0) +#87 := (or z3name!0 #86) +#84 := (not z3name!0) +#85 := (or #84 #55) +#88 := (and #85 #87) +#89 := [intro-def]: #88 +#91 := [apply-def #89]: #90 +#94 := [monotonicity #91]: #93 +#97 := [monotonicity #94]: #96 +#10 := 1::int +#11 := (- 1::int) +#7 := (< 0::int #5) +#6 := (< #5 0::int) +#8 := (or #6 #7) +#9 := (forall (vars (?x1 int)) #8) +#13 := (ite #9 #11 3::int) +#14 := (< 0::int #13) +#15 := (not #14) +#77 := (iff #15 #64) +#35 := (ite #9 -1::int 3::int) +#38 := (< 0::int #35) +#41 := (not #38) +#75 := (iff #41 #64) +#65 := (not #64) +#70 := (not #65) +#73 := (iff #70 #64) +#74 := [rewrite]: #73 +#71 := (iff #41 #70) +#68 := (iff #38 #65) +#61 := (< 0::int #58) +#66 := (iff #61 #65) +#67 := [rewrite]: #66 +#62 := (iff #38 #61) +#59 := (= #35 #58) +#56 := (iff #9 #55) +#53 := (iff #8 #52) +#50 := (iff #7 #49) +#51 := [rewrite]: #50 +#46 := (iff #6 #44) +#47 := [rewrite]: #46 +#54 := [monotonicity #47 #51]: #53 +#57 := [quant-intro #54]: #56 +#60 := [monotonicity #57]: #59 +#63 := [monotonicity #60]: #62 +#69 := [trans #63 #67]: #68 +#72 := [monotonicity #69]: #71 +#76 := [trans #72 #74]: #75 +#42 := (iff #15 #41) +#39 := (iff #14 #38) +#36 := (= #13 #35) +#33 := (= #11 -1::int) +#34 := [rewrite]: #33 +#37 := [monotonicity #34]: #36 +#40 := [monotonicity #37]: #39 +#43 := [monotonicity #40]: #42 +#78 := [trans #43 #76]: #77 +#31 := [asserted]: #15 +#79 := [mp #31 #78]: #64 +#126 := [mp~ #79 #97]: #95 +#266 := (not #95) +#396 := (or #609 #266) +#603 := [th-lemma]: #396 +#277 := [unit-resolution #603 #126]: #609 +#278 := [hypothesis]: #290 +#611 := (or #610 #607) +#612 := [th-lemma]: #611 +#613 := [unit-resolution #612 #278 #277]: false +#608 := [lemma #613]: #610 +#289 := (or z3name!0 #290) +#293 := [def-axiom]: #289 +#308 := [unit-resolution #293 #608]: z3name!0 +#129 := (or #55 #84) +decl ?x1!1 :: int +#108 := ?x1!1 +#111 := (>= ?x1!1 0::int) +#112 := (not #111) +#109 := (<= ?x1!1 0::int) +#110 := (not #109) +#132 := (or #110 #112) +#135 := (not #132) +#138 := (or z3name!0 #135) +#141 := (and #129 #138) +#113 := (or #112 #110) +#114 := (not #113) +#119 := (or z3name!0 #114) +#122 := (and #85 #119) +#142 := (iff #122 #141) +#139 := (iff #119 #138) +#136 := (iff #114 #135) +#133 := (iff #113 #132) +#134 := [rewrite]: #133 +#137 := [monotonicity #134]: #136 +#140 := [monotonicity #137]: #139 +#130 := (iff #85 #129) +#131 := [rewrite]: #130 +#143 := [monotonicity #131 #140]: #142 +#123 := (~ #88 #122) +#120 := (~ #87 #119) +#115 := (~ #86 #114) +#116 := [sk]: #115 +#106 := (~ z3name!0 z3name!0) +#107 := [refl]: #106 +#121 := [monotonicity #107 #116]: #120 +#104 := (~ #85 #85) +#102 := (~ #55 #55) +#100 := (~ #52 #52) +#101 := [refl]: #100 +#103 := [nnf-pos #101]: #102 +#98 := (~ #84 #84) +#99 := [refl]: #98 +#105 := [monotonicity #99 #103]: #104 +#124 := [monotonicity #105 #121]: #123 +#125 := [mp~ #89 #124]: #122 +#127 := [mp #125 #143]: #141 +#128 := [and-elim #127]: #129 +#582 := [unit-resolution #128 #308]: #55 +[unit-resolution #582 #307]: false +unsat diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_arith_quant_14 --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_arith_quant_14 Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,4 @@ +(benchmark Isabelle +:assumption (not (not (exists (?x1 Int) (?x2 Int) (?x3 Int) (= (+ (* 4 ?x1) (* (~ 6) ?x2)) 1)))) +:formula true +) diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_arith_quant_14.proof --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_arith_quant_14.proof Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,103 @@ +#2 := false +#104 := -1::int +decl ?x1!1 :: int +#86 := ?x1!1 +#106 := -4::int +#107 := (* -4::int ?x1!1) +decl ?x2!0 :: int +#85 := ?x2!0 +#7 := 6::int +#105 := (* 6::int ?x2!0) +#108 := (+ #105 #107) +#168 := (<= #108 -1::int) +#109 := (= #108 -1::int) +#12 := 1::int +#33 := -6::int +#87 := (* -6::int ?x2!0) +#4 := 4::int +#88 := (* 4::int ?x1!1) +#89 := (+ #88 #87) +#90 := (= #89 1::int) +#112 := (iff #90 #109) +#98 := (+ #87 #88) +#101 := (= #98 1::int) +#110 := (iff #101 #109) +#111 := [rewrite]: #110 +#102 := (iff #90 #101) +#99 := (= #89 #98) +#100 := [rewrite]: #99 +#103 := [monotonicity #100]: #102 +#113 := [trans #103 #111]: #112 +#53 := (:var 0 int) +#54 := (* -6::int #53) +#9 := (:var 1 int) +#55 := (* 4::int #9) +#56 := (+ #55 #54) +#76 := (= #56 1::int) +#74 := (exists (vars (?x1 int) (?x2 int)) #76) +#91 := (~ #74 #90) +#92 := [sk]: #91 +#8 := (- 6::int) +#10 := (* #8 #9) +#5 := (:var 2 int) +#6 := (* 4::int #5) +#11 := (+ #6 #10) +#13 := (= #11 1::int) +#14 := (exists (vars (?x1 int) (?x2 int) (?x3 int)) #13) +#15 := (not #14) +#16 := (not #15) +#79 := (iff #16 #74) +#57 := (= 1::int #56) +#58 := (exists (vars (?x1 int) (?x2 int)) #57) +#77 := (iff #58 #74) +#75 := (iff #57 #76) +#73 := [rewrite]: #75 +#78 := [quant-intro #73]: #77 +#71 := (iff #16 #58) +#63 := (not #58) +#66 := (not #63) +#69 := (iff #66 #58) +#70 := [rewrite]: #69 +#67 := (iff #16 #66) +#64 := (iff #15 #63) +#61 := (iff #14 #58) +#36 := (* -6::int #9) +#39 := (+ #6 #36) +#45 := (= 1::int #39) +#50 := (exists (vars (?x1 int) (?x2 int) (?x3 int)) #45) +#59 := (iff #50 #58) +#60 := [elim-unused]: #59 +#51 := (iff #14 #50) +#48 := (iff #13 #45) +#42 := (= #39 1::int) +#46 := (iff #42 #45) +#47 := [rewrite]: #46 +#43 := (iff #13 #42) +#40 := (= #11 #39) +#37 := (= #10 #36) +#34 := (= #8 -6::int) +#35 := [rewrite]: #34 +#38 := [monotonicity #35]: #37 +#41 := [monotonicity #38]: #40 +#44 := [monotonicity #41]: #43 +#49 := [trans #44 #47]: #48 +#52 := [quant-intro #49]: #51 +#62 := [trans #52 #60]: #61 +#65 := [monotonicity #62]: #64 +#68 := [monotonicity #65]: #67 +#72 := [trans #68 #70]: #71 +#80 := [trans #72 #78]: #79 +#32 := [asserted]: #16 +#81 := [mp #32 #80]: #74 +#95 := [mp~ #81 #92]: #90 +#96 := [mp #95 #113]: #109 +#170 := (not #109) +#171 := (or #170 #168) +#172 := [th-lemma]: #171 +#173 := [unit-resolution #172 #96]: #168 +#169 := (>= #108 -1::int) +#174 := (or #170 #169) +#175 := [th-lemma]: #174 +#176 := [unit-resolution #175 #96]: #169 +[th-lemma #176 #173]: false +unsat diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_arith_quant_15 --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_arith_quant_15 Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,4 @@ +(benchmark Isabelle +:assumption (not (exists (?x1 Int) (forall (?x2 Int) (?x3 Int) (implies (and (< 0 ?x2) (< 0 ?x3)) (< 0 (+ ?x2 ?x3)))))) +:formula true +) diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_arith_quant_15.proof --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_arith_quant_15.proof Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,107 @@ +#2 := false +#4 := 0::int +decl ?x2!1 :: int +#83 := ?x2!1 +decl ?x3!0 :: int +#82 := ?x3!0 +#108 := (+ ?x3!0 ?x2!1) +#111 := (<= #108 0::int) +#114 := (not #111) +#89 := (<= ?x2!1 0::int) +#90 := (not #89) +#87 := (<= ?x3!0 0::int) +#88 := (not #87) +#102 := (and #88 #90) +#105 := (not #102) +#117 := (or #105 #114) +#120 := (not #117) +#84 := (+ ?x2!1 ?x3!0) +#85 := (<= #84 0::int) +#86 := (not #85) +#91 := (and #90 #88) +#92 := (not #91) +#93 := (or #92 #86) +#94 := (not #93) +#121 := (iff #94 #120) +#118 := (iff #93 #117) +#115 := (iff #86 #114) +#112 := (iff #85 #111) +#109 := (= #84 #108) +#110 := [rewrite]: #109 +#113 := [monotonicity #110]: #112 +#116 := [monotonicity #113]: #115 +#106 := (iff #92 #105) +#103 := (iff #91 #102) +#104 := [rewrite]: #103 +#107 := [monotonicity #104]: #106 +#119 := [monotonicity #107 #116]: #118 +#122 := [monotonicity #119]: #121 +#7 := (:var 0 int) +#5 := (:var 1 int) +#10 := (+ #5 #7) +#63 := (<= #10 0::int) +#64 := (not #63) +#53 := (<= #7 0::int) +#54 := (not #53) +#49 := (<= #5 0::int) +#50 := (not #49) +#57 := (and #50 #54) +#60 := (not #57) +#67 := (or #60 #64) +#70 := (forall (vars (?x2 int) (?x3 int)) #67) +#73 := (not #70) +#95 := (~ #73 #94) +#96 := [sk]: #95 +#11 := (< 0::int #10) +#8 := (< 0::int #7) +#6 := (< 0::int #5) +#9 := (and #6 #8) +#12 := (implies #9 #11) +#13 := (forall (vars (?x2 int) (?x3 int)) #12) +#14 := (exists (vars (?x1 int)) #13) +#15 := (not #14) +#76 := (iff #15 #73) +#32 := (not #9) +#33 := (or #32 #11) +#36 := (forall (vars (?x2 int) (?x3 int)) #33) +#46 := (not #36) +#74 := (iff #46 #73) +#71 := (iff #36 #70) +#68 := (iff #33 #67) +#65 := (iff #11 #64) +#66 := [rewrite]: #65 +#61 := (iff #32 #60) +#58 := (iff #9 #57) +#55 := (iff #8 #54) +#56 := [rewrite]: #55 +#51 := (iff #6 #50) +#52 := [rewrite]: #51 +#59 := [monotonicity #52 #56]: #58 +#62 := [monotonicity #59]: #61 +#69 := [monotonicity #62 #66]: #68 +#72 := [quant-intro #69]: #71 +#75 := [monotonicity #72]: #74 +#47 := (iff #15 #46) +#44 := (iff #14 #36) +#39 := (exists (vars (?x1 int)) #36) +#42 := (iff #39 #36) +#43 := [elim-unused]: #42 +#40 := (iff #14 #39) +#37 := (iff #13 #36) +#34 := (iff #12 #33) +#35 := [rewrite]: #34 +#38 := [quant-intro #35]: #37 +#41 := [quant-intro #38]: #40 +#45 := [trans #41 #43]: #44 +#48 := [monotonicity #45]: #47 +#77 := [trans #48 #75]: #76 +#31 := [asserted]: #15 +#78 := [mp #31 #77]: #73 +#99 := [mp~ #78 #96]: #94 +#100 := [mp #99 #122]: #120 +#125 := [not-or-elim #100]: #111 +#101 := [not-or-elim #100]: #102 +#124 := [and-elim #101]: #90 +#123 := [and-elim #101]: #88 +[th-lemma #123 #124 #125]: false +unsat diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_arith_quant_16 --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_arith_quant_16 Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,4 @@ +(benchmark Isabelle +:assumption (not (exists (?x1 Int) (forall (?x2 Int) (?x3 Real) (implies (and (< 0 ?x2) (< 0.0 ?x3)) (< (~ 1) ?x2))))) +:formula true +) diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_arith_quant_16.proof --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_arith_quant_16.proof Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,117 @@ +#2 := false +#4 := 0::int +decl ?x2!1 :: int +#91 := ?x2!1 +#98 := (<= ?x2!1 0::int) +#99 := (not #98) +#7 := 0::real +decl ?x3!0 :: real +#93 := ?x3!0 +#96 := (<= ?x3!0 0::real) +#97 := (not #96) +#111 := (and #97 #99) +#114 := (not #111) +#33 := -1::int +#94 := (<= ?x2!1 -1::int) +#95 := (not #94) +#120 := (or #95 #114) +#125 := (not #120) +#100 := (and #99 #97) +#101 := (not #100) +#102 := (or #101 #95) +#103 := (not #102) +#126 := (iff #103 #125) +#123 := (iff #102 #120) +#117 := (or #114 #95) +#121 := (iff #117 #120) +#122 := [rewrite]: #121 +#118 := (iff #102 #117) +#115 := (iff #101 #114) +#112 := (iff #100 #111) +#113 := [rewrite]: #112 +#116 := [monotonicity #113]: #115 +#119 := [monotonicity #116]: #118 +#124 := [trans #119 #122]: #123 +#127 := [monotonicity #124]: #126 +#5 := (:var 1 int) +#75 := (<= #5 -1::int) +#76 := (not #75) +#8 := (:var 0 real) +#65 := (<= #8 0::real) +#66 := (not #65) +#61 := (<= #5 0::int) +#62 := (not #61) +#69 := (and #62 #66) +#72 := (not #69) +#79 := (or #72 #76) +#82 := (forall (vars (?x2 int) (?x3 real)) #79) +#85 := (not #82) +#104 := (~ #85 #103) +#105 := [sk]: #104 +#11 := 1::int +#12 := (- 1::int) +#13 := (< #12 #5) +#9 := (< 0::real #8) +#6 := (< 0::int #5) +#10 := (and #6 #9) +#14 := (implies #10 #13) +#15 := (forall (vars (?x2 int) (?x3 real)) #14) +#16 := (exists (vars (?x1 int)) #15) +#17 := (not #16) +#88 := (iff #17 #85) +#36 := (< -1::int #5) +#42 := (not #10) +#43 := (or #42 #36) +#48 := (forall (vars (?x2 int) (?x3 real)) #43) +#58 := (not #48) +#86 := (iff #58 #85) +#83 := (iff #48 #82) +#80 := (iff #43 #79) +#77 := (iff #36 #76) +#78 := [rewrite]: #77 +#73 := (iff #42 #72) +#70 := (iff #10 #69) +#67 := (iff #9 #66) +#68 := [rewrite]: #67 +#63 := (iff #6 #62) +#64 := [rewrite]: #63 +#71 := [monotonicity #64 #68]: #70 +#74 := [monotonicity #71]: #73 +#81 := [monotonicity #74 #78]: #80 +#84 := [quant-intro #81]: #83 +#87 := [monotonicity #84]: #86 +#59 := (iff #17 #58) +#56 := (iff #16 #48) +#51 := (exists (vars (?x1 int)) #48) +#54 := (iff #51 #48) +#55 := [elim-unused]: #54 +#52 := (iff #16 #51) +#49 := (iff #15 #48) +#46 := (iff #14 #43) +#39 := (implies #10 #36) +#44 := (iff #39 #43) +#45 := [rewrite]: #44 +#40 := (iff #14 #39) +#37 := (iff #13 #36) +#34 := (= #12 -1::int) +#35 := [rewrite]: #34 +#38 := [monotonicity #35]: #37 +#41 := [monotonicity #38]: #40 +#47 := [trans #41 #45]: #46 +#50 := [quant-intro #47]: #49 +#53 := [quant-intro #50]: #52 +#57 := [trans #53 #55]: #56 +#60 := [monotonicity #57]: #59 +#89 := [trans #60 #87]: #88 +#32 := [asserted]: #17 +#90 := [mp #32 #89]: #85 +#108 := [mp~ #90 #105]: #103 +#109 := [mp #108 #127]: #125 +#128 := [not-or-elim #109]: #111 +#130 := [and-elim #128]: #99 +#110 := [not-or-elim #109]: #94 +#186 := (or #95 #98) +#187 := [th-lemma]: #186 +#188 := [unit-resolution #187 #110]: #98 +[unit-resolution #188 #130]: false +unsat diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_arith_quant_17 --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_arith_quant_17 Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,4 @@ +(benchmark Isabelle +:assumption (not (exists (?x1 Int) (implies (forall (?x2 Int) (implies (<= ?x1 ?x2) (< 0 ?x2))) (< 0 ?x1)))) +:formula true +) diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_arith_quant_17.proof --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_arith_quant_17.proof Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,148 @@ +#2 := false +#144 := (not false) +#7 := 0::int +#5 := (:var 0 int) +#52 := (<= #5 0::int) +#53 := (not #52) +#147 := (or #53 #144) +#150 := (not #147) +#153 := (forall (vars (?x1 int)) #150) +#180 := (iff #153 false) +#175 := (forall (vars (?x1 int)) false) +#178 := (iff #175 false) +#179 := [elim-unused]: #178 +#176 := (iff #153 #175) +#173 := (iff #150 false) +#1 := true +#168 := (not true) +#171 := (iff #168 false) +#172 := [rewrite]: #171 +#169 := (iff #150 #168) +#166 := (iff #147 true) +#161 := (or #53 true) +#164 := (iff #161 true) +#165 := [rewrite]: #164 +#162 := (iff #147 #161) +#159 := (iff #144 true) +#160 := [rewrite]: #159 +#163 := [monotonicity #160]: #162 +#167 := [trans #163 #165]: #166 +#170 := [monotonicity #167]: #169 +#174 := [trans #170 #172]: #173 +#177 := [quant-intro #174]: #176 +#181 := [trans #177 #179]: #180 +#56 := -1::int +#57 := (* -1::int #5) +#4 := (:var 1 int) +#58 := (+ #4 #57) +#59 := (<= #58 0::int) +#62 := (not #59) +#68 := (or #53 #62) +#73 := (forall (vars (?x2 int)) #68) +#76 := (not #73) +#79 := (or #53 #76) +#105 := (not #79) +#123 := (forall (vars (?x1 int)) #105) +#156 := (iff #123 #153) +#127 := (forall (vars (?x2 int)) #53) +#130 := (not #127) +#133 := (or #53 #130) +#136 := (not #133) +#139 := (forall (vars (?x1 int)) #136) +#154 := (iff #139 #153) +#155 := [rewrite]: #154 +#140 := (iff #123 #139) +#141 := [rewrite]: #140 +#157 := [trans #141 #155]: #156 +#116 := (and #52 #73) +#119 := (forall (vars (?x1 int)) #116) +#124 := (iff #119 #123) +#113 := (iff #116 #105) +#122 := [rewrite]: #113 +#125 := [quant-intro #122]: #124 +#94 := (not #53) +#104 := (and #94 #73) +#108 := (forall (vars (?x1 int)) #104) +#120 := (iff #108 #119) +#117 := (iff #104 #116) +#114 := (iff #94 #52) +#115 := [rewrite]: #114 +#118 := [monotonicity #115]: #117 +#121 := [quant-intro #118]: #120 +#82 := (exists (vars (?x1 int)) #79) +#85 := (not #82) +#109 := (~ #85 #108) +#106 := (~ #105 #104) +#101 := (not #76) +#102 := (~ #101 #73) +#99 := (~ #73 #73) +#97 := (~ #68 #68) +#98 := [refl]: #97 +#100 := [nnf-pos #98]: #99 +#103 := [nnf-neg #100]: #102 +#95 := (~ #94 #94) +#96 := [refl]: #95 +#107 := [nnf-neg #96 #103]: #106 +#110 := [nnf-neg #107]: #109 +#8 := (< 0::int #5) +#6 := (<= #4 #5) +#9 := (implies #6 #8) +#10 := (forall (vars (?x2 int)) #9) +#11 := (implies #10 #8) +#12 := (exists (vars (?x1 int)) #11) +#13 := (not #12) +#88 := (iff #13 #85) +#30 := (not #6) +#31 := (or #30 #8) +#34 := (forall (vars (?x2 int)) #31) +#40 := (not #34) +#41 := (or #8 #40) +#46 := (exists (vars (?x1 int)) #41) +#49 := (not #46) +#86 := (iff #49 #85) +#83 := (iff #46 #82) +#80 := (iff #41 #79) +#77 := (iff #40 #76) +#74 := (iff #34 #73) +#71 := (iff #31 #68) +#65 := (or #62 #53) +#69 := (iff #65 #68) +#70 := [rewrite]: #69 +#66 := (iff #31 #65) +#54 := (iff #8 #53) +#55 := [rewrite]: #54 +#63 := (iff #30 #62) +#60 := (iff #6 #59) +#61 := [rewrite]: #60 +#64 := [monotonicity #61]: #63 +#67 := [monotonicity #64 #55]: #66 +#72 := [trans #67 #70]: #71 +#75 := [quant-intro #72]: #74 +#78 := [monotonicity #75]: #77 +#81 := [monotonicity #55 #78]: #80 +#84 := [quant-intro #81]: #83 +#87 := [monotonicity #84]: #86 +#50 := (iff #13 #49) +#47 := (iff #12 #46) +#44 := (iff #11 #41) +#37 := (implies #34 #8) +#42 := (iff #37 #41) +#43 := [rewrite]: #42 +#38 := (iff #11 #37) +#35 := (iff #10 #34) +#32 := (iff #9 #31) +#33 := [rewrite]: #32 +#36 := [quant-intro #33]: #35 +#39 := [monotonicity #36]: #38 +#45 := [trans #39 #43]: #44 +#48 := [quant-intro #45]: #47 +#51 := [monotonicity #48]: #50 +#89 := [trans #51 #87]: #88 +#29 := [asserted]: #13 +#90 := [mp #29 #89]: #85 +#111 := [mp~ #90 #110]: #108 +#112 := [mp #111 #121]: #119 +#126 := [mp #112 #125]: #123 +#158 := [mp #126 #157]: #153 +[mp #158 #181]: false +unsat diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_arith_quant_18 --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_arith_quant_18 Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,7 @@ +(benchmark Isabelle +:extrafuns ( + (uf_1 Int) + ) +:assumption (not (forall (?x1 Int) (implies (< ?x1 uf_1) (< (* 2 ?x1) (* 2 uf_1))) :pat{ ?x1 })) +:formula true +) diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_arith_quant_18.proof --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_arith_quant_18.proof Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,112 @@ +#2 := false +#43 := 0::int +decl ?x1!0 :: int +#78 := ?x1!0 +#56 := -2::int +#113 := (* -2::int ?x1!0) +decl uf_1 :: int +#6 := uf_1 +#8 := 2::int +#10 := (* 2::int uf_1) +#114 := (+ #10 #113) +#115 := (<= #114 0::int) +#120 := (not #115) +#41 := -1::int +#100 := (* -1::int ?x1!0) +#101 := (+ uf_1 #100) +#102 := (<= #101 0::int) +#123 := (or #102 #120) +#126 := (not #123) +#59 := (* -2::int uf_1) +#79 := (* 2::int ?x1!0) +#80 := (+ #79 #59) +#81 := (>= #80 0::int) +#82 := (not #81) +#45 := (* -1::int uf_1) +#83 := (+ ?x1!0 #45) +#84 := (>= #83 0::int) +#85 := (or #84 #82) +#86 := (not #85) +#127 := (iff #86 #126) +#124 := (iff #85 #123) +#121 := (iff #82 #120) +#118 := (iff #81 #115) +#107 := (+ #59 #79) +#110 := (>= #107 0::int) +#116 := (iff #110 #115) +#117 := [rewrite]: #116 +#111 := (iff #81 #110) +#108 := (= #80 #107) +#109 := [rewrite]: #108 +#112 := [monotonicity #109]: #111 +#119 := [trans #112 #117]: #118 +#122 := [monotonicity #119]: #121 +#105 := (iff #84 #102) +#94 := (+ #45 ?x1!0) +#97 := (>= #94 0::int) +#103 := (iff #97 #102) +#104 := [rewrite]: #103 +#98 := (iff #84 #97) +#95 := (= #83 #94) +#96 := [rewrite]: #95 +#99 := [monotonicity #96]: #98 +#106 := [trans #99 #104]: #105 +#125 := [monotonicity #106 #122]: #124 +#128 := [monotonicity #125]: #127 +#4 := (:var 0 int) +#5 := (pattern #4) +#9 := (* 2::int #4) +#60 := (+ #9 #59) +#58 := (>= #60 0::int) +#57 := (not #58) +#46 := (+ #4 #45) +#44 := (>= #46 0::int) +#63 := (or #44 #57) +#66 := (forall (vars (?x1 int)) (:pat #5) #63) +#69 := (not #66) +#87 := (~ #69 #86) +#88 := [sk]: #87 +#11 := (< #9 #10) +#7 := (< #4 uf_1) +#12 := (implies #7 #11) +#13 := (forall (vars (?x1 int)) (:pat #5) #12) +#14 := (not #13) +#72 := (iff #14 #69) +#31 := (not #7) +#32 := (or #31 #11) +#35 := (forall (vars (?x1 int)) (:pat #5) #32) +#38 := (not #35) +#70 := (iff #38 #69) +#67 := (iff #35 #66) +#64 := (iff #32 #63) +#61 := (iff #11 #57) +#62 := [rewrite]: #61 +#54 := (iff #31 #44) +#42 := (not #44) +#49 := (not #42) +#52 := (iff #49 #44) +#53 := [rewrite]: #52 +#50 := (iff #31 #49) +#47 := (iff #7 #42) +#48 := [rewrite]: #47 +#51 := [monotonicity #48]: #50 +#55 := [trans #51 #53]: #54 +#65 := [monotonicity #55 #62]: #64 +#68 := [quant-intro #65]: #67 +#71 := [monotonicity #68]: #70 +#39 := (iff #14 #38) +#36 := (iff #13 #35) +#33 := (iff #12 #32) +#34 := [rewrite]: #33 +#37 := [quant-intro #34]: #36 +#40 := [monotonicity #37]: #39 +#73 := [trans #40 #71]: #72 +#30 := [asserted]: #14 +#74 := [mp #30 #73]: #69 +#91 := [mp~ #74 #88]: #86 +#92 := [mp #91 #128]: #126 +#130 := [not-or-elim #92]: #115 +#93 := (not #102) +#129 := [not-or-elim #92]: #93 +[th-lemma #129 #130]: false +unsat diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_bv_01 --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_bv_01 Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,12 @@ +(benchmark Isabelle +:extrafuns ( + (uf_1 BitVec[2] Int) + ) +:assumption (= (uf_1 bv0[2]) 0) +:assumption (= (uf_1 bv1[2]) 1) +:assumption (= (uf_1 bv2[2]) 2) +:assumption (= (uf_1 bv3[2]) 3) +:assumption (forall (?x1 BitVec[2]) (< 0 (uf_1 ?x1))) +:assumption (not (forall (?x2 Int) (implies (< ?x2 0) (forall (?x3 BitVec[2]) (< ?x2 (uf_1 ?x3)))))) +:formula true +) diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_bv_01.proof --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_bv_01.proof Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,43 @@ +#2 := false +#6 := 0::int +decl uf_1 :: (-> bv[2] int) +#4 := bv[0:2] +#5 := (uf_1 bv[0:2]) +#225 := (<= #5 0::int) +#311 := (not #225) +#20 := (:var 0 bv[2]) +#21 := (uf_1 #20) +#640 := (pattern #21) +#54 := (<= #21 0::int) +#55 := (not #54) +#641 := (forall (vars (?x1 bv[2])) (:pat #640) #55) +#58 := (forall (vars (?x1 bv[2])) #55) +#644 := (iff #58 #641) +#642 := (iff #55 #55) +#643 := [refl]: #642 +#645 := [quant-intro #643]: #644 +#113 := (~ #58 #58) +#115 := (~ #55 #55) +#116 := [refl]: #115 +#114 := [nnf-pos #116]: #113 +#22 := (< 0::int #21) +#23 := (forall (vars (?x1 bv[2])) #22) +#59 := (iff #23 #58) +#56 := (iff #22 #55) +#57 := [rewrite]: #56 +#60 := [quant-intro #57]: #59 +#51 := [asserted]: #23 +#61 := [mp #51 #60]: #58 +#111 := [mp~ #61 #114]: #58 +#646 := [mp #111 #645]: #641 +#227 := (not #641) +#313 := (or #227 #311) +#304 := [quant-inst]: #313 +#635 := [unit-resolution #304 #646]: #311 +#7 := (= #5 0::int) +#47 := [asserted]: #7 +#638 := (not #7) +#633 := (or #638 #225) +#639 := [th-lemma]: #633 +[unit-resolution #639 #47 #635]: false +unsat diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_bv_02 --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_bv_02 Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,12 @@ +(benchmark Isabelle +:extrasorts ( T2 T1) +:extrafuns ( + (uf_2 T1) + (uf_1 BitVec[4] BitVec[4] T1) + (uf_3 T1 T2) + (uf_4 BitVec[4]) + ) +:assumption (forall (?x1 BitVec[4]) (?x2 BitVec[4]) (iff (= (uf_1 ?x1 ?x2) uf_2) (bvule ?x1 ?x2))) +:assumption (not (= (uf_3 (uf_1 bv0[4] uf_4)) (uf_3 uf_2))) +:formula true +) diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_bv_02.proof --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_bv_02.proof Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,1 @@ +unsat diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_bv_arith_01 --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_bv_arith_01 Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,4 @@ +(benchmark Isabelle +:assumption (not (= bv27[4] (bvneg bv5[4]))) +:formula true +) diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_bv_arith_01.proof --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_bv_arith_01.proof Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,1 @@ +unsat diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_bv_arith_02 --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_bv_arith_02 Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,4 @@ +(benchmark Isabelle +:assumption (not (= bv27[4] bv11[4])) +:formula true +) diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_bv_arith_02.proof --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_bv_arith_02.proof Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,1 @@ +unsat diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_bv_arith_03 --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_bv_arith_03 Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,4 @@ +(benchmark Isabelle +:assumption (not (bvult bv23[8] bv27[8])) +:formula true +) diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_bv_arith_03.proof --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_bv_arith_03.proof Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,1 @@ +unsat diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_bv_arith_04 --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_bv_arith_04 Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,4 @@ +(benchmark Isabelle +:assumption (not (= (bvadd bv27[5] bv11[5]) bv6[5])) +:formula true +) diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_bv_arith_04.proof --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_bv_arith_04.proof Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,1 @@ +unsat diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_bv_arith_05 --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_bv_arith_05 Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,4 @@ +(benchmark Isabelle +:assumption (not (= (bvmul bv7[8] bv3[8]) bv21[8])) +:formula true +) diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_bv_arith_05.proof --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_bv_arith_05.proof Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,1 @@ +unsat diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_bv_arith_06 --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_bv_arith_06 Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,4 @@ +(benchmark Isabelle +:assumption (not (= (bvsub bv11[8] bv27[8]) (bvneg bv16[8]))) +:formula true +) diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_bv_arith_06.proof --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_bv_arith_06.proof Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,1 @@ +unsat diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_bv_arith_07 --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_bv_arith_07 Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,4 @@ +(benchmark Isabelle +:assumption (not (= (bvneg (bvneg bv11[5])) bv11[5])) +:formula true +) diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_bv_arith_07.proof --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_bv_arith_07.proof Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,1 @@ +unsat diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_bv_arith_08 --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_bv_arith_08 Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,4 @@ +(benchmark Isabelle +:assumption (not (= (bvadd (bvneg bv40[7]) bv1[7]) (bvneg bv39[7]))) +:formula true +) diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_bv_arith_08.proof --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_bv_arith_08.proof Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,1 @@ +unsat diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_bv_arith_09 --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_bv_arith_09 Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,9 @@ +(benchmark Isabelle +:extrafuns ( + (uf_1 BitVec[32]) + (uf_2 BitVec[32]) + (uf_3 BitVec[32]) + ) +:assumption (not (= (bvsub (bvadd (bvadd uf_1 (bvmul bv2[32] uf_2)) uf_3) uf_2) (bvadd (bvadd uf_2 uf_3) uf_1))) +:formula true +) diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_bv_arith_09.proof --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_bv_arith_09.proof Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,1 @@ +unsat diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_bv_arith_10 --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_bv_arith_10 Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,8 @@ +(benchmark Isabelle +:extrafuns ( + (uf_1 BitVec[4]) + ) +:assumption (= uf_1 bv5[4]) +:assumption (not (= (bvmul bv4[4] uf_1) bv4[4])) +:formula true +) diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_bv_arith_10.proof --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_bv_arith_10.proof Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,1 @@ +unsat diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_bv_bit_01 --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_bv_bit_01 Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,4 @@ +(benchmark Isabelle +:assumption (not (= (bvand bv6[32] bv5[32]) bv4[32])) +:formula true +) diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_bv_bit_01.proof --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_bv_bit_01.proof Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,1 @@ +unsat diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_bv_bit_02 --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_bv_bit_02 Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,4 @@ +(benchmark Isabelle +:assumption (not (= (bvor bv6[8] bv3[8]) bv7[8])) +:formula true +) diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_bv_bit_02.proof --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_bv_bit_02.proof Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,1 @@ +unsat diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_bv_bit_03 --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_bv_bit_03 Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,4 @@ +(benchmark Isabelle +:assumption (not (= (bvxor bv240[8] bv255[8]) bv15[8])) +:formula true +) diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_bv_bit_03.proof --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_bv_bit_03.proof Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,1 @@ +unsat diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_bv_bit_04 --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_bv_bit_04 Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,4 @@ +(benchmark Isabelle +:assumption (not (= (bvnot bv240[16]) bv65295[16])) +:formula true +) diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_bv_bit_04.proof --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_bv_bit_04.proof Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,1 @@ +unsat diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_bv_bit_05 --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_bv_bit_05 Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,4 @@ +(benchmark Isabelle +:assumption (not (= (concat bv27[4] bv27[8]) bv2843[12])) +:formula true +) diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_bv_bit_05.proof --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_bv_bit_05.proof Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,1 @@ +unsat diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_bv_bit_06 --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_bv_bit_06 Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,4 @@ +(benchmark Isabelle +:assumption (not (= (concat bv3[4] bv15[6]) bv207[10])) +:formula true +) diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_bv_bit_06.proof --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_bv_bit_06.proof Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,1 @@ +unsat diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_bv_bit_07 --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_bv_bit_07 Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,12 @@ +(benchmark Isabelle +:extrasorts ( T1) +:extrafuns ( + (uf_1 Int T1) + (uf_2 T1 Int) + ) +:assumption (forall (?x1 T1) (= (uf_1 (uf_2 ?x1)) ?x1)) +:assumption (forall (?x2 Int) (implies (<= 0 ?x2) (= (uf_2 (uf_1 ?x2)) ?x2))) +:assumption (forall (?x3 Int) (implies (< ?x3 0) (= (uf_2 (uf_1 ?x3)) 0))) +:assumption (not (= (extract[2:1] bv22[4]) bv3[2])) +:formula true +) diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_bv_bit_07.proof --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_bv_bit_07.proof Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,1 @@ +unsat diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_bv_bit_08 --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_bv_bit_08 Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,4 @@ +(benchmark Isabelle +:assumption (not (= (zero_extend[6] bv10[4]) bv10[10])) +:formula true +) diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_bv_bit_08.proof --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_bv_bit_08.proof Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,1 @@ +unsat diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_bv_bit_09 --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_bv_bit_09 Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,4 @@ +(benchmark Isabelle +:assumption (not (= (sign_extend[2] bv10[4]) bv58[6])) +:formula true +) diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_bv_bit_09.proof --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_bv_bit_09.proof Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,1 @@ +unsat diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_bv_bit_10 --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_bv_bit_10 Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,4 @@ +(benchmark Isabelle +:assumption (not (= (bvlshr bv19[8] bv2[8]) bv4[8])) +:formula true +) diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_bv_bit_10.proof --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_bv_bit_10.proof Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,1 @@ +unsat diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_bv_bit_11 --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_bv_bit_11 Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,4 @@ +(benchmark Isabelle +:assumption (not (= (bvashr bv19[8] bv2[8]) bv4[8])) +:formula true +) diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_bv_bit_11.proof --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_bv_bit_11.proof Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,1 @@ +unsat diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_bv_bit_12 --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_bv_bit_12 Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,12 @@ +(benchmark Isabelle +:extrasorts ( T1) +:extrafuns ( + (uf_1 Int T1) + (uf_2 T1 Int) + ) +:assumption (forall (?x1 T1) (= (uf_1 (uf_2 ?x1)) ?x1)) +:assumption (forall (?x2 Int) (implies (<= 0 ?x2) (= (uf_2 (uf_1 ?x2)) ?x2))) +:assumption (forall (?x3 Int) (implies (< ?x3 0) (= (uf_2 (uf_1 ?x3)) 0))) +:assumption (not (= (rotate_right[2] bv6[4]) bv9[4])) +:formula true +) diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_bv_bit_12.proof --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_bv_bit_12.proof Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,1 @@ +unsat diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_bv_bit_13 --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_bv_bit_13 Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,12 @@ +(benchmark Isabelle +:extrasorts ( T1) +:extrafuns ( + (uf_1 Int T1) + (uf_2 T1 Int) + ) +:assumption (forall (?x1 T1) (= (uf_1 (uf_2 ?x1)) ?x1)) +:assumption (forall (?x2 Int) (implies (<= 0 ?x2) (= (uf_2 (uf_1 ?x2)) ?x2))) +:assumption (forall (?x3 Int) (implies (< ?x3 0) (= (uf_2 (uf_1 ?x3)) 0))) +:assumption (not (= (rotate_left[1] bv14[4]) bv13[4])) +:formula true +) diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_bv_bit_13.proof --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_bv_bit_13.proof Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,1 @@ +unsat diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_bv_bit_14 --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_bv_bit_14 Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,7 @@ +(benchmark Isabelle +:extrafuns ( + (uf_1 BitVec[16]) + ) +:assumption (not (= (bvor (bvand uf_1 bv65280[16]) (bvand uf_1 bv255[16])) uf_1)) +:formula true +) diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_bv_bit_14.proof --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_bv_bit_14.proof Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,1 @@ +unsat diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_bv_bit_15 --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_bv_bit_15 Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,8 @@ +(benchmark Isabelle +:extrafuns ( + (uf_1 BitVec[16]) + ) +:assumption (bvult uf_1 bv256[16]) +:assumption (not (= (bvand uf_1 bv255[16]) uf_1)) +:formula true +) diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_bv_bit_15.proof --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_bv_bit_15.proof Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,1 @@ +unsat diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_fol_01 --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_fol_01 Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,8 @@ +(benchmark Isabelle +:extrasorts ( T1) +:extrapreds ( + (up_1 Int) + ) +:assumption (not (forall (?x1 Int) (implies (up_1 ?x1) (forall (?x2 Int) (or (up_1 ?x1) (up_1 ?x2)))))) +:formula true +) diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_fol_01.proof --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_fol_01.proof Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,78 @@ +#2 := false +decl up_1 :: (-> int bool) +decl ?x1!0 :: int +#54 := ?x1!0 +#55 := (up_1 ?x1!0) +#58 := (not #55) +decl ?x2!1 :: int +#66 := ?x2!1 +#67 := (up_1 ?x2!1) +#85 := (or #55 #67) +#88 := (not #85) +#91 := (and #55 #88) +#68 := (or #67 #55) +#69 := (not #68) +#63 := (not #58) +#75 := (and #63 #69) +#92 := (iff #75 #91) +#89 := (iff #69 #88) +#86 := (iff #68 #85) +#87 := [rewrite]: #86 +#90 := [monotonicity #87]: #89 +#83 := (iff #63 #55) +#84 := [rewrite]: #83 +#93 := [monotonicity #84 #90]: #92 +#6 := (:var 1 int) +#7 := (up_1 #6) +#4 := (:var 0 int) +#5 := (up_1 #4) +#29 := (or #5 #7) +#32 := (forall (vars (?x2 int)) #29) +#38 := (not #5) +#39 := (or #38 #32) +#44 := (forall (vars (?x1 int)) #39) +#47 := (not #44) +#78 := (~ #47 #75) +#56 := (or #5 #55) +#57 := (forall (vars (?x2 int)) #56) +#59 := (or #58 #57) +#60 := (not #59) +#76 := (~ #60 #75) +#70 := (not #57) +#71 := (~ #70 #69) +#72 := [sk]: #71 +#64 := (~ #63 #63) +#65 := [refl]: #64 +#77 := [nnf-neg #65 #72]: #76 +#61 := (~ #47 #60) +#62 := [sk]: #61 +#79 := [trans #62 #77]: #78 +#8 := (or #7 #5) +#9 := (forall (vars (?x2 int)) #8) +#10 := (implies #5 #9) +#11 := (forall (vars (?x1 int)) #10) +#12 := (not #11) +#48 := (iff #12 #47) +#45 := (iff #11 #44) +#42 := (iff #10 #39) +#35 := (implies #5 #32) +#40 := (iff #35 #39) +#41 := [rewrite]: #40 +#36 := (iff #10 #35) +#33 := (iff #9 #32) +#30 := (iff #8 #29) +#31 := [rewrite]: #30 +#34 := [quant-intro #31]: #33 +#37 := [monotonicity #34]: #36 +#43 := [trans #37 #41]: #42 +#46 := [quant-intro #43]: #45 +#49 := [monotonicity #46]: #48 +#28 := [asserted]: #12 +#52 := [mp #28 #49]: #47 +#80 := [mp~ #52 #79]: #75 +#81 := [mp #80 #93]: #91 +#94 := [and-elim #81]: #88 +#95 := [not-or-elim #94]: #58 +#82 := [and-elim #81]: #55 +[unit-resolution #82 #95]: false +unsat diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_fol_02 --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_fol_02 Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,14 @@ +(benchmark Isabelle +:extrasorts ( T2 T1) +:extrafuns ( + (uf_2 T1) + (uf_4 T2) + (uf_3 T1) + ) +:extrapreds ( + (up_1 T1 T2) + ) +:assumption (forall (?x1 T1) (?x2 T2) (iff (up_1 ?x1 ?x2) (= ?x1 uf_2))) +:assumption (not (iff (exists (?x3 T2) (up_1 uf_3 ?x3)) (up_1 uf_3 uf_4))) +:formula true +) diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_fol_02.proof --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_fol_02.proof Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,135 @@ +#2 := false +decl up_1 :: (-> T1 T2 bool) +#5 := (:var 0 T2) +decl uf_3 :: T1 +#11 := uf_3 +#12 := (up_1 uf_3 #5) +#560 := (pattern #12) +#57 := (not #12) +#561 := (forall (vars (?x3 T2)) (:pat #560) #57) +decl uf_4 :: T2 +#14 := uf_4 +#15 := (up_1 uf_3 uf_4) +decl uf_2 :: T1 +#7 := uf_2 +#136 := (= uf_2 uf_3) +#543 := (iff #15 #136) +#4 := (:var 1 T1) +#6 := (up_1 #4 #5) +#553 := (pattern #6) +#8 := (= #4 uf_2) +#9 := (iff #6 #8) +#554 := (forall (vars (?x1 T1) (?x2 T2)) (:pat #553) #9) +#10 := (forall (vars (?x1 T1) (?x2 T2)) #9) +#557 := (iff #10 #554) +#555 := (iff #9 #9) +#556 := [refl]: #555 +#558 := [quant-intro #556]: #557 +#47 := (~ #10 #10) +#45 := (~ #9 #9) +#46 := [refl]: #45 +#48 := [nnf-pos #46]: #47 +#33 := [asserted]: #10 +#49 := [mp~ #33 #48]: #10 +#559 := [mp #49 #558]: #554 +#227 := (not #554) +#185 := (or #227 #543) +#135 := (= uf_3 uf_2) +#205 := (iff #15 #135) +#528 := (or #227 #205) +#190 := (iff #528 #185) +#192 := (iff #185 #185) +#530 := [rewrite]: #192 +#201 := (iff #205 #543) +#223 := (iff #135 #136) +#137 := [rewrite]: #223 +#544 := [monotonicity #137]: #201 +#191 := [monotonicity #544]: #190 +#531 := [trans #191 #530]: #190 +#189 := [quant-inst]: #528 +#532 := [mp #189 #531]: #185 +#539 := [unit-resolution #532 #559]: #543 +decl ?x3!0 :: T2 +#50 := ?x3!0 +#51 := (up_1 uf_3 ?x3!0) +#224 := (iff #51 #136) +#155 := (or #227 #224) +#222 := (iff #51 #135) +#228 := (or #227 #222) +#229 := (iff #228 #155) +#545 := (iff #155 #155) +#547 := [rewrite]: #545 +#215 := (iff #222 #224) +#226 := [monotonicity #137]: #215 +#208 := [monotonicity #226]: #229 +#202 := [trans #208 #547]: #229 +#225 := [quant-inst]: #228 +#334 := [mp #225 #202]: #155 +#537 := [unit-resolution #334 #559]: #224 +#541 := (not #224) +#527 := (or #541 #136) +#63 := (not #15) +#540 := [hypothesis]: #63 +#68 := (or #15 #51) +#60 := (forall (vars (?x3 T2)) #57) +#69 := (or #63 #60) +#76 := (and #68 #69) +#70 := (and #69 #68) +#77 := (iff #70 #76) +#78 := [rewrite]: #77 +#13 := (exists (vars (?x3 T2)) #12) +#35 := (not #13) +#36 := (iff #15 #35) +#71 := (~ #36 #70) +#61 := (~ #35 #60) +#58 := (~ #57 #57) +#59 := [refl]: #58 +#62 := [nnf-neg #59]: #61 +#54 := (not #35) +#55 := (~ #54 #51) +#42 := (~ #13 #51) +#39 := [sk]: #42 +#56 := [nnf-neg #39]: #55 +#66 := (~ #15 #15) +#67 := [refl]: #66 +#64 := (~ #63 #63) +#65 := [refl]: #64 +#72 := [nnf-pos #65 #67 #56 #62]: #71 +#16 := (iff #13 #15) +#17 := (not #16) +#37 := (iff #17 #36) +#38 := [rewrite]: #37 +#34 := [asserted]: #17 +#41 := [mp #34 #38]: #36 +#73 := [mp~ #41 #72]: #70 +#74 := [mp #73 #78]: #76 +#75 := [and-elim #74]: #68 +#526 := [unit-resolution #75 #540]: #51 +#549 := (not #51) +#550 := (or #541 #549 #136) +#551 := [def-axiom]: #550 +#233 := [unit-resolution #551 #526]: #527 +#249 := [unit-resolution #233 #537]: #136 +#213 := (not #136) +#533 := (not #543) +#250 := (or #533 #213) +#534 := (or #533 #15 #213) +#529 := [def-axiom]: #534 +#251 := [unit-resolution #529 #540]: #250 +#237 := [unit-resolution #251 #249 #539]: false +#252 := [lemma #237]: #15 +#566 := (or #63 #561) +#567 := (iff #69 #566) +#564 := (iff #60 #561) +#562 := (iff #57 #57) +#563 := [refl]: #562 +#565 := [quant-intro #563]: #564 +#568 := [monotonicity #565]: #567 +#79 := [and-elim #74]: #69 +#569 := [mp #79 #568]: #566 +#535 := [unit-resolution #569 #252]: #561 +#536 := (not #561) +#538 := (or #536 #63) +#176 := [quant-inst]: #538 +[unit-resolution #176 #252 #535]: false +unsat diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_fol_03 --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_fol_03 Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,15 @@ +(benchmark Isabelle +:extrasorts ( T2 T1) +:extrafuns ( + (uf_2 T1) + (uf_3 T2) + (uf_4 T1) + ) +:extrapreds ( + (up_1 T1 T2) + ) +:assumption (forall (?x1 T1) (?x2 T2) (iff (up_1 ?x1 ?x2) (= ?x1 uf_2))) +:assumption (iff (forall (?x3 T1) (exists (?x4 T2) (up_1 ?x3 ?x4))) (forall (?x5 T1) (up_1 ?x5 uf_3))) +:assumption (not (iff (exists (?x6 T2) (up_1 uf_4 ?x6)) (up_1 uf_4 uf_3))) +:formula true +) diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_fol_03.proof --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_fol_03.proof Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,135 @@ +#2 := false +decl up_1 :: (-> T1 T2 bool) +#5 := (:var 0 T2) +decl uf_4 :: T1 +#18 := uf_4 +#19 := (up_1 uf_4 #5) +#635 := (pattern #19) +#116 := (not #19) +#636 := (forall (vars (?x6 T2)) (:pat #635) #116) +decl uf_3 :: T2 +#14 := uf_3 +#21 := (up_1 uf_4 uf_3) +decl uf_2 :: T1 +#7 := uf_2 +#195 := (= uf_2 uf_4) +#602 := (iff #21 #195) +#4 := (:var 1 T1) +#6 := (up_1 #4 #5) +#612 := (pattern #6) +#8 := (= #4 uf_2) +#9 := (iff #6 #8) +#613 := (forall (vars (?x1 T1) (?x2 T2)) (:pat #612) #9) +#10 := (forall (vars (?x1 T1) (?x2 T2)) #9) +#616 := (iff #10 #613) +#614 := (iff #9 #9) +#615 := [refl]: #614 +#617 := [quant-intro #615]: #616 +#56 := (~ #10 #10) +#54 := (~ #9 #9) +#55 := [refl]: #54 +#57 := [nnf-pos #55]: #56 +#39 := [asserted]: #10 +#58 := [mp~ #39 #57]: #10 +#618 := [mp #58 #617]: #613 +#286 := (not #613) +#244 := (or #286 #602) +#194 := (= uf_4 uf_2) +#264 := (iff #21 #194) +#587 := (or #286 #264) +#249 := (iff #587 #244) +#251 := (iff #244 #244) +#589 := [rewrite]: #251 +#260 := (iff #264 #602) +#282 := (iff #194 #195) +#196 := [rewrite]: #282 +#603 := [monotonicity #196]: #260 +#250 := [monotonicity #603]: #249 +#590 := [trans #250 #589]: #249 +#248 := [quant-inst]: #587 +#591 := [mp #248 #590]: #244 +#598 := [unit-resolution #591 #618]: #602 +decl ?x6!3 :: T2 +#63 := ?x6!3 +#64 := (up_1 uf_4 ?x6!3) +#283 := (iff #64 #195) +#214 := (or #286 #283) +#281 := (iff #64 #194) +#287 := (or #286 #281) +#288 := (iff #287 #214) +#604 := (iff #214 #214) +#606 := [rewrite]: #604 +#274 := (iff #281 #283) +#285 := [monotonicity #196]: #274 +#267 := [monotonicity #285]: #288 +#261 := [trans #267 #606]: #288 +#284 := [quant-inst]: #287 +#393 := [mp #284 #261]: #214 +#596 := [unit-resolution #393 #618]: #283 +#600 := (not #283) +#586 := (or #600 #195) +#122 := (not #21) +#599 := [hypothesis]: #122 +#127 := (or #21 #64) +#119 := (forall (vars (?x6 T2)) #116) +#128 := (or #122 #119) +#135 := (and #127 #128) +#129 := (and #128 #127) +#136 := (iff #129 #135) +#137 := [rewrite]: #136 +#20 := (exists (vars (?x6 T2)) #19) +#42 := (not #20) +#43 := (iff #21 #42) +#130 := (~ #43 #129) +#120 := (~ #42 #119) +#117 := (~ #116 #116) +#118 := [refl]: #117 +#121 := [nnf-neg #118]: #120 +#113 := (not #42) +#114 := (~ #113 #64) +#88 := (~ #20 #64) +#89 := [sk]: #88 +#115 := [nnf-neg #89]: #114 +#125 := (~ #21 #21) +#126 := [refl]: #125 +#123 := (~ #122 #122) +#124 := [refl]: #123 +#131 := [nnf-pos #124 #126 #115 #121]: #130 +#22 := (iff #20 #21) +#23 := (not #22) +#44 := (iff #23 #43) +#45 := [rewrite]: #44 +#41 := [asserted]: #23 +#48 := [mp #41 #45]: #43 +#132 := [mp~ #48 #131]: #129 +#133 := [mp #132 #137]: #135 +#134 := [and-elim #133]: #127 +#585 := [unit-resolution #134 #599]: #64 +#608 := (not #64) +#609 := (or #600 #608 #195) +#610 := [def-axiom]: #609 +#292 := [unit-resolution #610 #585]: #586 +#308 := [unit-resolution #292 #596]: #195 +#272 := (not #195) +#592 := (not #602) +#309 := (or #592 #272) +#593 := (or #592 #21 #272) +#588 := [def-axiom]: #593 +#310 := [unit-resolution #588 #599]: #309 +#296 := [unit-resolution #310 #308 #598]: false +#311 := [lemma #296]: #21 +#641 := (or #122 #636) +#642 := (iff #128 #641) +#639 := (iff #119 #636) +#637 := (iff #116 #116) +#638 := [refl]: #637 +#640 := [quant-intro #638]: #639 +#643 := [monotonicity #640]: #642 +#138 := [and-elim #133]: #128 +#644 := [mp #138 #643]: #641 +#594 := [unit-resolution #644 #311]: #636 +#595 := (not #636) +#597 := (or #595 #122) +#235 := [quant-inst]: #597 +[unit-resolution #235 #311 #594]: false +unsat diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_fol_04 --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_fol_04 Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,13 @@ +(benchmark Isabelle +:extrasorts ( T1 T2) +:extrafuns ( + (uf_2 T1) + (uf_3 T1) + ) +:extrapreds ( + (up_1 T1) + ) +:assumption (if_then_else (up_1 uf_2) (not (exists (?x1 T1) (up_1 ?x1))) (forall (?x2 T1) (not (up_1 ?x2)))) +:assumption (not (implies (up_1 uf_2) (up_1 uf_3))) +:formula true +) diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_fol_04.proof --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_fol_04.proof Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,56 @@ +#2 := false +decl up_1 :: (-> T1 bool) +decl uf_2 :: T1 +#4 := uf_2 +#5 := (up_1 uf_2) +decl uf_3 :: T1 +#13 := uf_3 +#14 := (up_1 uf_3) +#34 := (not #5) +#35 := (or #34 #14) +#38 := (not #35) +#15 := (implies #5 #14) +#16 := (not #15) +#39 := (iff #16 #38) +#36 := (iff #15 #35) +#37 := [rewrite]: #36 +#40 := [monotonicity #37]: #39 +#33 := [asserted]: #16 +#43 := [mp #33 #40]: #38 +#41 := [not-or-elim #43]: #5 +#6 := (:var 0 T1) +#7 := (up_1 #6) +#536 := (pattern #7) +#10 := (not #7) +#537 := (forall (vars (?x2 T1)) (:pat #536) #10) +#11 := (forall (vars (?x2 T1)) #10) +#540 := (iff #11 #537) +#538 := (iff #10 #10) +#539 := [refl]: #538 +#541 := [quant-intro #539]: #540 +#8 := (exists (vars (?x1 T1)) #7) +#9 := (not #8) +#45 := (~ #9 #11) +#50 := (~ #10 #10) +#51 := [refl]: #50 +#59 := [nnf-neg #51]: #45 +#12 := (ite #5 #9 #11) +#57 := (iff #12 #9) +#1 := true +#52 := (ite true #9 #11) +#55 := (iff #52 #9) +#56 := [rewrite]: #55 +#53 := (iff #12 #52) +#48 := (iff #5 true) +#49 := [iff-true #41]: #48 +#54 := [monotonicity #49]: #53 +#58 := [trans #54 #56]: #57 +#32 := [asserted]: #12 +#47 := [mp #32 #58]: #9 +#60 := [mp~ #47 #59]: #11 +#542 := [mp #60 #541]: #537 +#119 := (not #537) +#206 := (or #119 #34) +#120 := [quant-inst]: #206 +[unit-resolution #120 #542 #41]: false +unsat diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_hol_01 --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_hol_01 Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,18 @@ +(benchmark Isabelle +:extrasorts ( T2 T3 T1) +:extrafuns ( + (uf_2 T1 T2 T3 T1) + (uf_1 T1 T2 T3) + (uf_6 T1) + (uf_3 T2) + (uf_4 T2) + (uf_5 T2) + (uf_7 T3) + (uf_8 T3) + ) +:assumption (forall (?x1 T1) (?x2 T2) (?x3 T3) (?x4 T2) (= (uf_1 (uf_2 ?x1 ?x2 ?x3) ?x4) (ite (= ?x4 ?x2) ?x3 (uf_1 ?x1 ?x4)))) +:assumption (forall (?x5 T1) (?x6 T2) (?x7 T3) (= (uf_1 (uf_2 ?x5 ?x6 ?x7) ?x6) ?x7)) +:assumption (and (not (= uf_3 uf_4)) (not (= uf_3 uf_5))) +:assumption (not (= (uf_1 (uf_2 (uf_2 uf_6 uf_4 uf_7) uf_5 uf_8) uf_3) (uf_1 uf_6 uf_3))) +:formula true +) diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_hol_01.proof --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_hol_01.proof Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,181 @@ +#2 := false +decl uf_1 :: (-> T1 T2 T3) +decl uf_3 :: T2 +#22 := uf_3 +decl uf_6 :: T1 +#30 := uf_6 +#36 := (uf_1 uf_6 uf_3) +decl uf_2 :: (-> T1 T2 T3 T1) +decl uf_8 :: T3 +#33 := uf_8 +decl uf_5 :: T2 +#26 := uf_5 +decl uf_7 :: T3 +#31 := uf_7 +decl uf_4 :: T2 +#23 := uf_4 +#32 := (uf_2 uf_6 uf_4 uf_7) +#34 := (uf_2 #32 uf_5 uf_8) +#35 := (uf_1 #34 uf_3) +#37 := (= #35 #36) +#223 := (uf_1 #32 uf_4) +#214 := (uf_2 uf_6 uf_4 #223) +#552 := (uf_1 #214 uf_3) +#555 := (= #552 #36) +#560 := (= #36 #552) +#556 := (= #223 #552) +#24 := (= uf_3 uf_4) +#561 := (ite #24 #556 #560) +#8 := (:var 0 T2) +#6 := (:var 1 T3) +#5 := (:var 2 T2) +#4 := (:var 3 T1) +#7 := (uf_2 #4 #5 #6) +#9 := (uf_1 #7 #8) +#575 := (pattern #9) +#11 := (uf_1 #4 #8) +#100 := (= #9 #11) +#99 := (= #6 #9) +#55 := (= #5 #8) +#83 := (ite #55 #99 #100) +#576 := (forall (vars (?x1 T1) (?x2 T2) (?x3 T3) (?x4 T2)) (:pat #575) #83) +#90 := (forall (vars (?x1 T1) (?x2 T2) (?x3 T3) (?x4 T2)) #83) +#579 := (iff #90 #576) +#577 := (iff #83 #83) +#578 := [refl]: #577 +#580 := [quant-intro #578]: #579 +#58 := (ite #55 #6 #11) +#61 := (= #9 #58) +#64 := (forall (vars (?x1 T1) (?x2 T2) (?x3 T3) (?x4 T2)) #61) +#87 := (iff #64 #90) +#84 := (iff #61 #83) +#89 := [rewrite]: #84 +#88 := [quant-intro #89]: #87 +#93 := (~ #64 #64) +#91 := (~ #61 #61) +#92 := [refl]: #91 +#94 := [nnf-pos #92]: #93 +#10 := (= #8 #5) +#12 := (ite #10 #6 #11) +#13 := (= #9 #12) +#14 := (forall (vars (?x1 T1) (?x2 T2) (?x3 T3) (?x4 T2)) #13) +#65 := (iff #14 #64) +#62 := (iff #13 #61) +#59 := (= #12 #58) +#56 := (iff #10 #55) +#57 := [rewrite]: #56 +#60 := [monotonicity #57]: #59 +#63 := [monotonicity #60]: #62 +#66 := [quant-intro #63]: #65 +#54 := [asserted]: #14 +#69 := [mp #54 #66]: #64 +#95 := [mp~ #69 #94]: #64 +#85 := [mp #95 #88]: #90 +#581 := [mp #85 #580]: #576 +#250 := (not #576) +#548 := (or #250 #561) +#551 := (= uf_4 uf_3) +#557 := (ite #551 #556 #555) +#549 := (or #250 #557) +#271 := (iff #549 #548) +#273 := (iff #548 #548) +#259 := [rewrite]: #273 +#559 := (iff #557 #561) +#198 := (iff #555 #560) +#199 := [rewrite]: #198 +#193 := (iff #551 #24) +#558 := [rewrite]: #193 +#562 := [monotonicity #558 #199]: #559 +#272 := [monotonicity #562]: #271 +#274 := [trans #272 #259]: #271 +#255 := [quant-inst]: #549 +#165 := [mp #255 #274]: #548 +#510 := [unit-resolution #165 #581]: #561 +#544 := (not #561) +#497 := (or #544 #560) +#25 := (not #24) +#27 := (= uf_3 uf_5) +#28 := (not #27) +#29 := (and #25 #28) +#75 := [asserted]: #29 +#79 := [and-elim #75]: #25 +#268 := (or #544 #24 #560) +#542 := [def-axiom]: #268 +#499 := [unit-resolution #542 #79]: #497 +#491 := [unit-resolution #499 #510]: #560 +#493 := [symm #491]: #555 +#494 := (= #35 #552) +#157 := (uf_1 #32 uf_3) +#503 := (= #157 #552) +#502 := (= #552 #157) +#509 := (= #214 #32) +#415 := (= #223 uf_7) +#566 := (= uf_7 #223) +#17 := (:var 0 T3) +#16 := (:var 1 T2) +#15 := (:var 2 T1) +#18 := (uf_2 #15 #16 #17) +#582 := (pattern #18) +#19 := (uf_1 #18 #16) +#68 := (= #17 #19) +#584 := (forall (vars (?x5 T1) (?x6 T2) (?x7 T3)) (:pat #582) #68) +#72 := (forall (vars (?x5 T1) (?x6 T2) (?x7 T3)) #68) +#583 := (iff #72 #584) +#586 := (iff #584 #584) +#587 := [rewrite]: #586 +#585 := [rewrite]: #583 +#588 := [trans #585 #587]: #583 +#82 := (~ #72 #72) +#96 := (~ #68 #68) +#97 := [refl]: #96 +#78 := [nnf-pos #97]: #82 +#20 := (= #19 #17) +#21 := (forall (vars (?x5 T1) (?x6 T2) (?x7 T3)) #20) +#73 := (iff #21 #72) +#70 := (iff #20 #68) +#71 := [rewrite]: #70 +#74 := [quant-intro #71]: #73 +#67 := [asserted]: #21 +#77 := [mp #67 #74]: #72 +#98 := [mp~ #77 #78]: #72 +#589 := [mp #98 #588]: #584 +#211 := (not #584) +#212 := (or #211 #566) +#213 := [quant-inst]: #212 +#414 := [unit-resolution #213 #589]: #566 +#416 := [symm #414]: #415 +#506 := [monotonicity #416]: #509 +#498 := [monotonicity #506]: #502 +#492 := [symm #498]: #503 +#244 := (= #35 #157) +#158 := (= uf_8 #35) +#248 := (ite #27 #158 #244) +#247 := (or #250 #248) +#245 := (= uf_5 uf_3) +#159 := (ite #245 #158 #244) +#251 := (or #250 #159) +#567 := (iff #251 #247) +#224 := (iff #247 #247) +#356 := [rewrite]: #224 +#249 := (iff #159 #248) +#246 := (iff #245 #27) +#237 := [rewrite]: #246 +#177 := [monotonicity #237]: #249 +#569 := [monotonicity #177]: #567 +#563 := [trans #569 #356]: #567 +#230 := [quant-inst]: #251 +#235 := [mp #230 #563]: #247 +#488 := [unit-resolution #235 #581]: #248 +#236 := (not #248) +#490 := (or #236 #244) +#80 := [and-elim #75]: #28 +#572 := (or #236 #27 #244) +#573 := [def-axiom]: #572 +#500 := [unit-resolution #573 #80]: #490 +#501 := [unit-resolution #500 #488]: #244 +#495 := [trans #501 #492]: #494 +#489 := [trans #495 #493]: #37 +#38 := (not #37) +#76 := [asserted]: #38 +[unit-resolution #76 #489]: false +unsat diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_hol_02 --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_hol_02 Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,13 @@ +(benchmark Isabelle +:extrasorts ( T2 T1 T3) +:extrafuns ( + (uf_2 T1) + (uf_3 T2) + ) +:extrapreds ( + (up_4 T1 T2) + (up_1 T1 T2) + ) +:assumption (not (or (iff (up_1 uf_2 uf_3) (and (up_4 uf_2 uf_3) true)) (or (iff (up_1 uf_2 uf_3) true) (iff (up_4 uf_2 uf_3) true)))) +:formula true +) diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_hol_02.proof --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_hol_02.proof Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,62 @@ +#2 := false +decl up_4 :: (-> T1 T2 bool) +decl uf_3 :: T2 +#5 := uf_3 +decl uf_2 :: T1 +#4 := uf_2 +#7 := (up_4 uf_2 uf_3) +#60 := (not #7) +decl up_1 :: (-> T1 T2 bool) +#6 := (up_1 uf_2 uf_3) +#33 := (iff #6 #7) +#49 := (or #6 #7 #33) +#52 := (not #49) +#1 := true +#11 := (iff #7 true) +#10 := (iff #6 true) +#12 := (or #10 #11) +#8 := (and #7 true) +#9 := (iff #6 #8) +#13 := (or #9 #12) +#14 := (not #13) +#55 := (iff #14 #52) +#40 := (or #6 #7) +#43 := (or #33 #40) +#46 := (not #43) +#53 := (iff #46 #52) +#50 := (iff #43 #49) +#51 := [rewrite]: #50 +#54 := [monotonicity #51]: #53 +#47 := (iff #14 #46) +#44 := (iff #13 #43) +#41 := (iff #12 #40) +#38 := (iff #11 #7) +#39 := [rewrite]: #38 +#36 := (iff #10 #6) +#37 := [rewrite]: #36 +#42 := [monotonicity #37 #39]: #41 +#34 := (iff #9 #33) +#31 := (iff #8 #7) +#32 := [rewrite]: #31 +#35 := [monotonicity #32]: #34 +#45 := [monotonicity #35 #42]: #44 +#48 := [monotonicity #45]: #47 +#56 := [trans #48 #54]: #55 +#30 := [asserted]: #14 +#57 := [mp #30 #56]: #52 +#61 := [not-or-elim #57]: #60 +#58 := (not #6) +#59 := [not-or-elim #57]: #58 +#72 := (or #7 #6) +#66 := (iff #7 #58) +#62 := (not #33) +#64 := (iff #62 #66) +#67 := [rewrite]: #64 +#63 := [not-or-elim #57]: #62 +#68 := [mp #63 #67]: #66 +#69 := (not #66) +#70 := (or #7 #6 #69) +#71 := [def-axiom]: #70 +#73 := [unit-resolution #71 #68]: #72 +[unit-resolution #73 #59 #61]: false +unsat diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_hol_03 --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_hol_03 Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,13 @@ +(benchmark Isabelle +:extrasorts ( T1 T2) +:extrafuns ( + (uf_3 T2) + (uf_1 T1 T1) + (uf_2 T2 T2) + (uf_4 T1) + ) +:assumption (forall (?x1 T1) (= (uf_1 ?x1) ?x1)) +:assumption (forall (?x2 T2) (iff (= (uf_2 ?x2) uf_3) (= ?x2 uf_3))) +:assumption (not (and (= (uf_1 uf_4) uf_4) (iff (= (uf_2 uf_3) uf_3) true))) +:formula true +) diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_hol_03.proof --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_hol_03.proof Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,120 @@ +#2 := false +decl uf_1 :: (-> T1 T1) +decl uf_4 :: T1 +#15 := uf_4 +#16 := (uf_1 uf_4) +#48 := (= uf_4 #16) +#83 := (not #48) +decl uf_2 :: (-> T2 T2) +decl uf_3 :: T2 +#10 := uf_3 +#18 := (uf_2 uf_3) +#51 := (= uf_3 #18) +#84 := (not #51) +#556 := [hypothesis]: #84 +#8 := (:var 0 T2) +#9 := (uf_2 #8) +#575 := (pattern #9) +#12 := (= #8 uf_3) +#11 := (= #9 uf_3) +#13 := (iff #11 #12) +#576 := (forall (vars (?x2 T2)) (:pat #575) #13) +#14 := (forall (vars (?x2 T2)) #13) +#579 := (iff #14 #576) +#577 := (iff #13 #13) +#578 := [refl]: #577 +#580 := [quant-intro #578]: #579 +#70 := (~ #14 #14) +#80 := (~ #13 #13) +#81 := [refl]: #80 +#67 := [nnf-pos #81]: #70 +#45 := [asserted]: #14 +#82 := [mp~ #45 #67]: #14 +#581 := [mp #82 #580]: #576 +#242 := (not #576) +#170 := (or #242 #51) +#150 := (= uf_3 uf_3) +#19 := (= #18 uf_3) +#237 := (iff #19 #150) +#243 := (or #242 #237) +#244 := (iff #243 #170) +#560 := (iff #170 #170) +#562 := [rewrite]: #560 +#230 := (iff #237 #51) +#1 := true +#54 := (iff #51 true) +#57 := (iff #54 #51) +#58 := [rewrite]: #57 +#152 := (iff #237 #54) +#151 := (iff #150 true) +#238 := [rewrite]: #151 +#52 := (iff #19 #51) +#53 := [rewrite]: #52 +#239 := [monotonicity #53 #238]: #152 +#241 := [trans #239 #58]: #230 +#223 := [monotonicity #241]: #244 +#217 := [trans #223 #562]: #244 +#240 := [quant-inst]: #243 +#349 := [mp #240 #217]: #170 +#228 := [unit-resolution #349 #581 #556]: false +#229 := [lemma #228]: #51 +#71 := (or #83 #84) +#61 := (and #48 #51) +#64 := (not #61) +#90 := (iff #64 #71) +#72 := (not #71) +#85 := (not #72) +#88 := (iff #85 #71) +#89 := [rewrite]: #88 +#86 := (iff #64 #85) +#73 := (iff #61 #72) +#74 := [rewrite]: #73 +#87 := [monotonicity #74]: #86 +#91 := [trans #87 #89]: #90 +#20 := (iff #19 true) +#17 := (= #16 uf_4) +#21 := (and #17 #20) +#22 := (not #21) +#65 := (iff #22 #64) +#62 := (iff #21 #61) +#59 := (iff #20 #51) +#55 := (iff #20 #54) +#56 := [monotonicity #53]: #55 +#60 := [trans #56 #58]: #59 +#49 := (iff #17 #48) +#50 := [rewrite]: #49 +#63 := [monotonicity #50 #60]: #62 +#66 := [monotonicity #63]: #65 +#46 := [asserted]: #22 +#69 := [mp #46 #66]: #64 +#92 := [mp #69 #91]: #71 +#563 := [unit-resolution #92 #229]: #83 +#4 := (:var 0 T1) +#5 := (uf_1 #4) +#568 := (pattern #5) +#39 := (= #4 #5) +#569 := (forall (vars (?x1 T1)) (:pat #568) #39) +#42 := (forall (vars (?x1 T1)) #39) +#572 := (iff #42 #569) +#570 := (iff #39 #39) +#571 := [refl]: #570 +#573 := [quant-intro #571]: #572 +#77 := (~ #42 #42) +#75 := (~ #39 #39) +#76 := [refl]: #75 +#78 := [nnf-pos #76]: #77 +#6 := (= #5 #4) +#7 := (forall (vars (?x1 T1)) #6) +#43 := (iff #7 #42) +#40 := (iff #6 #39) +#41 := [rewrite]: #40 +#44 := [quant-intro #41]: #43 +#38 := [asserted]: #7 +#47 := [mp #38 #44]: #42 +#79 := [mp~ #47 #78]: #42 +#574 := [mp #79 #573]: #569 +#565 := (not #569) +#566 := (or #565 #48) +#561 := [quant-inst]: #566 +[unit-resolution #561 #574 #563]: false +unsat diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_hol_04 --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_hol_04 Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,18 @@ +(benchmark Isabelle +:extrasorts ( T2 T3 T1) +:extrafuns ( + (uf_2 T1 T2 T3 T1) + (uf_1 T1 T2 T3) + (uf_6 T1) + (uf_3 T2) + (uf_4 T2) + (uf_5 T2) + (uf_7 T3) + (uf_8 T3) + ) +:assumption (forall (?x1 T1) (?x2 T2) (?x3 T3) (?x4 T2) (= (uf_1 (uf_2 ?x1 ?x2 ?x3) ?x4) (ite (= ?x4 ?x2) ?x3 (uf_1 ?x1 ?x4)))) +:assumption (forall (?x5 T1) (?x6 T2) (?x7 T3) (= (uf_1 (uf_2 ?x5 ?x6 ?x7) ?x6) ?x7)) +:assumption (and (not (= uf_3 uf_4)) (not (= uf_3 uf_5))) +:assumption (not (= (uf_1 (uf_2 (uf_2 uf_6 uf_4 uf_7) uf_5 uf_8) uf_3) (uf_1 uf_6 uf_3))) +:formula true +) diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_hol_04.proof --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_hol_04.proof Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,181 @@ +#2 := false +decl uf_1 :: (-> T1 T2 T3) +decl uf_3 :: T2 +#22 := uf_3 +decl uf_6 :: T1 +#30 := uf_6 +#36 := (uf_1 uf_6 uf_3) +decl uf_2 :: (-> T1 T2 T3 T1) +decl uf_8 :: T3 +#33 := uf_8 +decl uf_5 :: T2 +#26 := uf_5 +decl uf_7 :: T3 +#31 := uf_7 +decl uf_4 :: T2 +#23 := uf_4 +#32 := (uf_2 uf_6 uf_4 uf_7) +#34 := (uf_2 #32 uf_5 uf_8) +#35 := (uf_1 #34 uf_3) +#37 := (= #35 #36) +#223 := (uf_1 #32 uf_4) +#214 := (uf_2 uf_6 uf_4 #223) +#552 := (uf_1 #214 uf_3) +#555 := (= #552 #36) +#560 := (= #36 #552) +#556 := (= #223 #552) +#24 := (= uf_3 uf_4) +#561 := (ite #24 #556 #560) +#8 := (:var 0 T2) +#6 := (:var 1 T3) +#5 := (:var 2 T2) +#4 := (:var 3 T1) +#7 := (uf_2 #4 #5 #6) +#9 := (uf_1 #7 #8) +#575 := (pattern #9) +#11 := (uf_1 #4 #8) +#100 := (= #9 #11) +#99 := (= #6 #9) +#55 := (= #5 #8) +#83 := (ite #55 #99 #100) +#576 := (forall (vars (?x1 T1) (?x2 T2) (?x3 T3) (?x4 T2)) (:pat #575) #83) +#90 := (forall (vars (?x1 T1) (?x2 T2) (?x3 T3) (?x4 T2)) #83) +#579 := (iff #90 #576) +#577 := (iff #83 #83) +#578 := [refl]: #577 +#580 := [quant-intro #578]: #579 +#58 := (ite #55 #6 #11) +#61 := (= #9 #58) +#64 := (forall (vars (?x1 T1) (?x2 T2) (?x3 T3) (?x4 T2)) #61) +#87 := (iff #64 #90) +#84 := (iff #61 #83) +#89 := [rewrite]: #84 +#88 := [quant-intro #89]: #87 +#93 := (~ #64 #64) +#91 := (~ #61 #61) +#92 := [refl]: #91 +#94 := [nnf-pos #92]: #93 +#10 := (= #8 #5) +#12 := (ite #10 #6 #11) +#13 := (= #9 #12) +#14 := (forall (vars (?x1 T1) (?x2 T2) (?x3 T3) (?x4 T2)) #13) +#65 := (iff #14 #64) +#62 := (iff #13 #61) +#59 := (= #12 #58) +#56 := (iff #10 #55) +#57 := [rewrite]: #56 +#60 := [monotonicity #57]: #59 +#63 := [monotonicity #60]: #62 +#66 := [quant-intro #63]: #65 +#54 := [asserted]: #14 +#69 := [mp #54 #66]: #64 +#95 := [mp~ #69 #94]: #64 +#85 := [mp #95 #88]: #90 +#581 := [mp #85 #580]: #576 +#250 := (not #576) +#548 := (or #250 #561) +#551 := (= uf_4 uf_3) +#557 := (ite #551 #556 #555) +#549 := (or #250 #557) +#271 := (iff #549 #548) +#273 := (iff #548 #548) +#259 := [rewrite]: #273 +#559 := (iff #557 #561) +#198 := (iff #555 #560) +#199 := [rewrite]: #198 +#193 := (iff #551 #24) +#558 := [rewrite]: #193 +#562 := [monotonicity #558 #199]: #559 +#272 := [monotonicity #562]: #271 +#274 := [trans #272 #259]: #271 +#255 := [quant-inst]: #549 +#165 := [mp #255 #274]: #548 +#510 := [unit-resolution #165 #581]: #561 +#544 := (not #561) +#497 := (or #544 #560) +#25 := (not #24) +#27 := (= uf_3 uf_5) +#28 := (not #27) +#29 := (and #25 #28) +#75 := [asserted]: #29 +#79 := [and-elim #75]: #25 +#268 := (or #544 #24 #560) +#542 := [def-axiom]: #268 +#499 := [unit-resolution #542 #79]: #497 +#491 := [unit-resolution #499 #510]: #560 +#493 := [symm #491]: #555 +#494 := (= #35 #552) +#157 := (uf_1 #32 uf_3) +#503 := (= #157 #552) +#502 := (= #552 #157) +#509 := (= #214 #32) +#415 := (= #223 uf_7) +#566 := (= uf_7 #223) +#17 := (:var 0 T3) +#16 := (:var 1 T2) +#15 := (:var 2 T1) +#18 := (uf_2 #15 #16 #17) +#582 := (pattern #18) +#19 := (uf_1 #18 #16) +#68 := (= #17 #19) +#584 := (forall (vars (?x5 T1) (?x6 T2) (?x7 T3)) (:pat #582) #68) +#72 := (forall (vars (?x5 T1) (?x6 T2) (?x7 T3)) #68) +#583 := (iff #72 #584) +#586 := (iff #584 #584) +#587 := [rewrite]: #586 +#585 := [rewrite]: #583 +#588 := [trans #585 #587]: #583 +#82 := (~ #72 #72) +#96 := (~ #68 #68) +#97 := [refl]: #96 +#78 := [nnf-pos #97]: #82 +#20 := (= #19 #17) +#21 := (forall (vars (?x5 T1) (?x6 T2) (?x7 T3)) #20) +#73 := (iff #21 #72) +#70 := (iff #20 #68) +#71 := [rewrite]: #70 +#74 := [quant-intro #71]: #73 +#67 := [asserted]: #21 +#77 := [mp #67 #74]: #72 +#98 := [mp~ #77 #78]: #72 +#589 := [mp #98 #588]: #584 +#211 := (not #584) +#212 := (or #211 #566) +#213 := [quant-inst]: #212 +#414 := [unit-resolution #213 #589]: #566 +#416 := [symm #414]: #415 +#506 := [monotonicity #416]: #509 +#498 := [monotonicity #506]: #502 +#492 := [symm #498]: #503 +#244 := (= #35 #157) +#158 := (= uf_8 #35) +#248 := (ite #27 #158 #244) +#247 := (or #250 #248) +#245 := (= uf_5 uf_3) +#159 := (ite #245 #158 #244) +#251 := (or #250 #159) +#567 := (iff #251 #247) +#224 := (iff #247 #247) +#356 := [rewrite]: #224 +#249 := (iff #159 #248) +#246 := (iff #245 #27) +#237 := [rewrite]: #246 +#177 := [monotonicity #237]: #249 +#569 := [monotonicity #177]: #567 +#563 := [trans #569 #356]: #567 +#230 := [quant-inst]: #251 +#235 := [mp #230 #563]: #247 +#488 := [unit-resolution #235 #581]: #248 +#236 := (not #248) +#490 := (or #236 #244) +#80 := [and-elim #75]: #28 +#572 := (or #236 #27 #244) +#573 := [def-axiom]: #572 +#500 := [unit-resolution #573 #80]: #490 +#501 := [unit-resolution #500 #488]: #244 +#495 := [trans #501 #492]: #494 +#489 := [trans #495 #493]: #37 +#38 := (not #37) +#76 := [asserted]: #38 +[unit-resolution #76 #489]: false +unsat diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_hol_05 --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_hol_05 Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,20 @@ +(benchmark Isabelle +:extrasorts ( T1 T2 T3) +:extrafuns ( + (uf_6 Int T2) + (uf_1 T1 T3 T3) + (uf_4 T3) + (uf_2 T2 T3 T3) + (uf_3 T1 T2 T2) + (uf_7 T2 Int) + (uf_5 T1) + ) +:assumption (forall (?x1 T1) (?x2 T2) (?x3 T3) (= (uf_1 ?x1 (uf_2 ?x2 ?x3)) (uf_2 (uf_3 ?x1 ?x2) (uf_1 ?x1 ?x3)))) +:assumption (forall (?x4 T1) (= (uf_1 ?x4 uf_4) uf_4)) +:assumption (forall (?x5 T2) (= (uf_3 uf_5 ?x5) (uf_6 (+ (uf_7 ?x5) 1)))) +:assumption (forall (?x6 T2) (= (uf_6 (uf_7 ?x6)) ?x6)) +:assumption (forall (?x7 Int) (implies (<= 0 ?x7) (= (uf_7 (uf_6 ?x7)) ?x7))) +:assumption (forall (?x8 Int) (implies (< ?x8 0) (= (uf_7 (uf_6 ?x8)) 0))) +:assumption (not (= (uf_1 uf_5 (uf_2 (uf_6 0) (uf_2 (uf_6 1) uf_4))) (uf_2 (uf_6 1) (uf_2 (uf_6 2) uf_4)))) +:formula true +) diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_hol_05.proof --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_hol_05.proof Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,464 @@ +#2 := false +decl uf_2 :: (-> T2 T3 T3) +decl uf_4 :: T3 +#15 := uf_4 +decl uf_6 :: (-> int T2) +#48 := 2::int +#49 := (uf_6 2::int) +#50 := (uf_2 #49 uf_4) +#23 := 1::int +#44 := (uf_6 1::int) +#51 := (uf_2 #44 #50) +decl uf_1 :: (-> T1 T3 T3) +#45 := (uf_2 #44 uf_4) +#31 := 0::int +#43 := (uf_6 0::int) +#46 := (uf_2 #43 #45) +decl uf_5 :: T1 +#19 := uf_5 +#47 := (uf_1 uf_5 #46) +#52 := (= #47 #51) +#266 := (uf_1 uf_5 #45) +decl uf_3 :: (-> T1 T2 T2) +#352 := (uf_3 uf_5 #43) +#267 := (uf_2 #352 #266) +#797 := (= #267 #51) +#795 := (= #51 #267) +#758 := (= #50 #266) +#521 := (uf_1 uf_5 uf_4) +#522 := (uf_3 uf_5 #44) +#523 := (uf_2 #522 #521) +#756 := (= #523 #266) +#616 := (= #266 #523) +#6 := (:var 0 T3) +#4 := (:var 2 T1) +#10 := (uf_1 #4 #6) +#5 := (:var 1 T2) +#9 := (uf_3 #4 #5) +#11 := (uf_2 #9 #10) +#683 := (pattern #11) +#7 := (uf_2 #5 #6) +#8 := (uf_1 #4 #7) +#682 := (pattern #8) +#12 := (= #8 #11) +#684 := (forall (vars (?x1 T1) (?x2 T2) (?x3 T3)) (:pat #682 #683) #12) +#13 := (forall (vars (?x1 T1) (?x2 T2) (?x3 T3)) #12) +#687 := (iff #13 #684) +#685 := (iff #12 #12) +#686 := [refl]: #685 +#688 := [quant-intro #686]: #687 +#195 := (~ #13 #13) +#193 := (~ #12 #12) +#194 := [refl]: #193 +#196 := [nnf-pos #194]: #195 +#69 := [asserted]: #13 +#197 := [mp~ #69 #196]: #13 +#689 := [mp #197 #688]: #684 +#345 := (not #684) +#604 := (or #345 #616) +#606 := [quant-inst]: #604 +#277 := [unit-resolution #606 #689]: #616 +#757 := [symm #277]: #756 +#754 := (= #50 #523) +#569 := (= uf_4 #521) +#14 := (:var 0 T1) +#16 := (uf_1 #14 uf_4) +#690 := (pattern #16) +#71 := (= uf_4 #16) +#691 := (forall (vars (?x4 T1)) (:pat #690) #71) +#74 := (forall (vars (?x4 T1)) #71) +#694 := (iff #74 #691) +#692 := (iff #71 #71) +#693 := [refl]: #692 +#695 := [quant-intro #693]: #694 +#180 := (~ #74 #74) +#198 := (~ #71 #71) +#199 := [refl]: #198 +#178 := [nnf-pos #199]: #180 +#17 := (= #16 uf_4) +#18 := (forall (vars (?x4 T1)) #17) +#75 := (iff #18 #74) +#72 := (iff #17 #71) +#73 := [rewrite]: #72 +#76 := [quant-intro #73]: #75 +#70 := [asserted]: #18 +#79 := [mp #70 #76]: #74 +#200 := [mp~ #79 #178]: #74 +#696 := [mp #200 #695]: #691 +#572 := (not #691) +#573 := (or #572 #569) +#574 := [quant-inst]: #573 +#282 := [unit-resolution #574 #696]: #569 +#752 := (= #49 #522) +decl uf_7 :: (-> T2 int) +#666 := (uf_7 #44) +#595 := (+ 1::int #666) +#597 := (uf_6 #595) +#748 := (= #597 #522) +#605 := (= #522 #597) +#20 := (:var 0 T2) +#22 := (uf_7 #20) +#698 := (pattern #22) +#21 := (uf_3 uf_5 #20) +#697 := (pattern #21) +#78 := (+ 1::int #22) +#82 := (uf_6 #78) +#85 := (= #21 #82) +#699 := (forall (vars (?x5 T2)) (:pat #697 #698) #85) +#88 := (forall (vars (?x5 T2)) #85) +#702 := (iff #88 #699) +#700 := (iff #85 #85) +#701 := [refl]: #700 +#703 := [quant-intro #701]: #702 +#181 := (~ #88 #88) +#201 := (~ #85 #85) +#202 := [refl]: #201 +#182 := [nnf-pos #202]: #181 +#24 := (+ #22 1::int) +#25 := (uf_6 #24) +#26 := (= #21 #25) +#27 := (forall (vars (?x5 T2)) #26) +#89 := (iff #27 #88) +#86 := (iff #26 #85) +#83 := (= #25 #82) +#80 := (= #24 #78) +#81 := [rewrite]: #80 +#84 := [monotonicity #81]: #83 +#87 := [monotonicity #84]: #86 +#90 := [quant-intro #87]: #89 +#77 := [asserted]: #27 +#93 := [mp #77 #90]: #88 +#203 := [mp~ #93 #182]: #88 +#704 := [mp #203 #703]: #699 +#607 := (not #699) +#600 := (or #607 #605) +#601 := [quant-inst]: #600 +#269 := [unit-resolution #601 #704]: #605 +#749 := [symm #269]: #748 +#750 := (= #49 #597) +#499 := (uf_7 #597) +#337 := (uf_6 #499) +#318 := (= #337 #597) +#28 := (uf_6 #22) +#92 := (= #20 #28) +#705 := (forall (vars (?x6 T2)) (:pat #698) #92) +#96 := (forall (vars (?x6 T2)) #92) +#706 := (iff #96 #705) +#708 := (iff #705 #705) +#709 := [rewrite]: #708 +#707 := [rewrite]: #706 +#710 := [trans #707 #709]: #706 +#183 := (~ #96 #96) +#204 := (~ #92 #92) +#205 := [refl]: #204 +#184 := [nnf-pos #205]: #183 +#29 := (= #28 #20) +#30 := (forall (vars (?x6 T2)) #29) +#97 := (iff #30 #96) +#94 := (iff #29 #92) +#95 := [rewrite]: #94 +#98 := [quant-intro #95]: #97 +#91 := [asserted]: #30 +#101 := [mp #91 #98]: #96 +#206 := [mp~ #101 #184]: #96 +#711 := [mp #206 #710]: #705 +#376 := (not #705) +#325 := (or #376 #318) +#316 := (= #597 #337) +#326 := (or #376 #316) +#328 := (iff #326 #325) +#329 := (iff #325 #325) +#310 := [rewrite]: #329 +#323 := (iff #316 #318) +#324 := [rewrite]: #323 +#317 := [monotonicity #324]: #328 +#312 := [trans #317 #310]: #328 +#327 := [quant-inst]: #326 +#313 := [mp #327 #312]: #325 +#271 := [unit-resolution #313 #711]: #318 +#746 := (= #49 #337) +#744 := (= 2::int #499) +#742 := (= #499 2::int) +#578 := -1::int +#513 := (* -1::int #666) +#514 := (+ #499 #513) +#474 := (<= #514 1::int) +#512 := (= #514 1::int) +#504 := (>= #666 -1::int) +#586 := (>= #666 1::int) +#378 := (= #666 1::int) +#32 := (:var 0 int) +#34 := (uf_6 #32) +#712 := (pattern #34) +#118 := (>= #32 0::int) +#119 := (not #118) +#35 := (uf_7 #34) +#100 := (= #32 #35) +#125 := (or #100 #119) +#713 := (forall (vars (?x7 int)) (:pat #712) #125) +#130 := (forall (vars (?x7 int)) #125) +#716 := (iff #130 #713) +#714 := (iff #125 #125) +#715 := [refl]: #714 +#717 := [quant-intro #715]: #716 +#185 := (~ #130 #130) +#207 := (~ #125 #125) +#208 := [refl]: #207 +#186 := [nnf-pos #208]: #185 +#36 := (= #35 #32) +#33 := (<= 0::int #32) +#37 := (implies #33 #36) +#38 := (forall (vars (?x7 int)) #37) +#133 := (iff #38 #130) +#107 := (not #33) +#108 := (or #107 #100) +#113 := (forall (vars (?x7 int)) #108) +#131 := (iff #113 #130) +#128 := (iff #108 #125) +#122 := (or #119 #100) +#126 := (iff #122 #125) +#127 := [rewrite]: #126 +#123 := (iff #108 #122) +#120 := (iff #107 #119) +#116 := (iff #33 #118) +#117 := [rewrite]: #116 +#121 := [monotonicity #117]: #120 +#124 := [monotonicity #121]: #123 +#129 := [trans #124 #127]: #128 +#132 := [quant-intro #129]: #131 +#114 := (iff #38 #113) +#111 := (iff #37 #108) +#104 := (implies #33 #100) +#109 := (iff #104 #108) +#110 := [rewrite]: #109 +#105 := (iff #37 #104) +#102 := (iff #36 #100) +#103 := [rewrite]: #102 +#106 := [monotonicity #103]: #105 +#112 := [trans #106 #110]: #111 +#115 := [quant-intro #112]: #114 +#134 := [trans #115 #132]: #133 +#99 := [asserted]: #38 +#135 := [mp #99 #134]: #130 +#209 := [mp~ #135 #186]: #130 +#718 := [mp #209 #717]: #713 +#673 := (not #713) +#365 := (or #673 #378) +#307 := (>= 1::int 0::int) +#668 := (not #307) +#669 := (= 1::int #666) +#655 := (or #669 #668) +#366 := (or #673 #655) +#645 := (iff #366 #365) +#360 := (iff #365 #365) +#643 := [rewrite]: #360 +#654 := (iff #655 #378) +#374 := (or #378 false) +#653 := (iff #374 #378) +#650 := [rewrite]: #653 +#375 := (iff #655 #374) +#651 := (iff #668 false) +#1 := true +#670 := (not true) +#677 := (iff #670 false) +#678 := [rewrite]: #677 +#381 := (iff #668 #670) +#379 := (iff #307 true) +#380 := [rewrite]: #379 +#274 := [monotonicity #380]: #381 +#652 := [trans #274 #678]: #651 +#656 := (iff #669 #378) +#363 := [rewrite]: #656 +#649 := [monotonicity #363 #652]: #375 +#364 := [trans #649 #650]: #654 +#646 := [monotonicity #364]: #645 +#647 := [trans #646 #643]: #645 +#367 := [quant-inst]: #366 +#644 := [mp #367 #647]: #365 +#272 := [unit-resolution #644 #718]: #378 +#270 := (not #378) +#273 := (or #270 #586) +#725 := [th-lemma]: #273 +#726 := [unit-resolution #725 #272]: #586 +#727 := (not #586) +#728 := (or #727 #504) +#729 := [th-lemma]: #728 +#730 := [unit-resolution #729 #726]: #504 +#481 := (not #504) +#496 := (or #673 #481 #512) +#509 := (>= #595 0::int) +#468 := (not #509) +#501 := (= #595 #499) +#503 := (or #501 #468) +#497 := (or #673 #503) +#470 := (iff #497 #496) +#491 := (or #481 #512) +#498 := (or #673 #491) +#467 := (iff #498 #496) +#469 := [rewrite]: #467 +#459 := (iff #497 #498) +#494 := (iff #503 #491) +#488 := (or #512 #481) +#492 := (iff #488 #491) +#493 := [rewrite]: #492 +#489 := (iff #503 #488) +#486 := (iff #468 #481) +#525 := (iff #509 #504) +#480 := [rewrite]: #525 +#487 := [monotonicity #480]: #486 +#510 := (iff #501 #512) +#524 := [rewrite]: #510 +#490 := [monotonicity #524 #487]: #489 +#495 := [trans #490 #493]: #494 +#460 := [monotonicity #495]: #459 +#471 := [trans #460 #469]: #470 +#482 := [quant-inst]: #497 +#473 := [mp #482 #471]: #496 +#731 := [unit-resolution #473 #718 #730]: #512 +#732 := (not #512) +#733 := (or #732 #474) +#734 := [th-lemma]: #733 +#735 := [unit-resolution #734 #731]: #474 +#475 := (>= #514 1::int) +#736 := (or #732 #475) +#737 := [th-lemma]: #736 +#738 := [unit-resolution #737 #731]: #475 +#582 := (<= #666 1::int) +#739 := (or #270 #582) +#740 := [th-lemma]: #739 +#741 := [unit-resolution #740 #272]: #582 +#743 := [th-lemma #726 #741 #738 #735]: #742 +#745 := [symm #743]: #744 +#747 := [monotonicity #745]: #746 +#751 := [trans #747 #271]: #750 +#753 := [trans #751 #749]: #752 +#755 := [monotonicity #753 #282]: #754 +#759 := [trans #755 #757]: #758 +#792 := (= #44 #352) +#358 := (uf_7 #43) +#613 := (+ 1::int #358) +#617 := (uf_6 #613) +#788 := (= #617 #352) +#598 := (= #352 #617) +#608 := (or #607 #598) +#609 := [quant-inst]: #608 +#760 := [unit-resolution #609 #704]: #598 +#789 := [symm #760]: #788 +#790 := (= #44 #617) +#575 := (uf_7 #617) +#390 := (uf_6 #575) +#382 := (= #390 #617) +#385 := (or #376 #382) +#392 := (= #617 #390) +#386 := (or #376 #392) +#387 := (iff #386 #385) +#369 := (iff #385 #385) +#370 := [rewrite]: #369 +#383 := (iff #392 #382) +#384 := [rewrite]: #383 +#368 := [monotonicity #384]: #387 +#361 := [trans #368 #370]: #387 +#377 := [quant-inst]: #386 +#371 := [mp #377 #361]: #385 +#761 := [unit-resolution #371 #711]: #382 +#786 := (= #44 #390) +#784 := (= 1::int #575) +#782 := (= #575 1::int) +#568 := (* -1::int #575) +#579 := (+ #358 #568) +#535 := (<= #579 -1::int) +#557 := (= #579 -1::int) +#561 := (>= #358 -1::int) +#585 := (>= #358 0::int) +#676 := (= #358 0::int) +#315 := (or #673 #676) +#268 := (>= 0::int 0::int) +#354 := (not #268) +#355 := (= 0::int #358) +#359 := (or #355 #354) +#657 := (or #673 #359) +#320 := (iff #657 #315) +#322 := (iff #315 #315) +#659 := [rewrite]: #322 +#672 := (iff #359 #676) +#675 := (or #676 false) +#330 := (iff #675 #676) +#335 := [rewrite]: #330 +#681 := (iff #359 #675) +#679 := (iff #354 false) +#343 := (iff #354 #670) +#332 := (iff #268 true) +#463 := [rewrite]: #332 +#344 := [monotonicity #463]: #343 +#680 := [trans #344 #678]: #679 +#338 := (iff #355 #676) +#674 := [rewrite]: #338 +#671 := [monotonicity #674 #680]: #681 +#331 := [trans #671 #335]: #672 +#321 := [monotonicity #331]: #320 +#660 := [trans #321 #659]: #320 +#319 := [quant-inst]: #657 +#661 := [mp #319 #660]: #315 +#762 := [unit-resolution #661 #718]: #676 +#763 := (not #676) +#764 := (or #763 #585) +#765 := [th-lemma]: #764 +#766 := [unit-resolution #765 #762]: #585 +#767 := (not #585) +#768 := (or #767 #561) +#769 := [th-lemma]: #768 +#770 := [unit-resolution #769 #766]: #561 +#564 := (not #561) +#549 := (or #673 #557 #564) +#570 := (>= #613 0::int) +#571 := (not #570) +#576 := (= #613 #575) +#577 := (or #576 #571) +#552 := (or #673 #577) +#530 := (iff #552 #549) +#551 := (or #557 #564) +#554 := (or #673 #551) +#556 := (iff #554 #549) +#529 := [rewrite]: #556 +#555 := (iff #552 #554) +#547 := (iff #577 #551) +#559 := (iff #571 #564) +#562 := (iff #570 #561) +#563 := [rewrite]: #562 +#565 := [monotonicity #563]: #559 +#558 := (iff #576 #557) +#560 := [rewrite]: #558 +#548 := [monotonicity #560 #565]: #547 +#550 := [monotonicity #548]: #555 +#531 := [trans #550 #529]: #530 +#553 := [quant-inst]: #552 +#424 := [mp #553 #531]: #549 +#771 := [unit-resolution #424 #718 #770]: #557 +#772 := (not #557) +#773 := (or #772 #535) +#774 := [th-lemma]: #773 +#775 := [unit-resolution #774 #771]: #535 +#536 := (>= #579 -1::int) +#776 := (or #772 #536) +#777 := [th-lemma]: #776 +#778 := [unit-resolution #777 #771]: #536 +#584 := (<= #358 0::int) +#779 := (or #763 #584) +#780 := [th-lemma]: #779 +#781 := [unit-resolution #780 #762]: #584 +#783 := [th-lemma #766 #781 #778 #775]: #782 +#785 := [symm #783]: #784 +#787 := [monotonicity #785]: #786 +#791 := [trans #787 #761]: #790 +#793 := [trans #791 #789]: #792 +#796 := [monotonicity #793 #759]: #795 +#798 := [symm #796]: #797 +#353 := (= #47 #267) +#356 := (or #345 #353) +#357 := [quant-inst]: #356 +#794 := [unit-resolution #357 #689]: #353 +#799 := [trans #794 #798]: #52 +#53 := (not #52) +#177 := [asserted]: #53 +[unit-resolution #177 #799]: false +unsat diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_hol_06 --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_hol_06 Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,8 @@ +(benchmark Isabelle +:extrasorts ( T1 T2) +:extrapreds ( + (up_1 T1) + ) +:assumption (not (or (forall (?x1 T1) (up_1 ?x1)) (not (forall (?x2 T1) (up_1 ?x2))))) +:formula true +) diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_hol_06.proof --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_hol_06.proof Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,21 @@ +#2 := false +decl up_1 :: (-> T1 bool) +#4 := (:var 0 T1) +#5 := (up_1 #4) +#6 := (forall (vars (?x1 T1)) #5) +#7 := (not #6) +#8 := (or #6 #7) +#9 := (not #8) +#33 := (iff #9 false) +#1 := true +#28 := (not true) +#31 := (iff #28 false) +#32 := [rewrite]: #31 +#29 := (iff #9 #28) +#26 := (iff #8 true) +#27 := [rewrite]: #26 +#30 := [monotonicity #27]: #29 +#34 := [trans #30 #32]: #33 +#25 := [asserted]: #9 +[mp #25 #34]: false +unsat diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_hol_07 --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_hol_07 Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,14 @@ +(benchmark Isabelle +:extrasorts ( T1) +:extrafuns ( + (uf_1 Int T1) + (uf_3 T1 T1) + (uf_2 T1 Int) + ) +:assumption (forall (?x1 T1) (= (uf_1 (uf_2 ?x1)) ?x1)) +:assumption (forall (?x2 Int) (implies (<= 0 ?x2) (= (uf_2 (uf_1 ?x2)) ?x2))) +:assumption (forall (?x3 Int) (implies (< ?x3 0) (= (uf_2 (uf_1 ?x3)) 0))) +:assumption (forall (?x4 T1) (= (uf_3 ?x4) (ite (< (uf_2 ?x4) 10) ?x4 (uf_3 (uf_1 (- (uf_2 ?x4) 10)))))) +:assumption (not (= (uf_3 (uf_1 (* 4 (uf_2 (uf_3 (uf_1 4)))))) (uf_1 6))) +:formula true +) diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_hol_07.proof --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_hol_07.proof Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,373 @@ +#2 := false +decl uf_1 :: (-> int T1) +#37 := 6::int +#38 := (uf_1 6::int) +decl uf_3 :: (-> T1 T1) +decl uf_2 :: (-> T1 int) +#30 := 4::int +#31 := (uf_1 4::int) +#32 := (uf_3 #31) +#33 := (uf_2 #32) +#34 := (* 4::int #33) +#35 := (uf_1 #34) +#36 := (uf_3 #35) +#39 := (= #36 #38) +#548 := (uf_3 #38) +#394 := (= #548 #38) +#549 := (= #38 #548) +#523 := (uf_2 #38) +#142 := -10::int +#513 := (+ -10::int #523) +#537 := (uf_1 #513) +#538 := (uf_3 #537) +#514 := (= #538 #548) +#22 := 10::int +#539 := (>= #523 10::int) +#506 := (ite #539 #514 #549) +#4 := (:var 0 T1) +#21 := (uf_3 #4) +#708 := (pattern #21) +#5 := (uf_2 #4) +#687 := (pattern #5) +#209 := (= #4 #21) +#143 := (+ -10::int #5) +#146 := (uf_1 #143) +#149 := (uf_3 #146) +#208 := (= #21 #149) +#163 := (>= #5 10::int) +#190 := (ite #163 #208 #209) +#709 := (forall (vars (?x4 T1)) (:pat #687 #708) #190) +#193 := (forall (vars (?x4 T1)) #190) +#712 := (iff #193 #709) +#710 := (iff #190 #190) +#711 := [refl]: #710 +#713 := [quant-intro #711]: #712 +#168 := (ite #163 #149 #4) +#173 := (= #21 #168) +#176 := (forall (vars (?x4 T1)) #173) +#210 := (iff #176 #193) +#191 := (iff #173 #190) +#192 := [rewrite]: #191 +#211 := [quant-intro #192]: #210 +#188 := (~ #176 #176) +#205 := (~ #173 #173) +#206 := [refl]: #205 +#189 := [nnf-pos #206]: #188 +#24 := (- #5 10::int) +#25 := (uf_1 #24) +#26 := (uf_3 #25) +#23 := (< #5 10::int) +#27 := (ite #23 #4 #26) +#28 := (= #21 #27) +#29 := (forall (vars (?x4 T1)) #28) +#179 := (iff #29 #176) +#152 := (ite #23 #4 #149) +#155 := (= #21 #152) +#158 := (forall (vars (?x4 T1)) #155) +#177 := (iff #158 #176) +#174 := (iff #155 #173) +#171 := (= #152 #168) +#161 := (not #163) +#165 := (ite #161 #4 #149) +#169 := (= #165 #168) +#170 := [rewrite]: #169 +#166 := (= #152 #165) +#162 := (iff #23 #161) +#164 := [rewrite]: #162 +#167 := [monotonicity #164]: #166 +#172 := [trans #167 #170]: #171 +#175 := [monotonicity #172]: #174 +#178 := [quant-intro #175]: #177 +#159 := (iff #29 #158) +#156 := (iff #28 #155) +#153 := (= #27 #152) +#150 := (= #26 #149) +#147 := (= #25 #146) +#144 := (= #24 #143) +#145 := [rewrite]: #144 +#148 := [monotonicity #145]: #147 +#151 := [monotonicity #148]: #150 +#154 := [monotonicity #151]: #153 +#157 := [monotonicity #154]: #156 +#160 := [quant-intro #157]: #159 +#180 := [trans #160 #178]: #179 +#141 := [asserted]: #29 +#181 := [mp #141 #180]: #176 +#207 := [mp~ #181 #189]: #176 +#212 := [mp #207 #211]: #193 +#714 := [mp #212 #713]: #709 +#681 := (not #709) +#517 := (or #681 #506) +#533 := (= #548 #538) +#507 := (ite #539 #533 #549) +#518 := (or #681 #507) +#529 := (iff #518 #517) +#530 := (iff #517 #517) +#485 := [rewrite]: #530 +#508 := (iff #507 #506) +#473 := (iff #533 #514) +#504 := [rewrite]: #473 +#515 := [monotonicity #504]: #508 +#509 := [monotonicity #515]: #529 +#486 := [trans #509 #485]: #529 +#519 := [quant-inst]: #518 +#491 := [mp #519 #486]: #517 +#484 := [unit-resolution #491 #714]: #506 +#493 := (not #539) +#465 := (<= #523 6::int) +#526 := (= #523 6::int) +#10 := (:var 0 int) +#12 := (uf_1 #10) +#695 := (pattern #12) +#9 := 0::int +#82 := (>= #10 0::int) +#83 := (not #82) +#13 := (uf_2 #12) +#64 := (= #10 #13) +#89 := (or #64 #83) +#696 := (forall (vars (?x2 int)) (:pat #695) #89) +#94 := (forall (vars (?x2 int)) #89) +#699 := (iff #94 #696) +#697 := (iff #89 #89) +#698 := [refl]: #697 +#700 := [quant-intro #698]: #699 +#185 := (~ #94 #94) +#199 := (~ #89 #89) +#200 := [refl]: #199 +#183 := [nnf-pos #200]: #185 +#14 := (= #13 #10) +#11 := (<= 0::int #10) +#15 := (implies #11 #14) +#16 := (forall (vars (?x2 int)) #15) +#97 := (iff #16 #94) +#71 := (not #11) +#72 := (or #71 #64) +#77 := (forall (vars (?x2 int)) #72) +#95 := (iff #77 #94) +#92 := (iff #72 #89) +#86 := (or #83 #64) +#90 := (iff #86 #89) +#91 := [rewrite]: #90 +#87 := (iff #72 #86) +#84 := (iff #71 #83) +#80 := (iff #11 #82) +#81 := [rewrite]: #80 +#85 := [monotonicity #81]: #84 +#88 := [monotonicity #85]: #87 +#93 := [trans #88 #91]: #92 +#96 := [quant-intro #93]: #95 +#78 := (iff #16 #77) +#75 := (iff #15 #72) +#68 := (implies #11 #64) +#73 := (iff #68 #72) +#74 := [rewrite]: #73 +#69 := (iff #15 #68) +#66 := (iff #14 #64) +#67 := [rewrite]: #66 +#70 := [monotonicity #67]: #69 +#76 := [trans #70 #74]: #75 +#79 := [quant-intro #76]: #78 +#98 := [trans #79 #96]: #97 +#63 := [asserted]: #16 +#99 := [mp #63 #98]: #94 +#201 := [mp~ #99 #183]: #94 +#701 := [mp #201 #700]: #696 +#671 := (not #696) +#615 := (or #671 #526) +#520 := (>= 6::int 0::int) +#522 := (not #520) +#516 := (= 6::int #523) +#524 := (or #516 #522) +#604 := (or #671 #524) +#606 := (iff #604 #615) +#601 := (iff #615 #615) +#608 := [rewrite]: #601 +#614 := (iff #524 #526) +#603 := (or #526 false) +#612 := (iff #603 #526) +#613 := [rewrite]: #612 +#600 := (iff #524 #603) +#609 := (iff #522 false) +#1 := true +#327 := (not true) +#666 := (iff #327 false) +#667 := [rewrite]: #666 +#618 := (iff #522 #327) +#528 := (iff #520 true) +#621 := [rewrite]: #528 +#622 := [monotonicity #621]: #618 +#611 := [trans #622 #667]: #609 +#525 := (iff #516 #526) +#527 := [rewrite]: #525 +#602 := [monotonicity #527 #611]: #600 +#610 := [trans #602 #613]: #614 +#607 := [monotonicity #610]: #606 +#592 := [trans #607 #608]: #606 +#605 := [quant-inst]: #604 +#593 := [mp #605 #592]: #615 +#454 := [unit-resolution #593 #701]: #526 +#303 := (not #526) +#462 := (or #303 #465) +#458 := [th-lemma]: #462 +#463 := [unit-resolution #458 #454]: #465 +#442 := (not #465) +#445 := (or #442 #493) +#449 := [th-lemma]: #445 +#451 := [unit-resolution #449 #463]: #493 +#492 := (not #506) +#496 := (or #492 #539 #549) +#497 := [def-axiom]: #496 +#452 := [unit-resolution #497 #451 #484]: #549 +#395 := [symm #452]: #394 +#397 := (= #36 #548) +#372 := (uf_2 #35) +#576 := (+ -10::int #372) +#568 := (uf_1 #576) +#569 := (uf_3 #568) +#408 := (= #569 #548) +#401 := (= #568 #38) +#422 := (= #576 6::int) +#677 := (uf_2 #31) +#365 := -1::int +#478 := (* -1::int #677) +#479 := (+ #33 #478) +#480 := (<= #479 0::int) +#476 := (= #33 #677) +#431 := (= #32 #31) +#589 := (= #31 #32) +#590 := (+ -10::int #677) +#587 := (uf_1 #590) +#591 := (uf_3 #587) +#571 := (= #32 #591) +#572 := (>= #677 10::int) +#574 := (ite #572 #571 #589) +#577 := (or #681 #574) +#578 := [quant-inst]: #577 +#450 := [unit-resolution #578 #714]: #574 +#580 := (not #572) +#552 := (<= #677 4::int) +#324 := (= #677 4::int) +#674 := (or #671 #324) +#343 := (>= 4::int 0::int) +#679 := (not #343) +#336 := (= 4::int #677) +#678 := (or #336 #679) +#660 := (or #671 #678) +#368 := (iff #660 #674) +#384 := (iff #674 #674) +#385 := [rewrite]: #384 +#312 := (iff #678 #324) +#669 := (or #324 false) +#672 := (iff #669 #324) +#311 := [rewrite]: #672 +#306 := (iff #678 #669) +#668 := (iff #679 false) +#664 := (iff #679 #327) +#325 := (iff #343 true) +#326 := [rewrite]: #325 +#665 := [monotonicity #326]: #664 +#663 := [trans #665 #667]: #668 +#320 := (iff #336 #324) +#662 := [rewrite]: #320 +#670 := [monotonicity #662 #663]: #306 +#673 := [trans #670 #311]: #312 +#383 := [monotonicity #673]: #368 +#386 := [trans #383 #385]: #368 +#661 := [quant-inst]: #660 +#278 := [mp #661 #386]: #674 +#453 := [unit-resolution #278 #701]: #324 +#441 := (not #324) +#444 := (or #441 #552) +#446 := [th-lemma]: #444 +#447 := [unit-resolution #446 #453]: #552 +#443 := (not #552) +#448 := (or #443 #580) +#438 := [th-lemma]: #448 +#428 := [unit-resolution #438 #447]: #580 +#579 := (not #574) +#583 := (or #579 #572 #589) +#573 := [def-axiom]: #583 +#430 := [unit-resolution #573 #428 #450]: #589 +#434 := [symm #430]: #431 +#435 := [monotonicity #434]: #476 +#439 := (not #476) +#432 := (or #439 #480) +#440 := [th-lemma]: #432 +#433 := [unit-resolution #440 #435]: #480 +#481 := (>= #479 0::int) +#436 := (or #439 #481) +#437 := [th-lemma]: #436 +#423 := [unit-resolution #437 #435]: #481 +#553 := (>= #677 4::int) +#425 := (or #441 #553) +#426 := [th-lemma]: #425 +#424 := [unit-resolution #426 #453]: #553 +#648 := (* -1::int #372) +#652 := (+ #34 #648) +#631 := (<= #652 0::int) +#649 := (= #652 0::int) +#370 := (>= #34 0::int) +#409 := (not #481) +#427 := (not #553) +#411 := (or #370 #427 #409) +#412 := [th-lemma]: #411 +#413 := [unit-resolution #412 #424 #423]: #370 +#371 := (not #370) +#640 := (or #371 #649) +#488 := (or #671 #371 #649) +#650 := (= #34 #372) +#651 := (or #650 #371) +#489 := (or #671 #651) +#630 := (iff #489 #488) +#632 := (or #671 #640) +#635 := (iff #632 #488) +#629 := [rewrite]: #635 +#633 := (iff #489 #632) +#641 := (iff #651 #640) +#643 := (or #649 #371) +#645 := (iff #643 #640) +#646 := [rewrite]: #645 +#644 := (iff #651 #643) +#653 := (iff #650 #649) +#642 := [rewrite]: #653 +#639 := [monotonicity #642]: #644 +#647 := [trans #639 #646]: #641 +#634 := [monotonicity #647]: #633 +#636 := [trans #634 #629]: #630 +#490 := [quant-inst]: #489 +#637 := [mp #490 #636]: #488 +#414 := [unit-resolution #637 #701]: #640 +#415 := [unit-resolution #414 #413]: #649 +#416 := (not #649) +#417 := (or #416 #631) +#418 := [th-lemma]: #417 +#419 := [unit-resolution #418 #415]: #631 +#638 := (>= #652 0::int) +#420 := (or #416 #638) +#421 := [th-lemma]: #420 +#410 := [unit-resolution #421 #415]: #638 +#399 := [th-lemma #410 #419 #424 #447 #423 #433]: #422 +#402 := [monotonicity #399]: #401 +#393 := [monotonicity #402]: #408 +#564 := (= #36 #569) +#575 := (= #35 #36) +#570 := (>= #372 10::int) +#556 := (ite #570 #564 #575) +#554 := (or #681 #556) +#557 := [quant-inst]: #554 +#403 := [unit-resolution #557 #714]: #556 +#404 := (not #631) +#405 := (or #570 #404 #427 #409) +#406 := [th-lemma]: #405 +#407 := [unit-resolution #406 #419 #424 #423]: #570 +#559 := (not #570) +#558 := (not #556) +#560 := (or #558 #559 #564) +#555 := [def-axiom]: #560 +#400 := [unit-resolution #555 #407 #403]: #564 +#396 := [trans #400 #393]: #397 +#398 := [trans #396 #395]: #39 +#40 := (not #39) +#182 := [asserted]: #40 +[unit-resolution #182 #398]: false +unsat diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_hol_08 --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_hol_08 Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,24 @@ +(benchmark Isabelle +:extrasorts ( T2 T1 T4 T3) +:extrafuns ( + (uf_3 Int T1) + (uf_7 T2 T4 T4) + (uf_1 T2 T1 T1) + (uf_6 T3 T4 Int) + (uf_4 T1 Int) + (uf_5 T2) + (uf_2 T2) + (uf_10 Int) + (uf_8 T3) + (uf_9 T4) + ) +:assumption (forall (?x1 T1) (= (uf_1 uf_2 ?x1) (uf_3 (div (uf_4 ?x1) 2)))) +:assumption (forall (?x2 T1) (= (uf_1 uf_5 ?x2) (uf_3 (mod (uf_4 ?x2) 2)))) +:assumption (forall (?x3 T1) (= (uf_3 (uf_4 ?x3)) ?x3)) +:assumption (forall (?x4 Int) (implies (<= 0 ?x4) (= (uf_4 (uf_3 ?x4)) ?x4))) +:assumption (forall (?x5 Int) (implies (< ?x5 0) (= (uf_4 (uf_3 ?x5)) 0))) +:assumption (forall (?x6 T3) (?x7 T4) (= (mod (uf_6 ?x6 ?x7) 2) (mod (uf_6 ?x6 (uf_7 uf_5 ?x7)) 2))) +:assumption (forall (?x8 T3) (?x9 T4) (= (+ (* (uf_6 ?x8 (uf_7 uf_2 ?x9)) 2) (uf_6 ?x8 (uf_7 uf_5 ?x9))) (uf_6 ?x8 ?x9))) +:assumption (iff (= (uf_6 uf_8 uf_9) uf_10) (implies (= (mod (uf_6 uf_8 (uf_7 uf_5 uf_9)) 2) (mod uf_10 2)) (not (= (uf_6 uf_8 (uf_7 uf_2 uf_9)) (div (- uf_10 (uf_6 uf_8 (uf_7 uf_5 uf_9))) 2))))) +:formula true +) diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_hol_08.proof --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_hol_08.proof Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,410 @@ +#2 := false +#22 := 0::int +decl uf_6 :: (-> T3 T4 int) +decl uf_7 :: (-> T2 T4 T4) +decl uf_9 :: T4 +#50 := uf_9 +decl uf_2 :: T2 +#4 := uf_2 +#59 := (uf_7 uf_2 uf_9) +decl uf_8 :: T3 +#49 := uf_8 +#60 := (uf_6 uf_8 #59) +#204 := -2::int +#683 := (* -2::int #60) +decl uf_5 :: T2 +#13 := uf_5 +#54 := (uf_7 uf_5 uf_9) +#55 := (uf_6 uf_8 #54) +#172 := -1::int +#218 := (* -1::int #55) +#685 := (+ #218 #683) +#51 := (uf_6 uf_8 uf_9) +#686 := (+ #51 #685) +#679 := (>= #686 0::int) +#687 := (= #686 0::int) +#35 := (:var 0 T4) +#43 := (uf_7 uf_2 #35) +#34 := (:var 1 T3) +#44 := (uf_6 #34 #43) +#819 := (pattern #44) +#38 := (uf_7 uf_5 #35) +#39 := (uf_6 #34 #38) +#812 := (pattern #39) +#205 := (* -2::int #44) +#203 := (* -1::int #39) +#206 := (+ #203 #205) +#36 := (uf_6 #34 #35) +#207 := (+ #36 #206) +#208 := (= #207 0::int) +#820 := (forall (vars (?x8 T3) (?x9 T4)) (:pat #812 #819) #208) +#211 := (forall (vars (?x8 T3) (?x9 T4)) #208) +#823 := (iff #211 #820) +#821 := (iff #208 #208) +#822 := [refl]: #821 +#824 := [quant-intro #822]: #823 +#279 := (~ #211 #211) +#305 := (~ #208 #208) +#306 := [refl]: #305 +#280 := [nnf-pos #306]: #279 +#8 := 2::int +#45 := (* #44 2::int) +#46 := (+ #45 #39) +#47 := (= #46 #36) +#48 := (forall (vars (?x8 T3) (?x9 T4)) #47) +#214 := (iff #48 #211) +#171 := (* 2::int #44) +#187 := (+ #39 #171) +#195 := (= #36 #187) +#200 := (forall (vars (?x8 T3) (?x9 T4)) #195) +#212 := (iff #200 #211) +#209 := (iff #195 #208) +#210 := [rewrite]: #209 +#213 := [quant-intro #210]: #212 +#201 := (iff #48 #200) +#198 := (iff #47 #195) +#192 := (= #187 #36) +#196 := (iff #192 #195) +#197 := [rewrite]: #196 +#193 := (iff #47 #192) +#190 := (= #46 #187) +#184 := (+ #171 #39) +#188 := (= #184 #187) +#189 := [rewrite]: #188 +#185 := (= #46 #184) +#182 := (= #45 #171) +#183 := [rewrite]: #182 +#186 := [monotonicity #183]: #185 +#191 := [trans #186 #189]: #190 +#194 := [monotonicity #191]: #193 +#199 := [trans #194 #197]: #198 +#202 := [quant-intro #199]: #201 +#215 := [trans #202 #213]: #214 +#170 := [asserted]: #48 +#216 := [mp #170 #215]: #211 +#307 := [mp~ #216 #280]: #211 +#825 := [mp #307 #824]: #820 +#689 := (not #820) +#675 := (or #689 #687) +#676 := [quant-inst]: #675 +#536 := [unit-resolution #676 #825]: #687 +#537 := (not #687) +#533 := (or #537 #679) +#538 := [th-lemma]: #533 +#528 := [unit-resolution #538 #536]: #679 +decl uf_10 :: int +#52 := uf_10 +#219 := (+ uf_10 #218) +#222 := (div #219 2::int) +#251 := (* -1::int #222) +#252 := (+ #60 #251) +#449 := (<= #252 0::int) +#399 := (not #449) +#253 := (= #252 0::int) +#256 := (not #253) +#57 := (mod uf_10 2::int) +#243 := (* -1::int #57) +#56 := (mod #55 2::int) +#244 := (+ #56 #243) +#245 := (= #244 0::int) +#448 := (>= #244 0::int) +#688 := (mod #51 2::int) +#666 := (* -1::int #688) +#667 := (+ #56 #666) +#660 := (>= #667 0::int) +#668 := (= #667 0::int) +#40 := (mod #39 2::int) +#173 := (* -1::int #40) +#37 := (mod #36 2::int) +#174 := (+ #37 #173) +#175 := (= #174 0::int) +#813 := (forall (vars (?x6 T3) (?x7 T4)) (:pat #812) #175) +#178 := (forall (vars (?x6 T3) (?x7 T4)) #175) +#816 := (iff #178 #813) +#814 := (iff #175 #175) +#815 := [refl]: #814 +#817 := [quant-intro #815]: #816 +#277 := (~ #178 #178) +#302 := (~ #175 #175) +#303 := [refl]: #302 +#278 := [nnf-pos #303]: #277 +#41 := (= #37 #40) +#42 := (forall (vars (?x6 T3) (?x7 T4)) #41) +#179 := (iff #42 #178) +#176 := (iff #41 #175) +#177 := [rewrite]: #176 +#180 := [quant-intro #177]: #179 +#169 := [asserted]: #42 +#181 := [mp #169 #180]: #178 +#304 := [mp~ #181 #278]: #178 +#818 := [mp #304 #817]: #813 +#673 := (not #813) +#663 := (or #673 #668) +#756 := (* -1::int #56) +#684 := (+ #688 #756) +#680 := (= #684 0::int) +#674 := (or #673 #680) +#653 := (iff #674 #663) +#656 := (iff #663 #663) +#657 := [rewrite]: #656 +#671 := (iff #680 #668) +#677 := (+ #756 #688) +#662 := (= #677 0::int) +#669 := (iff #662 #668) +#670 := [rewrite]: #669 +#664 := (iff #680 #662) +#681 := (= #684 #677) +#661 := [rewrite]: #681 +#665 := [monotonicity #661]: #664 +#672 := [trans #665 #670]: #671 +#655 := [monotonicity #672]: #653 +#658 := [trans #655 #657]: #653 +#652 := [quant-inst]: #674 +#659 := [mp #652 #658]: #663 +#394 := [unit-resolution #659 #818]: #668 +#552 := (not #668) +#514 := (or #552 #660) +#517 := [th-lemma]: #514 +#499 := [unit-resolution #517 #394]: #660 +#503 := (not #448) +#414 := [hypothesis]: #503 +#561 := (+ #57 #666) +#709 := (<= #561 0::int) +#602 := (= #57 #688) +#468 := (= #688 #57) +#53 := (= #51 uf_10) +#248 := (not #245) +#259 := (or #248 #256) +#362 := (mod #219 2::int) +#699 := (>= #362 0::int) +#1 := true +#81 := [true-axiom]: true +#604 := (or false #699) +#506 := [th-lemma]: #604 +#507 := [unit-resolution #506 #81]: #699 +#628 := (* -1::int uf_10) +#623 := (+ #51 #628) +#629 := (<= #623 0::int) +#498 := (not #629) +#597 := (>= #623 0::int) +#381 := (not #259) +#508 := [hypothesis]: #381 +#450 := (or #259 #245) +#441 := [def-axiom]: #450 +#509 := [unit-resolution #441 #508]: #245 +#510 := (or #248 #448) +#511 := [th-lemma]: #510 +#500 := [unit-resolution #511 #509]: #448 +#743 := (div uf_10 2::int) +#723 := (* -2::int #743) +#545 := (* -2::int #688) +#546 := (+ #545 #723) +#646 := (div #51 2::int) +#645 := (* -2::int #646) +#547 := (+ #645 #546) +#605 := (* -2::int #57) +#549 := (+ #605 #547) +#594 := (* 2::int #56) +#550 := (+ #594 #549) +#598 := (* 2::int uf_10) +#551 := (+ #598 #550) +#563 := (>= #551 2::int) +#520 := (not #563) +#361 := (<= #244 0::int) +#512 := (or #248 #361) +#489 := [th-lemma]: #512 +#491 := [unit-resolution #489 #509]: #361 +#363 := (>= #252 0::int) +#452 := (or #259 #253) +#453 := [def-axiom]: #452 +#492 := [unit-resolution #453 #508]: #253 +#493 := (or #256 #363) +#494 := [th-lemma]: #493 +#495 := [unit-resolution #494 #492]: #363 +#556 := (not #361) +#573 := (not #363) +#521 := (or #520 #573 #556) +#703 := (>= #362 2::int) +#704 := (not #703) +#599 := (or false #704) +#620 := [th-lemma]: #599 +#575 := [unit-resolution #620 #81]: #704 +#654 := (<= #667 0::int) +#548 := (or #552 #654) +#553 := [th-lemma]: #548 +#532 := [unit-resolution #553 #394]: #654 +#651 := (+ #645 #666) +#624 := (+ #51 #651) +#626 := (<= #624 0::int) +#650 := (= #624 0::int) +#535 := (or false #650) +#539 := [th-lemma]: #535 +#541 := [unit-resolution #539 #81]: #650 +#542 := (not #650) +#540 := (or #542 #626) +#543 := [th-lemma]: #540 +#531 := [unit-resolution #543 #541]: #626 +#587 := [hypothesis]: #361 +#724 := (+ #243 #723) +#725 := (+ uf_10 #724) +#727 := (<= #725 0::int) +#722 := (= #725 0::int) +#576 := (or false #722) +#581 := [th-lemma]: #576 +#582 := [unit-resolution #581 #81]: #722 +#583 := (not #722) +#584 := (or #583 #727) +#585 := [th-lemma]: #584 +#586 := [unit-resolution #585 #582]: #727 +#534 := [hypothesis]: #563 +#555 := [hypothesis]: #363 +#616 := (* -1::int #362) +#615 := (* -2::int #222) +#617 := (+ #615 #616) +#618 := (+ #218 #617) +#711 := (+ uf_10 #618) +#708 := (<= #711 0::int) +#606 := (= #711 0::int) +#562 := (or false #606) +#564 := [th-lemma]: #562 +#565 := [unit-resolution #564 #81]: #606 +#566 := (not #606) +#568 := (or #566 #708) +#569 := [th-lemma]: #568 +#570 := [unit-resolution #569 #565]: #708 +#518 := [th-lemma #570 #555 #528 #534 #586 #587 #531 #532 #575]: false +#524 := [lemma #518]: #521 +#496 := [unit-resolution #524 #495 #491]: #520 +#504 := (or #597 #563 #503) +#529 := (not #597) +#522 := [hypothesis]: #529 +#519 := (>= #624 0::int) +#530 := (or #542 #519) +#523 := [th-lemma]: #530 +#526 := [unit-resolution #523 #541]: #519 +#527 := [hypothesis]: #448 +#721 := (>= #725 0::int) +#513 := (or #583 #721) +#515 := [th-lemma]: #513 +#516 := [unit-resolution #515 #582]: #721 +#501 := [th-lemma #499 #516 #527 #526 #522]: #563 +#525 := [hypothesis]: #520 +#502 := [unit-resolution #525 #501]: false +#505 := [lemma #502]: #504 +#497 := [unit-resolution #505 #496 #500]: #597 +#485 := (or #498 #529) +#558 := (not #53) +#440 := (or #558 #259) +#262 := (iff #53 #259) +#61 := (- uf_10 #55) +#62 := (div #61 2::int) +#63 := (= #60 #62) +#64 := (not #63) +#58 := (= #56 #57) +#65 := (implies #58 #64) +#66 := (iff #53 #65) +#265 := (iff #66 #262) +#225 := (= #60 #222) +#228 := (not #225) +#234 := (not #58) +#235 := (or #234 #228) +#240 := (iff #53 #235) +#263 := (iff #240 #262) +#260 := (iff #235 #259) +#257 := (iff #228 #256) +#254 := (iff #225 #253) +#255 := [rewrite]: #254 +#258 := [monotonicity #255]: #257 +#249 := (iff #234 #248) +#246 := (iff #58 #245) +#247 := [rewrite]: #246 +#250 := [monotonicity #247]: #249 +#261 := [monotonicity #250 #258]: #260 +#264 := [monotonicity #261]: #263 +#241 := (iff #66 #240) +#238 := (iff #65 #235) +#231 := (implies #58 #228) +#236 := (iff #231 #235) +#237 := [rewrite]: #236 +#232 := (iff #65 #231) +#229 := (iff #64 #228) +#226 := (iff #63 #225) +#223 := (= #62 #222) +#220 := (= #61 #219) +#221 := [rewrite]: #220 +#224 := [monotonicity #221]: #223 +#227 := [monotonicity #224]: #226 +#230 := [monotonicity #227]: #229 +#233 := [monotonicity #230]: #232 +#239 := [trans #233 #237]: #238 +#242 := [monotonicity #239]: #241 +#266 := [trans #242 #264]: #265 +#217 := [asserted]: #66 +#267 := [mp #217 #266]: #262 +#455 := (not #262) +#765 := (or #558 #259 #455) +#439 := [def-axiom]: #765 +#772 := [unit-resolution #439 #267]: #440 +#490 := [unit-resolution #772 #508]: #558 +#483 := (or #53 #498 #529) +#484 := [th-lemma]: #483 +#487 := [unit-resolution #484 #490]: #485 +#486 := [unit-resolution #487 #497]: #498 +#678 := (<= #686 0::int) +#488 := (or #537 #678) +#477 := [th-lemma]: #488 +#478 := [unit-resolution #477 #536]: #678 +#479 := (or #256 #449) +#471 := [th-lemma]: #479 +#480 := [unit-resolution #471 #492]: #449 +#712 := (>= #711 0::int) +#481 := (or #566 #712) +#472 := [th-lemma]: #481 +#482 := [unit-resolution #472 #565]: #712 +#463 := [th-lemma #482 #480 #478 #486 #507]: false +#464 := [lemma #463]: #259 +#771 := (or #53 #381) +#434 := (or #53 #381 #455) +#769 := [def-axiom]: #434 +#428 := [unit-resolution #769 #267]: #771 +#442 := [unit-resolution #428 #464]: #53 +#435 := [monotonicity #442]: #468 +#437 := [symm #435]: #602 +#438 := (not #602) +#419 := (or #438 #709) +#420 := [th-lemma]: #419 +#421 := [unit-resolution #420 #437]: #709 +#422 := [th-lemma #421 #414 #499]: false +#423 := [lemma #422]: #448 +#410 := (or #245 #503) +#611 := (>= #561 0::int) +#682 := (or #438 #611) +#447 := [th-lemma]: #682 +#430 := [unit-resolution #447 #437]: #611 +#432 := [hypothesis]: #556 +#433 := [th-lemma #532 #432 #430]: false +#412 := [lemma #433]: #361 +#409 := (or #245 #556 #503) +#407 := [th-lemma]: #409 +#398 := [unit-resolution #407 #412]: #410 +#400 := [unit-resolution #398 #423]: #245 +#454 := (or #381 #248 #256) +#451 := [def-axiom]: #454 +#401 := [unit-resolution #451 #464]: #259 +#404 := [unit-resolution #401 #400]: #256 +#384 := (or #253 #399) +#429 := [hypothesis]: #573 +#443 := (or #558 #597) +#444 := [th-lemma]: #443 +#445 := [unit-resolution #444 #442]: #597 +#446 := [th-lemma #445 #507 #482 #429 #478]: false +#436 := [lemma #446]: #363 +#405 := (or #253 #399 #573) +#379 := [th-lemma]: #405 +#385 := [unit-resolution #379 #436]: #384 +#390 := [unit-resolution #385 #404]: #399 +#392 := (or #558 #629) +#393 := [th-lemma]: #392 +#395 := [unit-resolution #393 #442]: #629 +[th-lemma #395 #575 #570 #390 #528]: false +unsat diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_linarith_01 --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_linarith_01 Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,4 @@ +(benchmark Isabelle +:assumption (not (= 3 3)) +:formula true +) diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_linarith_01.proof --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_linarith_01.proof Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,17 @@ +#2 := false +#4 := 3::int +#5 := (= 3::int 3::int) +#6 := (not #5) +#30 := (iff #6 false) +#1 := true +#25 := (not true) +#28 := (iff #25 false) +#29 := [rewrite]: #28 +#26 := (iff #6 #25) +#23 := (iff #5 true) +#24 := [rewrite]: #23 +#27 := [monotonicity #24]: #26 +#31 := [trans #27 #29]: #30 +#22 := [asserted]: #6 +[mp #22 #31]: false +unsat diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_linarith_02 --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_linarith_02 Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,4 @@ +(benchmark Isabelle +:assumption (not (= 3.0 3.0)) +:formula true +) diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_linarith_02.proof --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_linarith_02.proof Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,17 @@ +#2 := false +#4 := 3::real +#5 := (= 3::real 3::real) +#6 := (not #5) +#30 := (iff #6 false) +#1 := true +#25 := (not true) +#28 := (iff #25 false) +#29 := [rewrite]: #28 +#26 := (iff #6 #25) +#23 := (iff #5 true) +#24 := [rewrite]: #23 +#27 := [monotonicity #24]: #26 +#31 := [trans #27 #29]: #30 +#22 := [asserted]: #6 +[mp #22 #31]: false +unsat diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_linarith_03 --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_linarith_03 Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,4 @@ +(benchmark Isabelle +:assumption (not (= (+ 3 1) 4)) +:formula true +) diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_linarith_03.proof --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_linarith_03.proof Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,26 @@ +#2 := false +#7 := 4::int +#5 := 1::int +#4 := 3::int +#6 := (+ 3::int 1::int) +#8 := (= #6 4::int) +#9 := (not #8) +#39 := (iff #9 false) +#1 := true +#34 := (not true) +#37 := (iff #34 false) +#38 := [rewrite]: #37 +#35 := (iff #9 #34) +#32 := (iff #8 true) +#27 := (= 4::int 4::int) +#30 := (iff #27 true) +#31 := [rewrite]: #30 +#28 := (iff #8 #27) +#26 := [rewrite]: #8 +#29 := [monotonicity #26]: #28 +#33 := [trans #29 #31]: #32 +#36 := [monotonicity #33]: #35 +#40 := [trans #36 #38]: #39 +#25 := [asserted]: #9 +[mp #25 #40]: false +unsat diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_linarith_04 --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_linarith_04 Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,9 @@ +(benchmark Isabelle +:extrafuns ( + (uf_1 Int) + (uf_2 Int) + (uf_3 Int) + ) +:assumption (not (= (+ uf_1 (+ uf_2 uf_3)) (+ uf_2 (+ uf_3 uf_1)))) +:formula true +) diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_linarith_04.proof --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_linarith_04.proof Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,41 @@ +#2 := false +decl uf_1 :: int +#4 := uf_1 +decl uf_3 :: int +#6 := uf_3 +#9 := (+ uf_3 uf_1) +decl uf_2 :: int +#5 := uf_2 +#10 := (+ uf_2 #9) +#7 := (+ uf_2 uf_3) +#8 := (+ uf_1 #7) +#11 := (= #8 #10) +#12 := (not #11) +#51 := (iff #12 false) +#1 := true +#46 := (not true) +#49 := (iff #46 false) +#50 := [rewrite]: #49 +#47 := (iff #12 #46) +#44 := (iff #11 true) +#39 := (= #8 #8) +#42 := (iff #39 true) +#43 := [rewrite]: #42 +#40 := (iff #11 #39) +#37 := (= #10 #8) +#29 := (+ uf_1 uf_3) +#32 := (+ uf_2 #29) +#35 := (= #32 #8) +#36 := [rewrite]: #35 +#33 := (= #10 #32) +#30 := (= #9 #29) +#31 := [rewrite]: #30 +#34 := [monotonicity #31]: #33 +#38 := [trans #34 #36]: #37 +#41 := [monotonicity #38]: #40 +#45 := [trans #41 #43]: #44 +#48 := [monotonicity #45]: #47 +#52 := [trans #48 #50]: #51 +#28 := [asserted]: #12 +[mp #28 #52]: false +unsat diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_linarith_05 --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_linarith_05 Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,4 @@ +(benchmark Isabelle +:assumption (not (< 5 (ite (<= 3 8) 8 3))) +:formula true +) diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_linarith_05.proof --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_linarith_05.proof Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,35 @@ +#2 := false +#5 := 3::int +#6 := 8::int +#7 := (<= 3::int 8::int) +#8 := (ite #7 8::int 3::int) +#4 := 5::int +#9 := (< 5::int #8) +#10 := (not #9) +#50 := (iff #10 false) +#1 := true +#45 := (not true) +#48 := (iff #45 false) +#49 := [rewrite]: #48 +#46 := (iff #10 #45) +#43 := (iff #9 true) +#38 := (< 5::int 8::int) +#41 := (iff #38 true) +#42 := [rewrite]: #41 +#39 := (iff #9 #38) +#36 := (= #8 8::int) +#31 := (ite true 8::int 3::int) +#34 := (= #31 8::int) +#35 := [rewrite]: #34 +#32 := (= #8 #31) +#29 := (iff #7 true) +#30 := [rewrite]: #29 +#33 := [monotonicity #30]: #32 +#37 := [trans #33 #35]: #36 +#40 := [monotonicity #37]: #39 +#44 := [trans #40 #42]: #43 +#47 := [monotonicity #44]: #46 +#51 := [trans #47 #49]: #50 +#26 := [asserted]: #10 +[mp #26 #51]: false +unsat diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_linarith_06 --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_linarith_06 Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,8 @@ +(benchmark Isabelle +:extrafuns ( + (uf_1 Real) + (uf_2 Real) + ) +:assumption (not (<= (ite (< (+ uf_1 uf_2) 0.0) (~ (+ uf_1 uf_2)) (+ uf_1 uf_2)) (+ (ite (< uf_1 0.0) (~ uf_1) uf_1) (ite (< uf_2 0.0) (~ uf_2) uf_2)))) +:formula true +) diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_linarith_06.proof --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_linarith_06.proof Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,250 @@ +#2 := false +#7 := 0::real +decl uf_2 :: real +#5 := uf_2 +#143 := 2::real +#144 := (* 2::real uf_2) +#165 := (<= #144 0::real) +#188 := (not #165) +#88 := (>= uf_2 0::real) +#166 := (or #88 #165) +#191 := (not #166) +decl uf_1 :: real +#4 := uf_1 +#76 := (>= uf_1 0::real) +#89 := (not #88) +#146 := (* 2::real uf_1) +#167 := (<= #146 0::real) +#199 := (not #167) +#263 := [hypothesis]: #88 +#147 := (+ #146 #144) +#168 := (<= #147 0::real) +#169 := (ite #88 #167 #168) +#194 := (not #169) +#186 := (or #166 #89) +#187 := [def-axiom]: #186 +#271 := [unit-resolution #187 #263]: #166 +#170 := (ite #76 #166 #169) +#205 := (not #170) +#6 := (+ uf_1 uf_2) +#64 := (>= #6 0::real) +#269 := (or #64 #89) +#65 := (not #64) +#262 := [hypothesis]: #65 +#174 := (>= #144 0::real) +#175 := (or #89 #174) +#230 := (not #175) +#257 := [hypothesis]: #230 +#225 := (or #175 #88) +#226 := [def-axiom]: #225 +#258 := [unit-resolution #226 #257]: #88 +#227 := (not #174) +#228 := (or #175 #227) +#229 := [def-axiom]: #228 +#259 := [unit-resolution #229 #257]: #227 +#260 := [th-lemma #259 #258]: false +#261 := [lemma #260]: #175 +#172 := (>= #146 0::real) +#171 := (>= #147 0::real) +#173 := (ite #88 #171 #172) +#176 := (ite #76 #173 #175) +#233 := (not #176) +#264 := (or #64 #233) +#177 := (ite #64 #170 #176) +#182 := (not #177) +#36 := -1::real +#38 := (* -1::real uf_2) +#95 := (ite #88 uf_2 #38) +#107 := (* -1::real #95) +#37 := (* -1::real uf_1) +#83 := (ite #76 uf_1 #37) +#106 := (* -1::real #83) +#108 := (+ #106 #107) +#39 := (+ #37 #38) +#71 := (ite #64 #6 #39) +#109 := (+ #71 #108) +#110 := (<= #109 0::real) +#115 := (not #110) +#183 := (iff #115 #182) +#180 := (iff #110 #177) +#150 := -2::real +#152 := (* -2::real uf_2) +#155 := (ite #88 #152 0::real) +#151 := (* -2::real uf_1) +#153 := (+ #151 #152) +#154 := (ite #88 #153 #151) +#156 := (ite #76 #154 #155) +#148 := (ite #88 #146 #147) +#145 := (ite #88 0::real #144) +#149 := (ite #76 #145 #148) +#157 := (ite #64 #149 #156) +#162 := (<= #157 0::real) +#178 := (iff #162 #177) +#179 := [rewrite]: #178 +#163 := (iff #110 #162) +#160 := (= #109 #157) +#133 := (+ uf_1 #38) +#134 := (ite #88 #133 #6) +#131 := (+ #37 uf_2) +#132 := (ite #88 #39 #131) +#135 := (ite #76 #132 #134) +#140 := (+ #71 #135) +#158 := (= #140 #157) +#159 := [rewrite]: #158 +#141 := (= #109 #140) +#138 := (= #108 #135) +#125 := (ite #88 #38 uf_2) +#123 := (ite #76 #37 uf_1) +#128 := (+ #123 #125) +#136 := (= #128 #135) +#137 := [rewrite]: #136 +#129 := (= #108 #128) +#126 := (= #107 #125) +#127 := [rewrite]: #126 +#121 := (= #106 #123) +#124 := [rewrite]: #121 +#130 := [monotonicity #124 #127]: #129 +#139 := [trans #130 #137]: #138 +#142 := [monotonicity #139]: #141 +#161 := [trans #142 #159]: #160 +#164 := [monotonicity #161]: #163 +#181 := [trans #164 #179]: #180 +#184 := [monotonicity #181]: #183 +#15 := (- uf_2) +#14 := (< uf_2 0::real) +#16 := (ite #14 #15 uf_2) +#12 := (- uf_1) +#11 := (< uf_1 0::real) +#13 := (ite #11 #12 uf_1) +#17 := (+ #13 #16) +#9 := (- #6) +#8 := (< #6 0::real) +#10 := (ite #8 #9 #6) +#18 := (<= #10 #17) +#19 := (not #18) +#118 := (iff #19 #115) +#52 := (ite #14 #38 uf_2) +#47 := (ite #11 #37 uf_1) +#55 := (+ #47 #52) +#42 := (ite #8 #39 #6) +#58 := (<= #42 #55) +#61 := (not #58) +#116 := (iff #61 #115) +#113 := (iff #58 #110) +#100 := (+ #83 #95) +#103 := (<= #71 #100) +#111 := (iff #103 #110) +#112 := [rewrite]: #111 +#104 := (iff #58 #103) +#101 := (= #55 #100) +#98 := (= #52 #95) +#92 := (ite #89 #38 uf_2) +#96 := (= #92 #95) +#97 := [rewrite]: #96 +#93 := (= #52 #92) +#90 := (iff #14 #89) +#91 := [rewrite]: #90 +#94 := [monotonicity #91]: #93 +#99 := [trans #94 #97]: #98 +#86 := (= #47 #83) +#77 := (not #76) +#80 := (ite #77 #37 uf_1) +#84 := (= #80 #83) +#85 := [rewrite]: #84 +#81 := (= #47 #80) +#78 := (iff #11 #77) +#79 := [rewrite]: #78 +#82 := [monotonicity #79]: #81 +#87 := [trans #82 #85]: #86 +#102 := [monotonicity #87 #99]: #101 +#74 := (= #42 #71) +#68 := (ite #65 #39 #6) +#72 := (= #68 #71) +#73 := [rewrite]: #72 +#69 := (= #42 #68) +#66 := (iff #8 #65) +#67 := [rewrite]: #66 +#70 := [monotonicity #67]: #69 +#75 := [trans #70 #73]: #74 +#105 := [monotonicity #75 #102]: #104 +#114 := [trans #105 #112]: #113 +#117 := [monotonicity #114]: #116 +#62 := (iff #19 #61) +#59 := (iff #18 #58) +#56 := (= #17 #55) +#53 := (= #16 #52) +#50 := (= #15 #38) +#51 := [rewrite]: #50 +#54 := [monotonicity #51]: #53 +#48 := (= #13 #47) +#45 := (= #12 #37) +#46 := [rewrite]: #45 +#49 := [monotonicity #46]: #48 +#57 := [monotonicity #49 #54]: #56 +#43 := (= #10 #42) +#40 := (= #9 #39) +#41 := [rewrite]: #40 +#44 := [monotonicity #41]: #43 +#60 := [monotonicity #44 #57]: #59 +#63 := [monotonicity #60]: #62 +#119 := [trans #63 #117]: #118 +#35 := [asserted]: #19 +#120 := [mp #35 #119]: #115 +#185 := [mp #120 #184]: #182 +#248 := (or #177 #64 #233) +#249 := [def-axiom]: #248 +#265 := [unit-resolution #249 #185]: #264 +#266 := [unit-resolution #265 #262]: #233 +#240 := (or #176 #76 #230) +#241 := [def-axiom]: #240 +#267 := [unit-resolution #241 #266 #261]: #76 +#268 := [th-lemma #267 #263 #262]: false +#270 := [lemma #268]: #269 +#272 := [unit-resolution #270 #263]: #64 +#273 := (or #65 #205) +#246 := (or #177 #65 #205) +#247 := [def-axiom]: #246 +#274 := [unit-resolution #247 #185]: #273 +#275 := [unit-resolution #274 #272]: #205 +#255 := (or #170 #194 #191) +#250 := [hypothesis]: #169 +#251 := [hypothesis]: #205 +#252 := [hypothesis]: #166 +#210 := (or #170 #77 #191) +#211 := [def-axiom]: #210 +#253 := [unit-resolution #211 #251 #252]: #77 +#212 := (or #170 #76 #194) +#213 := [def-axiom]: #212 +#254 := [unit-resolution #213 #253 #251 #250]: false +#256 := [lemma #254]: #255 +#276 := [unit-resolution #256 #275 #271]: #194 +#200 := (or #169 #89 #199) +#201 := [def-axiom]: #200 +#277 := [unit-resolution #201 #276 #263]: #199 +#278 := [unit-resolution #211 #275 #271]: #77 +#279 := [th-lemma #278 #277]: false +#280 := [lemma #279]: #89 +#281 := [hypothesis]: #77 +#282 := [unit-resolution #241 #281 #261]: #176 +#283 := [unit-resolution #265 #282]: #64 +#284 := [th-lemma #281 #283 #280]: false +#285 := [lemma #284]: #76 +#222 := (not #172) +#286 := [hypothesis]: #222 +#287 := [th-lemma #285 #286]: false +#288 := [lemma #287]: #172 +#223 := (or #173 #88 #222) +#224 := [def-axiom]: #223 +#289 := [unit-resolution #224 #288 #280]: #173 +#214 := (not #173) +#238 := (or #176 #77 #214) +#239 := [def-axiom]: #238 +#290 := [unit-resolution #239 #289 #285]: #176 +#291 := [unit-resolution #265 #290]: #64 +#292 := [unit-resolution #274 #291]: #205 +#293 := [unit-resolution #211 #292 #285]: #191 +#189 := (or #166 #188) +#190 := [def-axiom]: #189 +#294 := [unit-resolution #190 #293]: #188 +[th-lemma #280 #294]: false +unsat diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_linarith_07 --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_linarith_07 Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,11 @@ +(benchmark Isabelle +:extrasorts ( T2 T1) +:extrafuns ( + (uf_2 T1) + (uf_1 Int Int T1) + (uf_3 T1 T2) + ) +:assumption (forall (?x1 Int) (?x2 Int) (iff (= (uf_1 ?x1 ?x2) uf_2) (< ?x1 ?x2))) +:assumption (not (= (uf_3 (uf_1 2 3)) (uf_3 uf_2))) +:formula true +) diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_linarith_07.proof --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_linarith_07.proof Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,105 @@ +#2 := false +decl uf_3 :: (-> T1 T2) +decl uf_2 :: T1 +#7 := uf_2 +#16 := (uf_3 uf_2) +decl uf_1 :: (-> int int T1) +#13 := 3::int +#12 := 2::int +#14 := (uf_1 2::int 3::int) +#15 := (uf_3 #14) +#17 := (= #15 #16) +#516 := (= #16 #15) +#194 := (= uf_2 #14) +#5 := (:var 0 int) +#4 := (:var 1 int) +#6 := (uf_1 #4 #5) +#530 := (pattern #6) +#39 := 0::int +#37 := -1::int +#41 := (* -1::int #5) +#42 := (+ #4 #41) +#40 := (>= #42 0::int) +#38 := (not #40) +#8 := (= #6 uf_2) +#45 := (iff #8 #38) +#531 := (forall (vars (?x1 int) (?x2 int)) (:pat #530) #45) +#48 := (forall (vars (?x1 int) (?x2 int)) #45) +#534 := (iff #48 #531) +#532 := (iff #45 #45) +#533 := [refl]: #532 +#535 := [quant-intro #533]: #534 +#58 := (~ #48 #48) +#56 := (~ #45 #45) +#57 := [refl]: #56 +#59 := [nnf-pos #57]: #58 +#9 := (< #4 #5) +#10 := (iff #8 #9) +#11 := (forall (vars (?x1 int) (?x2 int)) #10) +#49 := (iff #11 #48) +#46 := (iff #10 #45) +#43 := (iff #9 #38) +#44 := [rewrite]: #43 +#47 := [monotonicity #44]: #46 +#50 := [quant-intro #47]: #49 +#34 := [asserted]: #11 +#51 := [mp #34 #50]: #48 +#60 := [mp~ #51 #59]: #48 +#536 := [mp #60 #535]: #531 +#508 := (not #531) +#509 := (or #508 #194) +#201 := (* -1::int 3::int) +#115 := (+ 2::int #201) +#202 := (>= #115 0::int) +#116 := (not #202) +#114 := (= #14 uf_2) +#203 := (iff #114 #116) +#510 := (or #508 #203) +#506 := (iff #510 #509) +#150 := (iff #509 #509) +#513 := [rewrite]: #150 +#171 := (iff #203 #194) +#1 := true +#164 := (iff #194 true) +#169 := (iff #164 #194) +#170 := [rewrite]: #169 +#505 := (iff #203 #164) +#180 := (iff #116 true) +#529 := (not false) +#184 := (iff #529 true) +#520 := [rewrite]: #184 +#519 := (iff #116 #529) +#528 := (iff #202 false) +#192 := (>= -1::int 0::int) +#526 := (iff #192 false) +#527 := [rewrite]: #526 +#193 := (iff #202 #192) +#311 := (= #115 -1::int) +#134 := -3::int +#208 := (+ 2::int -3::int) +#524 := (= #208 -1::int) +#181 := [rewrite]: #524 +#187 := (= #115 #208) +#207 := (= #201 -3::int) +#204 := [rewrite]: #207 +#522 := [monotonicity #204]: #187 +#518 := [trans #522 #181]: #311 +#525 := [monotonicity #518]: #193 +#523 := [trans #525 #527]: #528 +#179 := [monotonicity #523]: #519 +#521 := [trans #179 #520]: #180 +#205 := (iff #114 #194) +#206 := [rewrite]: #205 +#168 := [monotonicity #206 #521]: #505 +#507 := [trans #168 #170]: #171 +#512 := [monotonicity #507]: #506 +#515 := [trans #512 #513]: #506 +#511 := [quant-inst]: #510 +#155 := [mp #511 #515]: #509 +#156 := [unit-resolution #155 #536]: #194 +#514 := [monotonicity #156]: #516 +#517 := [symm #514]: #17 +#18 := (not #17) +#35 := [asserted]: #18 +[unit-resolution #35 #517]: false +unsat diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_linarith_08 --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_linarith_08 Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,7 @@ +(benchmark Isabelle +:extrafuns ( + (uf_1 Int) + ) +:assumption (not (or (<= 4 (+ uf_1 3)) (< uf_1 1))) +:formula true +) diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_linarith_08.proof --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_linarith_08.proof Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,54 @@ +#2 := false +#9 := 1::int +decl uf_1 :: int +#5 := uf_1 +#10 := (< uf_1 1::int) +#6 := 3::int +#7 := (+ uf_1 3::int) +#4 := 4::int +#8 := (<= 4::int #7) +#11 := (or #8 #10) +#12 := (not #11) +#66 := (iff #12 false) +#29 := (+ 3::int uf_1) +#32 := (<= 4::int #29) +#38 := (or #10 #32) +#43 := (not #38) +#64 := (iff #43 false) +#1 := true +#59 := (not true) +#62 := (iff #59 false) +#63 := [rewrite]: #62 +#60 := (iff #43 #59) +#57 := (iff #38 true) +#48 := (>= uf_1 1::int) +#46 := (not #48) +#52 := (or #46 #48) +#55 := (iff #52 true) +#56 := [rewrite]: #55 +#53 := (iff #38 #52) +#50 := (iff #32 #48) +#51 := [rewrite]: #50 +#47 := (iff #10 #46) +#49 := [rewrite]: #47 +#54 := [monotonicity #49 #51]: #53 +#58 := [trans #54 #56]: #57 +#61 := [monotonicity #58]: #60 +#65 := [trans #61 #63]: #64 +#44 := (iff #12 #43) +#41 := (iff #11 #38) +#35 := (or #32 #10) +#39 := (iff #35 #38) +#40 := [rewrite]: #39 +#36 := (iff #11 #35) +#33 := (iff #8 #32) +#30 := (= #7 #29) +#31 := [rewrite]: #30 +#34 := [monotonicity #31]: #33 +#37 := [monotonicity #34]: #36 +#42 := [trans #37 #40]: #41 +#45 := [monotonicity #42]: #44 +#67 := [trans #45 #65]: #66 +#28 := [asserted]: #12 +[mp #28 #67]: false +unsat diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_linarith_09 --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_linarith_09 Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,10 @@ +(benchmark Isabelle +:extrafuns ( + (uf_1 Int) + (uf_2 Int) + ) +:assumption (<= 3 uf_1) +:assumption (= uf_2 (+ uf_1 4)) +:assumption (not (< 0 (- uf_2 uf_1))) +:formula true +) diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_linarith_09.proof --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_linarith_09.proof Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,63 @@ +#2 := false +#11 := 0::int +decl uf_2 :: int +#7 := uf_2 +#42 := -1::int +#45 := (* -1::int uf_2) +decl uf_1 :: int +#5 := uf_1 +#46 := (+ uf_1 #45) +#63 := (>= #46 0::int) +#83 := (iff #63 false) +#44 := -4::int +#79 := (>= -4::int 0::int) +#81 := (iff #79 false) +#82 := [rewrite]: #81 +#77 := (iff #63 #79) +#47 := (= #46 -4::int) +#8 := 4::int +#9 := (+ uf_1 4::int) +#10 := (= uf_2 #9) +#49 := (iff #10 #47) +#32 := (+ 4::int uf_1) +#39 := (= uf_2 #32) +#43 := (iff #39 #47) +#48 := [rewrite]: #43 +#40 := (iff #10 #39) +#37 := (= #9 #32) +#38 := [rewrite]: #37 +#41 := [monotonicity #38]: #40 +#50 := [trans #41 #48]: #49 +#31 := [asserted]: #10 +#51 := [mp #31 #50]: #47 +#80 := [monotonicity #51]: #77 +#84 := [trans #80 #82]: #83 +#12 := (- uf_2 uf_1) +#13 := (< 0::int #12) +#14 := (not #13) +#74 := (iff #14 #63) +#53 := (* -1::int uf_1) +#54 := (+ #53 uf_2) +#57 := (< 0::int #54) +#60 := (not #57) +#72 := (iff #60 #63) +#64 := (not #63) +#67 := (not #64) +#70 := (iff #67 #63) +#71 := [rewrite]: #70 +#68 := (iff #60 #67) +#65 := (iff #57 #64) +#66 := [rewrite]: #65 +#69 := [monotonicity #66]: #68 +#73 := [trans #69 #71]: #72 +#61 := (iff #14 #60) +#58 := (iff #13 #57) +#55 := (= #12 #54) +#56 := [rewrite]: #55 +#59 := [monotonicity #56]: #58 +#62 := [monotonicity #59]: #61 +#75 := [trans #62 #73]: #74 +#52 := [asserted]: #14 +#76 := [mp #52 #75]: #63 +[mp #76 #84]: false +unsat diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_linarith_10 --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_linarith_10 Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,4 @@ +(benchmark Isabelle +:assumption (not (let (?x1 2) (not (= (+ ?x1 ?x1) 5)))) +:formula true +) diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_linarith_10.proof --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_linarith_10.proof Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,35 @@ +#2 := false +#6 := 5::int +#4 := 2::int +#5 := (+ 2::int 2::int) +#7 := (= #5 5::int) +#8 := (not #7) +#9 := (not #8) +#48 := (iff #9 false) +#1 := true +#43 := (not true) +#46 := (iff #43 false) +#47 := [rewrite]: #46 +#44 := (iff #9 #43) +#41 := (iff #8 true) +#36 := (not false) +#39 := (iff #36 true) +#40 := [rewrite]: #39 +#37 := (iff #8 #36) +#34 := (iff #7 false) +#26 := 4::int +#29 := (= 4::int 5::int) +#32 := (iff #29 false) +#33 := [rewrite]: #32 +#30 := (iff #7 #29) +#27 := (= #5 4::int) +#28 := [rewrite]: #27 +#31 := [monotonicity #28]: #30 +#35 := [trans #31 #33]: #34 +#38 := [monotonicity #35]: #37 +#42 := [trans #38 #40]: #41 +#45 := [monotonicity #42]: #44 +#49 := [trans #45 #47]: #48 +#25 := [asserted]: #9 +[mp #25 #49]: false +unsat diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_linarith_11 --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_linarith_11 Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,10 @@ +(benchmark Isabelle +:extrafuns ( + (uf_2 Real) + (uf_1 Real) + ) +:assumption (< (+ (* 3.0 uf_1) (* 7.0 uf_2)) 4.0) +:assumption (< 3.0 (* 2.0 uf_1)) +:assumption (not (< uf_2 0.0)) +:formula true +) diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_linarith_11.proof --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_linarith_11.proof Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,45 @@ +#2 := false +#11 := 4::real +decl uf_2 :: real +#8 := uf_2 +#7 := 7::real +#9 := (* 7::real uf_2) +decl uf_1 :: real +#5 := uf_1 +#4 := 3::real +#6 := (* 3::real uf_1) +#10 := (+ #6 #9) +#41 := (>= #10 4::real) +#39 := (not #41) +#12 := (< #10 4::real) +#40 := (iff #12 #39) +#37 := [rewrite]: #40 +#34 := [asserted]: #12 +#38 := [mp #34 #37]: #39 +#13 := 2::real +#14 := (* 2::real uf_1) +#43 := (<= #14 3::real) +#44 := (not #43) +#15 := (< 3::real #14) +#45 := (iff #15 #44) +#46 := [rewrite]: #45 +#35 := [asserted]: #15 +#47 := [mp #35 #46]: #44 +#16 := 0::real +#51 := (>= uf_2 0::real) +#17 := (< uf_2 0::real) +#18 := (not #17) +#58 := (iff #18 #51) +#49 := (not #51) +#53 := (not #49) +#56 := (iff #53 #51) +#57 := [rewrite]: #56 +#54 := (iff #18 #53) +#50 := (iff #17 #49) +#52 := [rewrite]: #50 +#55 := [monotonicity #52]: #54 +#59 := [trans #55 #57]: #58 +#36 := [asserted]: #18 +#60 := [mp #36 #59]: #51 +[th-lemma #60 #47 #38]: false +unsat diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_linarith_12 --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_linarith_12 Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,8 @@ +(benchmark Isabelle +:extrafuns ( + (uf_2 Int) + (uf_1 Int) + ) +:assumption (not (iff (or (<= 0 (+ uf_1 (* (~ 1) uf_2))) (or (not (<= 0 uf_2)) (<= 0 uf_2))) (not false))) +:formula true +) diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_linarith_12.proof --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_linarith_12.proof Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,59 @@ +#2 := false +#16 := (not false) +decl uf_2 :: int +#8 := uf_2 +#4 := 0::int +#12 := (<= 0::int uf_2) +#13 := (not #12) +#14 := (or #13 #12) +#6 := 1::int +#7 := (- 1::int) +#9 := (* #7 uf_2) +decl uf_1 :: int +#5 := uf_1 +#10 := (+ uf_1 #9) +#11 := (<= 0::int #10) +#15 := (or #11 #14) +#17 := (iff #15 #16) +#18 := (not #17) +#70 := (iff #18 false) +#1 := true +#65 := (not true) +#68 := (iff #65 false) +#69 := [rewrite]: #68 +#66 := (iff #18 #65) +#63 := (iff #17 true) +#58 := (iff true true) +#61 := (iff #58 true) +#62 := [rewrite]: #61 +#59 := (iff #17 #58) +#56 := (iff #16 true) +#57 := [rewrite]: #56 +#54 := (iff #15 true) +#35 := -1::int +#38 := (* -1::int uf_2) +#41 := (+ uf_1 #38) +#44 := (<= 0::int #41) +#49 := (or #44 true) +#52 := (iff #49 true) +#53 := [rewrite]: #52 +#50 := (iff #15 #49) +#47 := (iff #14 true) +#48 := [rewrite]: #47 +#45 := (iff #11 #44) +#42 := (= #10 #41) +#39 := (= #9 #38) +#36 := (= #7 -1::int) +#37 := [rewrite]: #36 +#40 := [monotonicity #37]: #39 +#43 := [monotonicity #40]: #42 +#46 := [monotonicity #43]: #45 +#51 := [monotonicity #46 #48]: #50 +#55 := [trans #51 #53]: #54 +#60 := [monotonicity #55 #57]: #59 +#64 := [trans #60 #62]: #63 +#67 := [monotonicity #64]: #66 +#71 := [trans #67 #69]: #70 +#34 := [asserted]: #18 +[mp #34 #71]: false +unsat diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_linarith_13 --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_linarith_13 Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,13 @@ +(benchmark Isabelle +:extrasorts ( T1) +:extrafuns ( + (uf_2 T1) + (uf_3 Int Int T1) + (uf_1 Int Int T1) + (uf_4 Int) + ) +:assumption (forall (?x1 Int) (?x2 Int) (iff (= (uf_1 ?x1 ?x2) uf_2) (<= ?x1 ?x2))) +:assumption (forall (?x3 Int) (?x4 Int) (iff (= (uf_3 ?x3 ?x4) uf_2) (< ?x3 ?x4))) +:assumption (not (distinct (uf_3 uf_4 3) (uf_1 3 uf_4))) +:formula true +) diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_linarith_13.proof --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_linarith_13.proof Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,212 @@ +#2 := false +decl uf_3 :: (-> int int T1) +#18 := 3::int +decl uf_4 :: int +#17 := uf_4 +#19 := (uf_3 uf_4 3::int) +decl uf_2 :: T1 +#7 := uf_2 +#221 := (= uf_2 #19) +decl uf_1 :: (-> int int T1) +#20 := (uf_1 3::int uf_4) +#256 := (= uf_2 #20) +#531 := (iff #256 #221) +#529 := (iff #221 #256) +#87 := (= #19 #20) +#21 := (distinct #19 #20) +#22 := (not #21) +#96 := (iff #22 #87) +#88 := (not #87) +#91 := (not #88) +#94 := (iff #91 #87) +#95 := [rewrite]: #94 +#92 := (iff #22 #91) +#89 := (iff #21 #88) +#90 := [rewrite]: #89 +#93 := [monotonicity #90]: #92 +#97 := [trans #93 #95]: #96 +#86 := [asserted]: #22 +#100 := [mp #86 #97]: #87 +#530 := [monotonicity #100]: #529 +#525 := [symm #530]: #531 +#548 := (not #221) +#232 := (not #256) +#526 := (iff #232 #548) +#532 := [monotonicity #525]: #526 +#536 := [hypothesis]: #232 +#533 := [mp #536 #532]: #548 +#259 := (>= uf_4 3::int) +#576 := (not #259) +#542 := (or #256 #576) +#257 := (iff #256 #259) +#5 := (:var 0 int) +#4 := (:var 1 int) +#6 := (uf_1 #4 #5) +#583 := (pattern #6) +#44 := 0::int +#41 := -1::int +#42 := (* -1::int #5) +#43 := (+ #4 #42) +#45 := (<= #43 0::int) +#8 := (= #6 uf_2) +#48 := (iff #8 #45) +#584 := (forall (vars (?x1 int) (?x2 int)) (:pat #583) #48) +#51 := (forall (vars (?x1 int) (?x2 int)) #48) +#587 := (iff #51 #584) +#585 := (iff #48 #48) +#586 := [refl]: #585 +#588 := [quant-intro #586]: #587 +#108 := (~ #51 #51) +#106 := (~ #48 #48) +#107 := [refl]: #106 +#109 := [nnf-pos #107]: #108 +#9 := (<= #4 #5) +#10 := (iff #8 #9) +#11 := (forall (vars (?x1 int) (?x2 int)) #10) +#52 := (iff #11 #51) +#49 := (iff #10 #48) +#46 := (iff #9 #45) +#47 := [rewrite]: #46 +#50 := [monotonicity #47]: #49 +#53 := [quant-intro #50]: #52 +#38 := [asserted]: #11 +#54 := [mp #38 #53]: #51 +#110 := [mp~ #54 #109]: #51 +#589 := [mp #110 #588]: #584 +#575 := (not #584) +#577 := (or #575 #257) +#167 := (* -1::int uf_4) +#254 := (+ 3::int #167) +#168 := (<= #254 0::int) +#255 := (= #20 uf_2) +#169 := (iff #255 #168) +#234 := (or #575 #169) +#571 := (iff #234 #577) +#246 := (iff #577 #577) +#578 := [rewrite]: #246 +#261 := (iff #169 #257) +#187 := (iff #168 #259) +#260 := [rewrite]: #187 +#247 := (iff #255 #256) +#258 := [rewrite]: #247 +#240 := [monotonicity #258 #260]: #261 +#245 := [monotonicity #240]: #571 +#579 := [trans #245 #578]: #571 +#364 := [quant-inst]: #234 +#580 := [mp #364 #579]: #577 +#541 := [unit-resolution #580 #589]: #257 +#581 := (not #257) +#582 := (or #581 #256 #576) +#572 := [def-axiom]: #582 +#537 := [unit-resolution #572 #541]: #542 +#543 := [unit-resolution #537 #536]: #576 +#385 := (or #221 #259) +#552 := (iff #221 #576) +#12 := (uf_3 #4 #5) +#590 := (pattern #12) +#69 := (>= #43 0::int) +#68 := (not #69) +#40 := (= uf_2 #12) +#75 := (iff #40 #68) +#591 := (forall (vars (?x3 int) (?x4 int)) (:pat #590) #75) +#80 := (forall (vars (?x3 int) (?x4 int)) #75) +#594 := (iff #80 #591) +#592 := (iff #75 #75) +#593 := [refl]: #592 +#595 := [quant-intro #593]: #594 +#101 := (~ #80 #80) +#111 := (~ #75 #75) +#112 := [refl]: #111 +#98 := [nnf-pos #112]: #101 +#14 := (< #4 #5) +#13 := (= #12 uf_2) +#15 := (iff #13 #14) +#16 := (forall (vars (?x3 int) (?x4 int)) #15) +#83 := (iff #16 #80) +#60 := (iff #14 #40) +#65 := (forall (vars (?x3 int) (?x4 int)) #60) +#81 := (iff #65 #80) +#78 := (iff #60 #75) +#72 := (iff #68 #40) +#76 := (iff #72 #75) +#77 := [rewrite]: #76 +#73 := (iff #60 #72) +#70 := (iff #14 #68) +#71 := [rewrite]: #70 +#74 := [monotonicity #71]: #73 +#79 := [trans #74 #77]: #78 +#82 := [quant-intro #79]: #81 +#66 := (iff #16 #65) +#63 := (iff #15 #60) +#57 := (iff #40 #14) +#61 := (iff #57 #60) +#62 := [rewrite]: #61 +#58 := (iff #15 #57) +#55 := (iff #13 #40) +#56 := [rewrite]: #55 +#59 := [monotonicity #56]: #58 +#64 := [trans #59 #62]: #63 +#67 := [quant-intro #64]: #66 +#84 := [trans #67 #82]: #83 +#39 := [asserted]: #16 +#85 := [mp #39 #84]: #80 +#113 := [mp~ #85 #98]: #80 +#596 := [mp #113 #595]: #591 +#276 := (not #591) +#550 := (or #276 #552) +#222 := (* -1::int 3::int) +#223 := (+ uf_4 #222) +#224 := (>= #223 0::int) +#560 := (not #224) +#561 := (iff #221 #560) +#554 := (or #276 #561) +#555 := (iff #554 #550) +#266 := (iff #550 #550) +#267 := [rewrite]: #266 +#553 := (iff #561 #552) +#282 := (iff #560 #576) +#280 := (iff #224 #259) +#562 := -3::int +#566 := (+ -3::int uf_4) +#567 := (>= #566 0::int) +#557 := (iff #567 #259) +#279 := [rewrite]: #557 +#570 := (iff #224 #567) +#209 := (= #223 #566) +#559 := (+ uf_4 -3::int) +#568 := (= #559 #566) +#208 := [rewrite]: #568 +#565 := (= #223 #559) +#563 := (= #222 -3::int) +#564 := [rewrite]: #563 +#203 := [monotonicity #564]: #565 +#569 := [trans #203 #208]: #209 +#556 := [monotonicity #569]: #570 +#281 := [trans #556 #279]: #280 +#175 := [monotonicity #281]: #282 +#275 := [monotonicity #175]: #553 +#265 := [monotonicity #275]: #555 +#268 := [trans #265 #267]: #555 +#551 := [quant-inst]: #554 +#546 := [mp #551 #268]: #550 +#384 := [unit-resolution #546 #596]: #552 +#547 := (not #552) +#262 := (or #547 #221 #259) +#544 := [def-axiom]: #262 +#386 := [unit-resolution #544 #384]: #385 +#528 := [unit-resolution #386 #543]: #221 +#527 := [unit-resolution #528 #533]: false +#534 := [lemma #527]: #256 +#523 := [mp #534 #525]: #221 +#363 := (or #232 #259) +#237 := (or #581 #232 #259) +#573 := [def-axiom]: #237 +#365 := [unit-resolution #573 #541]: #363 +#366 := [unit-resolution #365 #534]: #259 +#519 := (or #548 #576) +#545 := (or #547 #548 #576) +#549 := [def-axiom]: #545 +#520 := [unit-resolution #549 #384]: #519 +#522 := [unit-resolution #520 #366]: #548 +[unit-resolution #522 #523]: false +unsat diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_linarith_14 --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_linarith_14 Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,8 @@ +(benchmark Isabelle +:extrafuns ( + (uf_1 Int) + ) +:assumption (< 0 uf_1) +:assumption (not (distinct uf_1 (* uf_1 2) (- uf_1 uf_1))) +:formula true +) diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_linarith_14.proof --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_linarith_14.proof Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,86 @@ +#2 := false +decl uf_1 :: int +#5 := uf_1 +#7 := 2::int +#29 := (* 2::int uf_1) +#4 := 0::int +#54 := (= 0::int #29) +#55 := (not #54) +#61 := (= #29 0::int) +#104 := (not #61) +#110 := (iff #104 #55) +#108 := (iff #61 #54) +#109 := [commutativity]: #108 +#111 := [monotonicity #109]: #110 +#62 := (<= #29 0::int) +#100 := (not #62) +#30 := (<= uf_1 0::int) +#31 := (not #30) +#6 := (< 0::int uf_1) +#32 := (iff #6 #31) +#33 := [rewrite]: #32 +#27 := [asserted]: #6 +#34 := [mp #27 #33]: #31 +#101 := (or #100 #30) +#102 := [th-lemma]: #101 +#103 := [unit-resolution #102 #34]: #100 +#105 := (or #104 #62) +#106 := [th-lemma]: #105 +#107 := [unit-resolution #106 #103]: #104 +#112 := [mp #107 #111]: #55 +#56 := (= uf_1 #29) +#57 := (not #56) +#53 := (= 0::int uf_1) +#50 := (not #53) +#58 := (and #50 #55 #57) +#69 := (not #58) +#42 := (distinct 0::int uf_1 #29) +#47 := (not #42) +#9 := (- uf_1 uf_1) +#8 := (* uf_1 2::int) +#10 := (distinct uf_1 #8 #9) +#11 := (not #10) +#48 := (iff #11 #47) +#45 := (iff #10 #42) +#39 := (distinct uf_1 #29 0::int) +#43 := (iff #39 #42) +#44 := [rewrite]: #43 +#40 := (iff #10 #39) +#37 := (= #9 0::int) +#38 := [rewrite]: #37 +#35 := (= #8 #29) +#36 := [rewrite]: #35 +#41 := [monotonicity #36 #38]: #40 +#46 := [trans #41 #44]: #45 +#49 := [monotonicity #46]: #48 +#28 := [asserted]: #11 +#52 := [mp #28 #49]: #47 +#80 := (or #42 #69) +#81 := [def-axiom]: #80 +#82 := [unit-resolution #81 #52]: #69 +#59 := (= uf_1 0::int) +#83 := (not #59) +#89 := (iff #83 #50) +#87 := (iff #59 #53) +#88 := [commutativity]: #87 +#90 := [monotonicity #88]: #89 +#84 := (or #83 #30) +#85 := [th-lemma]: #84 +#86 := [unit-resolution #85 #34]: #83 +#91 := [mp #86 #90]: #50 +#64 := -1::int +#65 := (* -1::int #29) +#66 := (+ uf_1 #65) +#68 := (>= #66 0::int) +#92 := (not #68) +#93 := (or #92 #30) +#94 := [th-lemma]: #93 +#95 := [unit-resolution #94 #34]: #92 +#96 := (or #57 #68) +#97 := [th-lemma]: #96 +#98 := [unit-resolution #97 #95]: #57 +#76 := (or #58 #53 #54 #56) +#77 := [def-axiom]: #76 +#99 := [unit-resolution #77 #98 #91 #82]: #54 +[unit-resolution #99 #112]: false +unsat diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_linarith_15 --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_linarith_15 Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,9 @@ +(benchmark Isabelle +:extrafuns ( + (uf_2 Int) + (uf_1 Int) + (uf_3 Int) + ) +:assumption (not (or (and (< uf_1 uf_2) (< uf_2 uf_3)) (or (and (< uf_1 uf_2) (= uf_2 uf_3)) (or (and (< uf_1 uf_3) (< uf_3 uf_2)) (or (and (= uf_1 uf_3) (< uf_3 uf_2)) (or (and (= uf_1 uf_2) (< uf_2 uf_3)) (or (and (< uf_3 uf_2) (< uf_2 uf_1)) (or (and (< uf_3 uf_2) (= uf_2 uf_1)) (or (and (< uf_3 uf_1) (< uf_1 uf_2)) (or (and (= uf_3 uf_1) (< uf_1 uf_2)) (or (and (= uf_3 uf_2) (< uf_2 uf_1)) (or (and (< uf_2 uf_1) (< uf_1 uf_3)) (or (and (< uf_2 uf_1) (= uf_3 uf_1)) (or (and (< uf_2 uf_3) (< uf_3 uf_1)) (or (and (= uf_2 uf_1) (< uf_1 uf_3)) (or (and (= uf_2 uf_3) (< uf_3 uf_1)) (and (= uf_3 uf_2) (= uf_2 uf_1)))))))))))))))))) +:formula true +) diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_linarith_15.proof --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_linarith_15.proof Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,673 @@ +#2 := false +#169 := 0::int +decl uf_2 :: int +#5 := uf_2 +#166 := -1::int +#202 := (* -1::int uf_2) +decl uf_1 :: int +#4 := uf_1 +#203 := (+ uf_1 #202) +#218 := (>= #203 0::int) +decl uf_3 :: int +#7 := uf_3 +#167 := (* -1::int uf_3) +#168 := (+ uf_1 #167) +#178 := (>= #168 0::int) +#217 := (not #218) +#204 := (<= #203 0::int) +#205 := (not #204) +#692 := [hypothesis]: #205 +#177 := (not #178) +#693 := (or #177 #204) +#170 := (<= #168 0::int) +#191 := (+ uf_2 #167) +#237 := (<= #191 0::int) +#238 := (not #237) +#171 := (not #170) +#685 := [hypothesis]: #171 +#190 := (>= #191 0::int) +#455 := (or #170 #190) +#189 := (not #190) +#197 := (and #171 #189) +#354 := (not #197) +#464 := (iff #354 #455) +#456 := (not #455) +#459 := (not #456) +#462 := (iff #459 #455) +#463 := [rewrite]: #462 +#460 := (iff #354 #459) +#457 := (iff #197 #456) +#458 := [rewrite]: #457 +#461 := [monotonicity #458]: #460 +#465 := [trans #461 #463]: #464 +#287 := (and #189 #217) +#10 := (= uf_2 uf_3) +#279 := (and #10 #217) +#273 := (and #177 #238) +#15 := (= uf_1 uf_3) +#268 := (and #15 #238) +#17 := (= uf_1 uf_2) +#260 := (and #17 #189) +#252 := (and #205 #238) +#244 := (and #17 #238) +#232 := (and #171 #217) +#224 := (and #15 #217) +#214 := (and #10 #205) +#211 := (and #177 #205) +#208 := (and #15 #205) +#184 := (and #17 #177) +#174 := (and #10 #171) +#115 := (and #10 #17) +#337 := (or #115 #174 #184 #197 #208 #211 #214 #224 #232 #244 #252 #260 #268 #273 #279 #287) +#342 := (not #337) +#21 := (= uf_2 uf_1) +#27 := (= uf_3 uf_2) +#34 := (and #27 #21) +#23 := (< uf_3 uf_1) +#33 := (and #10 #23) +#35 := (or #33 #34) +#12 := (< uf_1 uf_3) +#32 := (and #21 #12) +#36 := (or #32 #35) +#8 := (< uf_2 uf_3) +#31 := (and #8 #23) +#37 := (or #31 #36) +#25 := (= uf_3 uf_1) +#19 := (< uf_2 uf_1) +#30 := (and #19 #25) +#38 := (or #30 #37) +#29 := (and #19 #12) +#39 := (or #29 #38) +#28 := (and #27 #19) +#40 := (or #28 #39) +#6 := (< uf_1 uf_2) +#26 := (and #25 #6) +#41 := (or #26 #40) +#24 := (and #23 #6) +#42 := (or #24 #41) +#13 := (< uf_3 uf_2) +#22 := (and #13 #21) +#43 := (or #22 #42) +#20 := (and #13 #19) +#44 := (or #20 #43) +#18 := (and #17 #8) +#45 := (or #18 #44) +#16 := (and #15 #13) +#46 := (or #16 #45) +#14 := (and #12 #13) +#47 := (or #14 #46) +#11 := (and #6 #10) +#48 := (or #11 #47) +#9 := (and #6 #8) +#49 := (or #9 #48) +#50 := (not #49) +#345 := (iff #50 #342) +#118 := (or #33 #115) +#110 := (and #12 #17) +#121 := (or #110 #118) +#124 := (or #31 #121) +#102 := (and #15 #19) +#127 := (or #102 #124) +#96 := (and #12 #19) +#130 := (or #96 #127) +#93 := (and #10 #19) +#133 := (or #93 #130) +#86 := (and #6 #15) +#136 := (or #86 #133) +#78 := (and #6 #23) +#139 := (or #78 #136) +#75 := (and #13 #17) +#142 := (or #75 #139) +#145 := (or #20 #142) +#70 := (and #8 #17) +#148 := (or #70 #145) +#67 := (and #13 #15) +#151 := (or #67 #148) +#154 := (or #14 #151) +#157 := (or #11 #154) +#160 := (or #9 #157) +#163 := (not #160) +#343 := (iff #163 #342) +#340 := (iff #160 #337) +#292 := (or #174 #115) +#295 := (or #184 #292) +#298 := (or #197 #295) +#301 := (or #208 #298) +#304 := (or #211 #301) +#307 := (or #214 #304) +#310 := (or #224 #307) +#313 := (or #232 #310) +#316 := (or #244 #313) +#319 := (or #252 #316) +#322 := (or #260 #319) +#325 := (or #268 #322) +#328 := (or #273 #325) +#331 := (or #279 #328) +#334 := (or #287 #331) +#338 := (iff #334 #337) +#339 := [rewrite]: #338 +#335 := (iff #160 #334) +#332 := (iff #157 #331) +#329 := (iff #154 #328) +#326 := (iff #151 #325) +#323 := (iff #148 #322) +#320 := (iff #145 #319) +#317 := (iff #142 #316) +#314 := (iff #139 #313) +#311 := (iff #136 #310) +#308 := (iff #133 #307) +#305 := (iff #130 #304) +#302 := (iff #127 #301) +#299 := (iff #124 #298) +#296 := (iff #121 #295) +#293 := (iff #118 #292) +#175 := (iff #33 #174) +#172 := (iff #23 #171) +#173 := [rewrite]: #172 +#176 := [monotonicity #173]: #175 +#294 := [monotonicity #176]: #293 +#187 := (iff #110 #184) +#181 := (and #177 #17) +#185 := (iff #181 #184) +#186 := [rewrite]: #185 +#182 := (iff #110 #181) +#179 := (iff #12 #177) +#180 := [rewrite]: #179 +#183 := [monotonicity #180]: #182 +#188 := [trans #183 #186]: #187 +#297 := [monotonicity #188 #294]: #296 +#200 := (iff #31 #197) +#194 := (and #189 #171) +#198 := (iff #194 #197) +#199 := [rewrite]: #198 +#195 := (iff #31 #194) +#192 := (iff #8 #189) +#193 := [rewrite]: #192 +#196 := [monotonicity #193 #173]: #195 +#201 := [trans #196 #199]: #200 +#300 := [monotonicity #201 #297]: #299 +#209 := (iff #102 #208) +#206 := (iff #19 #205) +#207 := [rewrite]: #206 +#210 := [monotonicity #207]: #209 +#303 := [monotonicity #210 #300]: #302 +#212 := (iff #96 #211) +#213 := [monotonicity #180 #207]: #212 +#306 := [monotonicity #213 #303]: #305 +#215 := (iff #93 #214) +#216 := [monotonicity #207]: #215 +#309 := [monotonicity #216 #306]: #308 +#227 := (iff #86 #224) +#221 := (and #217 #15) +#225 := (iff #221 #224) +#226 := [rewrite]: #225 +#222 := (iff #86 #221) +#219 := (iff #6 #217) +#220 := [rewrite]: #219 +#223 := [monotonicity #220]: #222 +#228 := [trans #223 #226]: #227 +#312 := [monotonicity #228 #309]: #311 +#235 := (iff #78 #232) +#229 := (and #217 #171) +#233 := (iff #229 #232) +#234 := [rewrite]: #233 +#230 := (iff #78 #229) +#231 := [monotonicity #220 #173]: #230 +#236 := [trans #231 #234]: #235 +#315 := [monotonicity #236 #312]: #314 +#247 := (iff #75 #244) +#241 := (and #238 #17) +#245 := (iff #241 #244) +#246 := [rewrite]: #245 +#242 := (iff #75 #241) +#239 := (iff #13 #238) +#240 := [rewrite]: #239 +#243 := [monotonicity #240]: #242 +#248 := [trans #243 #246]: #247 +#318 := [monotonicity #248 #315]: #317 +#255 := (iff #20 #252) +#249 := (and #238 #205) +#253 := (iff #249 #252) +#254 := [rewrite]: #253 +#250 := (iff #20 #249) +#251 := [monotonicity #240 #207]: #250 +#256 := [trans #251 #254]: #255 +#321 := [monotonicity #256 #318]: #320 +#263 := (iff #70 #260) +#257 := (and #189 #17) +#261 := (iff #257 #260) +#262 := [rewrite]: #261 +#258 := (iff #70 #257) +#259 := [monotonicity #193]: #258 +#264 := [trans #259 #262]: #263 +#324 := [monotonicity #264 #321]: #323 +#271 := (iff #67 #268) +#265 := (and #238 #15) +#269 := (iff #265 #268) +#270 := [rewrite]: #269 +#266 := (iff #67 #265) +#267 := [monotonicity #240]: #266 +#272 := [trans #267 #270]: #271 +#327 := [monotonicity #272 #324]: #326 +#274 := (iff #14 #273) +#275 := [monotonicity #180 #240]: #274 +#330 := [monotonicity #275 #327]: #329 +#282 := (iff #11 #279) +#276 := (and #217 #10) +#280 := (iff #276 #279) +#281 := [rewrite]: #280 +#277 := (iff #11 #276) +#278 := [monotonicity #220]: #277 +#283 := [trans #278 #281]: #282 +#333 := [monotonicity #283 #330]: #332 +#290 := (iff #9 #287) +#284 := (and #217 #189) +#288 := (iff #284 #287) +#289 := [rewrite]: #288 +#285 := (iff #9 #284) +#286 := [monotonicity #220 #193]: #285 +#291 := [trans #286 #289]: #290 +#336 := [monotonicity #291 #333]: #335 +#341 := [trans #336 #339]: #340 +#344 := [monotonicity #341]: #343 +#164 := (iff #50 #163) +#161 := (iff #49 #160) +#158 := (iff #48 #157) +#155 := (iff #47 #154) +#152 := (iff #46 #151) +#149 := (iff #45 #148) +#146 := (iff #44 #145) +#143 := (iff #43 #142) +#140 := (iff #42 #139) +#137 := (iff #41 #136) +#134 := (iff #40 #133) +#131 := (iff #39 #130) +#128 := (iff #38 #127) +#125 := (iff #37 #124) +#122 := (iff #36 #121) +#119 := (iff #35 #118) +#116 := (iff #34 #115) +#73 := (iff #21 #17) +#74 := [rewrite]: #73 +#91 := (iff #27 #10) +#92 := [rewrite]: #91 +#117 := [monotonicity #92 #74]: #116 +#120 := [monotonicity #117]: #119 +#113 := (iff #32 #110) +#107 := (and #17 #12) +#111 := (iff #107 #110) +#112 := [rewrite]: #111 +#108 := (iff #32 #107) +#109 := [monotonicity #74]: #108 +#114 := [trans #109 #112]: #113 +#123 := [monotonicity #114 #120]: #122 +#126 := [monotonicity #123]: #125 +#105 := (iff #30 #102) +#99 := (and #19 #15) +#103 := (iff #99 #102) +#104 := [rewrite]: #103 +#100 := (iff #30 #99) +#81 := (iff #25 #15) +#82 := [rewrite]: #81 +#101 := [monotonicity #82]: #100 +#106 := [trans #101 #104]: #105 +#129 := [monotonicity #106 #126]: #128 +#97 := (iff #29 #96) +#98 := [rewrite]: #97 +#132 := [monotonicity #98 #129]: #131 +#94 := (iff #28 #93) +#95 := [monotonicity #92]: #94 +#135 := [monotonicity #95 #132]: #134 +#89 := (iff #26 #86) +#83 := (and #15 #6) +#87 := (iff #83 #86) +#88 := [rewrite]: #87 +#84 := (iff #26 #83) +#85 := [monotonicity #82]: #84 +#90 := [trans #85 #88]: #89 +#138 := [monotonicity #90 #135]: #137 +#79 := (iff #24 #78) +#80 := [rewrite]: #79 +#141 := [monotonicity #80 #138]: #140 +#76 := (iff #22 #75) +#77 := [monotonicity #74]: #76 +#144 := [monotonicity #77 #141]: #143 +#147 := [monotonicity #144]: #146 +#71 := (iff #18 #70) +#72 := [rewrite]: #71 +#150 := [monotonicity #72 #147]: #149 +#68 := (iff #16 #67) +#69 := [rewrite]: #68 +#153 := [monotonicity #69 #150]: #152 +#156 := [monotonicity #153]: #155 +#159 := [monotonicity #156]: #158 +#162 := [monotonicity #159]: #161 +#165 := [monotonicity #162]: #164 +#346 := [trans #165 #344]: #345 +#66 := [asserted]: #50 +#347 := [mp #66 #346]: #342 +#355 := [not-or-elim #347]: #354 +#466 := [mp #355 #465]: #455 +#686 := [unit-resolution #466 #685]: #190 +#427 := (or #170 #189 #238) +#350 := (not #174) +#430 := (iff #350 #427) +#382 := (or #189 #238) +#414 := (or #170 #382) +#428 := (iff #414 #427) +#429 := [rewrite]: #428 +#425 := (iff #350 #414) +#415 := (not #414) +#420 := (not #415) +#423 := (iff #420 #414) +#424 := [rewrite]: #423 +#421 := (iff #350 #420) +#418 := (iff #174 #415) +#380 := (not #382) +#411 := (and #380 #171) +#416 := (iff #411 #415) +#417 := [rewrite]: #416 +#412 := (iff #174 #411) +#383 := (iff #10 #380) +#384 := [rewrite]: #383 +#413 := [monotonicity #384]: #412 +#419 := [trans #413 #417]: #418 +#422 := [monotonicity #419]: #421 +#426 := [trans #422 #424]: #425 +#431 := [trans #426 #429]: #430 +#351 := [not-or-elim #347]: #350 +#432 := [mp #351 #431]: #427 +#687 := [unit-resolution #432 #686 #685]: #238 +#549 := (or #170 #218) +#364 := (not #232) +#558 := (iff #364 #549) +#550 := (not #549) +#553 := (not #550) +#556 := (iff #553 #549) +#557 := [rewrite]: #556 +#554 := (iff #364 #553) +#551 := (iff #232 #550) +#552 := [rewrite]: #551 +#555 := [monotonicity #552]: #554 +#559 := [trans #555 #557]: #558 +#365 := [not-or-elim #347]: #364 +#560 := [mp #365 #559]: #549 +#688 := [unit-resolution #560 #685]: #218 +#577 := (or #205 #217 #237) +#366 := (not #244) +#580 := (iff #366 #577) +#385 := (or #205 #217) +#564 := (or #237 #385) +#578 := (iff #564 #577) +#579 := [rewrite]: #578 +#575 := (iff #366 #564) +#565 := (not #564) +#570 := (not #565) +#573 := (iff #570 #564) +#574 := [rewrite]: #573 +#571 := (iff #366 #570) +#568 := (iff #244 #565) +#386 := (not #385) +#561 := (and #386 #238) +#566 := (iff #561 #565) +#567 := [rewrite]: #566 +#562 := (iff #244 #561) +#387 := (iff #17 #386) +#388 := [rewrite]: #387 +#563 := [monotonicity #388]: #562 +#569 := [trans #563 #567]: #568 +#572 := [monotonicity #569]: #571 +#576 := [trans #572 #574]: #575 +#581 := [trans #576 #579]: #580 +#367 := [not-or-elim #347]: #366 +#582 := [mp #367 #581]: #577 +#689 := [unit-resolution #582 #688 #687]: #205 +#583 := (or #204 #237) +#368 := (not #252) +#592 := (iff #368 #583) +#584 := (not #583) +#587 := (not #584) +#590 := (iff #587 #583) +#591 := [rewrite]: #590 +#588 := (iff #368 #587) +#585 := (iff #252 #584) +#586 := [rewrite]: #585 +#589 := [monotonicity #586]: #588 +#593 := [trans #589 #591]: #592 +#369 := [not-or-elim #347]: #368 +#594 := [mp #369 #593]: #583 +#690 := [unit-resolution #594 #689 #687]: false +#691 := [lemma #690]: #170 +#487 := (or #171 #177 #204) +#356 := (not #208) +#490 := (iff #356 #487) +#467 := (or #171 #177) +#474 := (or #204 #467) +#488 := (iff #474 #487) +#489 := [rewrite]: #488 +#485 := (iff #356 #474) +#475 := (not #474) +#480 := (not #475) +#483 := (iff #480 #474) +#484 := [rewrite]: #483 +#481 := (iff #356 #480) +#478 := (iff #208 #475) +#468 := (not #467) +#471 := (and #468 #205) +#476 := (iff #471 #475) +#477 := [rewrite]: #476 +#472 := (iff #208 #471) +#469 := (iff #15 #468) +#470 := [rewrite]: #469 +#473 := [monotonicity #470]: #472 +#479 := [trans #473 #477]: #478 +#482 := [monotonicity #479]: #481 +#486 := [trans #482 #484]: #485 +#491 := [trans #486 #489]: #490 +#357 := [not-or-elim #347]: #356 +#492 := [mp #357 #491]: #487 +#694 := [unit-resolution #492 #691]: #693 +#695 := [unit-resolution #694 #692]: #177 +#493 := (or #178 #204) +#358 := (not #211) +#502 := (iff #358 #493) +#494 := (not #493) +#497 := (not #494) +#500 := (iff #497 #493) +#501 := [rewrite]: #500 +#498 := (iff #358 #497) +#495 := (iff #211 #494) +#496 := [rewrite]: #495 +#499 := [monotonicity #496]: #498 +#503 := [trans #499 #501]: #502 +#359 := [not-or-elim #347]: #358 +#504 := [mp #359 #503]: #493 +#696 := [unit-resolution #504 #695 #692]: false +#697 := [lemma #696]: #204 +#698 := [hypothesis]: #177 +#449 := (or #178 #205 #217) +#352 := (not #184) +#452 := (iff #352 #449) +#436 := (or #178 #385) +#450 := (iff #436 #449) +#451 := [rewrite]: #450 +#447 := (iff #352 #436) +#437 := (not #436) +#442 := (not #437) +#445 := (iff #442 #436) +#446 := [rewrite]: #445 +#443 := (iff #352 #442) +#440 := (iff #184 #437) +#433 := (and #386 #177) +#438 := (iff #433 #437) +#439 := [rewrite]: #438 +#434 := (iff #184 #433) +#435 := [monotonicity #388]: #434 +#441 := [trans #435 #439]: #440 +#444 := [monotonicity #441]: #443 +#448 := [trans #444 #446]: #447 +#453 := [trans #448 #451]: #452 +#353 := [not-or-elim #347]: #352 +#454 := [mp #353 #453]: #449 +#699 := [unit-resolution #454 #698 #697]: #217 +#639 := (or #178 #237) +#374 := (not #273) +#648 := (iff #374 #639) +#640 := (not #639) +#643 := (not #640) +#646 := (iff #643 #639) +#647 := [rewrite]: #646 +#644 := (iff #374 #643) +#641 := (iff #273 #640) +#642 := [rewrite]: #641 +#645 := [monotonicity #642]: #644 +#649 := [trans #645 #647]: #648 +#375 := [not-or-elim #347]: #374 +#650 := [mp #375 #649]: #639 +#700 := [unit-resolution #650 #698]: #237 +#667 := (or #189 #218 #238) +#376 := (not #279) +#670 := (iff #376 #667) +#654 := (or #218 #382) +#668 := (iff #654 #667) +#669 := [rewrite]: #668 +#665 := (iff #376 #654) +#655 := (not #654) +#660 := (not #655) +#663 := (iff #660 #654) +#664 := [rewrite]: #663 +#661 := (iff #376 #660) +#658 := (iff #279 #655) +#651 := (and #380 #217) +#656 := (iff #651 #655) +#657 := [rewrite]: #656 +#652 := (iff #279 #651) +#653 := [monotonicity #384]: #652 +#659 := [trans #653 #657]: #658 +#662 := [monotonicity #659]: #661 +#666 := [trans #662 #664]: #665 +#671 := [trans #666 #669]: #670 +#377 := [not-or-elim #347]: #376 +#672 := [mp #377 #671]: #667 +#701 := [unit-resolution #672 #699 #700]: #189 +#673 := (or #190 #218) +#378 := (not #287) +#682 := (iff #378 #673) +#674 := (not #673) +#677 := (not #674) +#680 := (iff #677 #673) +#681 := [rewrite]: #680 +#678 := (iff #378 #677) +#675 := (iff #287 #674) +#676 := [rewrite]: #675 +#679 := [monotonicity #676]: #678 +#683 := [trans #679 #681]: #682 +#379 := [not-or-elim #347]: #378 +#684 := [mp #379 #683]: #673 +#702 := [unit-resolution #684 #701 #699]: false +#703 := [lemma #702]: #178 +#704 := (or #177 #218) +#543 := (or #171 #177 #218) +#362 := (not #224) +#546 := (iff #362 #543) +#530 := (or #218 #467) +#544 := (iff #530 #543) +#545 := [rewrite]: #544 +#541 := (iff #362 #530) +#531 := (not #530) +#536 := (not #531) +#539 := (iff #536 #530) +#540 := [rewrite]: #539 +#537 := (iff #362 #536) +#534 := (iff #224 #531) +#527 := (and #468 #217) +#532 := (iff #527 #531) +#533 := [rewrite]: #532 +#528 := (iff #224 #527) +#529 := [monotonicity #470]: #528 +#535 := [trans #529 #533]: #534 +#538 := [monotonicity #535]: #537 +#542 := [trans #538 #540]: #541 +#547 := [trans #542 #545]: #546 +#363 := [not-or-elim #347]: #362 +#548 := [mp #363 #547]: #543 +#705 := [unit-resolution #548 #691]: #704 +#706 := [unit-resolution #705 #703]: #218 +#707 := (or #177 #237) +#633 := (or #171 #177 #237) +#372 := (not #268) +#636 := (iff #372 #633) +#620 := (or #237 #467) +#634 := (iff #620 #633) +#635 := [rewrite]: #634 +#631 := (iff #372 #620) +#621 := (not #620) +#626 := (not #621) +#629 := (iff #626 #620) +#630 := [rewrite]: #629 +#627 := (iff #372 #626) +#624 := (iff #268 #621) +#617 := (and #468 #238) +#622 := (iff #617 #621) +#623 := [rewrite]: #622 +#618 := (iff #268 #617) +#619 := [monotonicity #470]: #618 +#625 := [trans #619 #623]: #624 +#628 := [monotonicity #625]: #627 +#632 := [trans #628 #630]: #631 +#637 := [trans #632 #635]: #636 +#373 := [not-or-elim #347]: #372 +#638 := [mp #373 #637]: #633 +#708 := [unit-resolution #638 #691]: #707 +#709 := [unit-resolution #708 #703]: #237 +#611 := (or #190 #205 #217) +#370 := (not #260) +#614 := (iff #370 #611) +#598 := (or #190 #385) +#612 := (iff #598 #611) +#613 := [rewrite]: #612 +#609 := (iff #370 #598) +#599 := (not #598) +#604 := (not #599) +#607 := (iff #604 #598) +#608 := [rewrite]: #607 +#605 := (iff #370 #604) +#602 := (iff #260 #599) +#595 := (and #386 #189) +#600 := (iff #595 #599) +#601 := [rewrite]: #600 +#596 := (iff #260 #595) +#597 := [monotonicity #388]: #596 +#603 := [trans #597 #601]: #602 +#606 := [monotonicity #603]: #605 +#610 := [trans #606 #608]: #609 +#615 := [trans #610 #613]: #614 +#371 := [not-or-elim #347]: #370 +#616 := [mp #371 #615]: #611 +#710 := [unit-resolution #616 #706 #697]: #190 +#405 := (or #189 #205 #217 #238) +#348 := (not #115) +#408 := (iff #348 #405) +#392 := (or #382 #385) +#406 := (iff #392 #405) +#407 := [rewrite]: #406 +#403 := (iff #348 #392) +#393 := (not #392) +#398 := (not #393) +#401 := (iff #398 #392) +#402 := [rewrite]: #401 +#399 := (iff #348 #398) +#396 := (iff #115 #393) +#389 := (and #380 #386) +#394 := (iff #389 #393) +#395 := [rewrite]: #394 +#390 := (iff #115 #389) +#391 := [monotonicity #384 #388]: #390 +#397 := [trans #391 #395]: #396 +#400 := [monotonicity #397]: #399 +#404 := [trans #400 #402]: #403 +#409 := [trans #404 #407]: #408 +#349 := [not-or-elim #347]: #348 +#410 := [mp #349 #409]: #405 +[unit-resolution #410 #710 #709 #697 #706]: false +unsat diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_linarith_16 --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_linarith_16 Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,26 @@ +(benchmark Isabelle +:extrafuns ( + (uf_3 Int) + (uf_2 Int) + (uf_1 Int) + (uf_4 Int) + (uf_5 Int) + (uf_6 Int) + (uf_7 Int) + (uf_8 Int) + (uf_9 Int) + (uf_10 Int) + (uf_11 Int) + ) +:assumption (= uf_1 (- (ite (< uf_2 0) (~ uf_2) uf_2) uf_3)) +:assumption (= uf_4 (- (ite (< uf_1 0) (~ uf_1) uf_1) uf_2)) +:assumption (= uf_5 (- (ite (< uf_4 0) (~ uf_4) uf_4) uf_1)) +:assumption (= uf_6 (- (ite (< uf_5 0) (~ uf_5) uf_5) uf_4)) +:assumption (= uf_7 (- (ite (< uf_6 0) (~ uf_6) uf_6) uf_5)) +:assumption (= uf_8 (- (ite (< uf_7 0) (~ uf_7) uf_7) uf_6)) +:assumption (= uf_9 (- (ite (< uf_8 0) (~ uf_8) uf_8) uf_7)) +:assumption (= uf_10 (- (ite (< uf_9 0) (~ uf_9) uf_9) uf_8)) +:assumption (= uf_11 (- (ite (< uf_10 0) (~ uf_10) uf_10) uf_9)) +:assumption (not (and (= uf_3 uf_10) (= uf_2 uf_11))) +:formula true +) diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_linarith_16.proof --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_linarith_16.proof Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,2291 @@ +#2 := false +#6 := 0::int +decl z3name!0 :: int +#647 := z3name!0 +#81 := -1::int +#656 := (* -1::int z3name!0) +decl uf_2 :: int +#5 := uf_2 +#882 := (+ uf_2 #656) +#883 := (<= #882 0::int) +#885 := (not #883) +#881 := (>= #882 0::int) +#884 := (not #881) +#886 := (or #884 #885) +decl uf_11 :: int +#55 := uf_11 +#513 := (* -1::int uf_11) +#514 := (+ uf_2 #513) +#515 := (<= #514 0::int) +decl z3name!5 :: int +#777 := z3name!5 +decl uf_7 :: int +#31 := uf_7 +#1083 := (+ uf_7 z3name!5) +#1084 := (<= #1083 0::int) +#335 := (>= uf_7 0::int) +#1085 := (>= #1083 0::int) +#1087 := (not #1085) +#1086 := (not #1084) +#1088 := (or #1086 #1087) +#2302 := [hypothesis]: #1086 +#1289 := (or #1088 #1084) +#1290 := [def-axiom]: #1289 +#2303 := [unit-resolution #1290 #2302]: #1088 +#1089 := (not #1088) +#1092 := (or #335 #1089) +#1099 := (not #1092) +#786 := (* -1::int z3name!5) +#1072 := (+ uf_7 #786) +#1073 := (<= #1072 0::int) +#1075 := (not #1073) +#1071 := (>= #1072 0::int) +#1074 := (not #1071) +#1076 := (or #1074 #1075) +#1077 := (not #1076) +#336 := (not #335) +#1080 := (or #336 #1077) +#1098 := (not #1080) +#1100 := (or #1098 #1099) +#1101 := (not #1100) +#318 := (* -1::int uf_7) +#780 := (= z3name!5 #318) +#781 := (or #335 #780) +#778 := (= z3name!5 uf_7) +#779 := (or #336 #778) +#782 := (and #779 #781) +#1104 := (iff #782 #1101) +#1095 := (and #1080 #1092) +#1102 := (iff #1095 #1101) +#1103 := [rewrite]: #1102 +#1096 := (iff #782 #1095) +#1093 := (iff #781 #1092) +#1090 := (iff #780 #1089) +#1091 := [rewrite]: #1090 +#1094 := [monotonicity #1091]: #1093 +#1081 := (iff #779 #1080) +#1078 := (iff #778 #1077) +#1079 := [rewrite]: #1078 +#1082 := [monotonicity #1079]: #1081 +#1097 := [monotonicity #1082 #1094]: #1096 +#1105 := [trans #1097 #1103]: #1104 +#783 := [intro-def]: #782 +#1106 := [mp #783 #1105]: #1101 +#1108 := [not-or-elim #1106]: #1092 +#2304 := [unit-resolution #1108 #2303]: #335 +decl uf_4 :: int +#13 := uf_4 +#194 := (>= uf_4 0::int) +decl uf_10 :: int +#49 := uf_10 +#459 := (* -1::int uf_10) +decl uf_3 :: int +#10 := uf_3 +#508 := (+ uf_3 #459) +#509 := (>= #508 0::int) +decl z3name!1 :: int +#673 := z3name!1 +#682 := (* -1::int z3name!1) +decl uf_1 :: int +#4 := uf_1 +#920 := (+ uf_1 #682) +#921 := (<= #920 0::int) +#931 := (+ uf_1 z3name!1) +#933 := (>= #931 0::int) +#935 := (not #933) +#932 := (<= #931 0::int) +#934 := (not #932) +#936 := (or #934 #935) +#937 := (not #936) +#147 := (>= uf_1 0::int) +#148 := (not #147) +#923 := (not #921) +#919 := (>= #920 0::int) +#922 := (not #919) +#924 := (or #922 #923) +#2022 := [hypothesis]: #923 +#1237 := (or #924 #921) +#1238 := [def-axiom]: #1237 +#2023 := [unit-resolution #1238 #2022]: #924 +#925 := (not #924) +#928 := (or #148 #925) +#940 := (or #147 #937) +#947 := (not #940) +#946 := (not #928) +#948 := (or #946 #947) +#949 := (not #948) +#130 := (* -1::int uf_1) +#676 := (= z3name!1 #130) +#677 := (or #147 #676) +#674 := (= z3name!1 uf_1) +#675 := (or #148 #674) +#678 := (and #675 #677) +#952 := (iff #678 #949) +#943 := (and #928 #940) +#950 := (iff #943 #949) +#951 := [rewrite]: #950 +#944 := (iff #678 #943) +#941 := (iff #677 #940) +#938 := (iff #676 #937) +#939 := [rewrite]: #938 +#942 := [monotonicity #939]: #941 +#929 := (iff #675 #928) +#926 := (iff #674 #925) +#927 := [rewrite]: #926 +#930 := [monotonicity #927]: #929 +#945 := [monotonicity #930 #942]: #944 +#953 := [trans #945 #951]: #952 +#679 := [intro-def]: #678 +#954 := [mp #679 #953]: #949 +#955 := [not-or-elim #954]: #928 +#2024 := [unit-resolution #955 #2023]: #148 +#956 := [not-or-elim #954]: #940 +#2025 := [unit-resolution #956 #2024]: #937 +#2026 := (or #921 #919) +#2027 := [th-lemma]: #2026 +#2028 := [unit-resolution #2027 #2022]: #919 +#2029 := (or #922 #147 #935) +#2030 := [th-lemma]: #2029 +#2031 := [unit-resolution #2030 #2024 #2028]: #935 +#1243 := (or #936 #933) +#1244 := [def-axiom]: #1243 +#2032 := [unit-resolution #1244 #2031 #2025]: false +#2033 := [lemma #2032]: #921 +decl z3name!7 :: int +#829 := z3name!7 +decl uf_9 :: int +#43 := uf_9 +#1159 := (+ uf_9 z3name!7) +#1160 := (<= #1159 0::int) +#838 := (* -1::int z3name!7) +#1148 := (+ uf_9 #838) +#1147 := (>= #1148 0::int) +decl z3name!4 :: int +#751 := z3name!4 +#760 := (* -1::int z3name!4) +decl uf_6 :: int +#25 := uf_6 +#1034 := (+ uf_6 #760) +#1033 := (>= #1034 0::int) +#1035 := (<= #1034 0::int) +#1037 := (not #1035) +#1036 := (not #1033) +#1038 := (or #1036 #1037) +#1039 := (not #1038) +#288 := (>= uf_6 0::int) +#893 := (+ uf_2 z3name!0) +#895 := (>= #893 0::int) +#897 := (not #895) +#894 := (<= #893 0::int) +#896 := (not #894) +#898 := (or #896 #897) +#899 := (not #898) +#100 := (>= uf_2 0::int) +#101 := (not #100) +#1736 := [hypothesis]: #885 +#1225 := (or #886 #883) +#1226 := [def-axiom]: #1225 +#1737 := [unit-resolution #1226 #1736]: #886 +#887 := (not #886) +#890 := (or #101 #887) +#902 := (or #100 #899) +#909 := (not #902) +#908 := (not #890) +#910 := (or #908 #909) +#911 := (not #910) +#82 := (* -1::int uf_2) +#650 := (= z3name!0 #82) +#651 := (or #100 #650) +#648 := (= z3name!0 uf_2) +#649 := (or #101 #648) +#652 := (and #649 #651) +#914 := (iff #652 #911) +#905 := (and #890 #902) +#912 := (iff #905 #911) +#913 := [rewrite]: #912 +#906 := (iff #652 #905) +#903 := (iff #651 #902) +#900 := (iff #650 #899) +#901 := [rewrite]: #900 +#904 := [monotonicity #901]: #903 +#891 := (iff #649 #890) +#888 := (iff #648 #887) +#889 := [rewrite]: #888 +#892 := [monotonicity #889]: #891 +#907 := [monotonicity #892 #904]: #906 +#915 := [trans #907 #913]: #914 +#653 := [intro-def]: #652 +#916 := [mp #653 #915]: #911 +#917 := [not-or-elim #916]: #890 +#1738 := [unit-resolution #917 #1737]: #101 +#918 := [not-or-elim #916]: #902 +#1739 := [unit-resolution #918 #1738]: #899 +#1231 := (or #898 #895) +#1232 := [def-axiom]: #1231 +#1740 := [unit-resolution #1232 #1739]: #895 +#1741 := [th-lemma #1736 #1738 #1740]: false +#1742 := [lemma #1741]: #883 +#1149 := (<= #1148 0::int) +#1151 := (not #1149) +#1150 := (not #1147) +#1152 := (or #1150 #1151) +#1153 := (not #1152) +#429 := (>= uf_9 0::int) +decl z3name!6 :: int +#803 := z3name!6 +#812 := (* -1::int z3name!6) +decl uf_8 :: int +#37 := uf_8 +#1110 := (+ uf_8 #812) +#1111 := (<= #1110 0::int) +#1113 := (not #1111) +#1109 := (>= #1110 0::int) +#1112 := (not #1109) +#1114 := (or #1112 #1113) +#1865 := [hypothesis]: #1113 +#1297 := (or #1114 #1111) +#1298 := [def-axiom]: #1297 +#1866 := [unit-resolution #1298 #1865]: #1114 +#382 := (>= uf_8 0::int) +#1685 := (or #1111 #1109) +#1686 := [th-lemma]: #1685 +#1867 := [unit-resolution #1686 #1865]: #1109 +#1734 := (or #382 #1112) +#1121 := (+ uf_8 z3name!6) +#1123 := (>= #1121 0::int) +#1125 := (not #1123) +#1122 := (<= #1121 0::int) +#1124 := (not #1122) +#1126 := (or #1124 #1125) +#1127 := (not #1126) +#383 := (not #382) +#1428 := [hypothesis]: #383 +#1130 := (or #382 #1127) +#1137 := (not #1130) +#1115 := (not #1114) +#1118 := (or #383 #1115) +#1136 := (not #1118) +#1138 := (or #1136 #1137) +#1139 := (not #1138) +#365 := (* -1::int uf_8) +#806 := (= z3name!6 #365) +#807 := (or #382 #806) +#804 := (= z3name!6 uf_8) +#805 := (or #383 #804) +#808 := (and #805 #807) +#1142 := (iff #808 #1139) +#1133 := (and #1118 #1130) +#1140 := (iff #1133 #1139) +#1141 := [rewrite]: #1140 +#1134 := (iff #808 #1133) +#1131 := (iff #807 #1130) +#1128 := (iff #806 #1127) +#1129 := [rewrite]: #1128 +#1132 := [monotonicity #1129]: #1131 +#1119 := (iff #805 #1118) +#1116 := (iff #804 #1115) +#1117 := [rewrite]: #1116 +#1120 := [monotonicity #1117]: #1119 +#1135 := [monotonicity #1120 #1132]: #1134 +#1143 := [trans #1135 #1141]: #1142 +#809 := [intro-def]: #808 +#1144 := [mp #809 #1143]: #1139 +#1146 := [not-or-elim #1144]: #1130 +#1729 := [unit-resolution #1146 #1428]: #1127 +#1637 := [hypothesis]: #1109 +#1730 := (or #1112 #1125 #382) +#1731 := [th-lemma]: #1730 +#1732 := [unit-resolution #1731 #1428 #1637]: #1125 +#1303 := (or #1126 #1123) +#1304 := [def-axiom]: #1303 +#1733 := [unit-resolution #1304 #1732 #1729]: false +#1735 := [lemma #1733]: #1734 +#1868 := [unit-resolution #1735 #1867]: #382 +#1145 := [not-or-elim #1144]: #1118 +#1869 := [unit-resolution #1145 #1868 #1866]: false +#1870 := [lemma #1869]: #1111 +#289 := (not #288) +#1405 := [hypothesis]: #289 +#1688 := (or #288 #429 #1113) +#815 := (+ uf_9 #812) +#818 := (+ uf_7 #815) +#825 := (>= #818 0::int) +#389 := (ite #382 uf_8 #365) +#400 := (* -1::int #389) +#401 := (+ uf_9 #400) +#402 := (+ uf_7 #401) +#599 := (>= #402 0::int) +#826 := (= #599 #825) +#819 := (~ #402 #818) +#816 := (~ #401 #815) +#813 := (~ #400 #812) +#810 := (~ #389 z3name!6) +#811 := [apply-def #809]: #810 +#814 := [monotonicity #811]: #813 +#817 := [monotonicity #814]: #816 +#820 := [monotonicity #817]: #819 +#827 := [monotonicity #820]: #826 +#601 := (not #599) +#598 := (<= #402 0::int) +#600 := (not #598) +#602 := (or #600 #601) +#603 := (not #602) +#403 := (= #402 0::int) +#604 := (iff #403 #603) +#605 := [rewrite]: #604 +#45 := (- uf_8) +#44 := (< uf_8 0::int) +#46 := (ite #44 #45 uf_8) +#47 := (- #46 uf_7) +#48 := (= uf_9 #47) +#408 := (iff #48 #403) +#368 := (ite #44 #365 uf_8) +#374 := (+ #318 #368) +#379 := (= uf_9 #374) +#406 := (iff #379 #403) +#394 := (+ #318 #389) +#397 := (= uf_9 #394) +#404 := (iff #397 #403) +#405 := [rewrite]: #404 +#398 := (iff #379 #397) +#395 := (= #374 #394) +#392 := (= #368 #389) +#386 := (ite #383 #365 uf_8) +#390 := (= #386 #389) +#391 := [rewrite]: #390 +#387 := (= #368 #386) +#384 := (iff #44 #383) +#385 := [rewrite]: #384 +#388 := [monotonicity #385]: #387 +#393 := [trans #388 #391]: #392 +#396 := [monotonicity #393]: #395 +#399 := [monotonicity #396]: #398 +#407 := [trans #399 #405]: #406 +#380 := (iff #48 #379) +#377 := (= #47 #374) +#371 := (- #368 uf_7) +#375 := (= #371 #374) +#376 := [rewrite]: #375 +#372 := (= #47 #371) +#369 := (= #46 #368) +#366 := (= #45 #365) +#367 := [rewrite]: #366 +#370 := [monotonicity #367]: #369 +#373 := [monotonicity #370]: #372 +#378 := [trans #373 #376]: #377 +#381 := [monotonicity #378]: #380 +#409 := [trans #381 #407]: #408 +#364 := [asserted]: #48 +#410 := [mp #364 #409]: #403 +#606 := [mp #410 #605]: #603 +#608 := [not-or-elim #606]: #599 +#828 := [mp~ #608 #827]: #825 +#1441 := [hypothesis]: #1075 +#1285 := (or #1076 #1073) +#1286 := [def-axiom]: #1285 +#1442 := [unit-resolution #1286 #1441]: #1076 +#1107 := [not-or-elim #1106]: #1080 +#1443 := [unit-resolution #1107 #1442]: #336 +#1444 := [unit-resolution #1108 #1443]: #1089 +#1291 := (or #1088 #1085) +#1292 := [def-axiom]: #1291 +#1445 := [unit-resolution #1292 #1444]: #1085 +#1446 := [th-lemma #1441 #1445 #1443]: false +#1447 := [lemma #1446]: #1073 +#789 := (+ uf_8 #786) +#792 := (+ uf_6 #789) +#799 := (>= #792 0::int) +#342 := (ite #335 uf_7 #318) +#353 := (* -1::int #342) +#354 := (+ uf_8 #353) +#355 := (+ uf_6 #354) +#588 := (>= #355 0::int) +#800 := (= #588 #799) +#793 := (~ #355 #792) +#790 := (~ #354 #789) +#787 := (~ #353 #786) +#784 := (~ #342 z3name!5) +#785 := [apply-def #783]: #784 +#788 := [monotonicity #785]: #787 +#791 := [monotonicity #788]: #790 +#794 := [monotonicity #791]: #793 +#801 := [monotonicity #794]: #800 +#590 := (not #588) +#587 := (<= #355 0::int) +#589 := (not #587) +#591 := (or #589 #590) +#592 := (not #591) +#356 := (= #355 0::int) +#593 := (iff #356 #592) +#594 := [rewrite]: #593 +#39 := (- uf_7) +#38 := (< uf_7 0::int) +#40 := (ite #38 #39 uf_7) +#41 := (- #40 uf_6) +#42 := (= uf_8 #41) +#361 := (iff #42 #356) +#321 := (ite #38 #318 uf_7) +#271 := (* -1::int uf_6) +#327 := (+ #271 #321) +#332 := (= uf_8 #327) +#359 := (iff #332 #356) +#347 := (+ #271 #342) +#350 := (= uf_8 #347) +#357 := (iff #350 #356) +#358 := [rewrite]: #357 +#351 := (iff #332 #350) +#348 := (= #327 #347) +#345 := (= #321 #342) +#339 := (ite #336 #318 uf_7) +#343 := (= #339 #342) +#344 := [rewrite]: #343 +#340 := (= #321 #339) +#337 := (iff #38 #336) +#338 := [rewrite]: #337 +#341 := [monotonicity #338]: #340 +#346 := [trans #341 #344]: #345 +#349 := [monotonicity #346]: #348 +#352 := [monotonicity #349]: #351 +#360 := [trans #352 #358]: #359 +#333 := (iff #42 #332) +#330 := (= #41 #327) +#324 := (- #321 uf_6) +#328 := (= #324 #327) +#329 := [rewrite]: #328 +#325 := (= #41 #324) +#322 := (= #40 #321) +#319 := (= #39 #318) +#320 := [rewrite]: #319 +#323 := [monotonicity #320]: #322 +#326 := [monotonicity #323]: #325 +#331 := [trans #326 #329]: #330 +#334 := [monotonicity #331]: #333 +#362 := [trans #334 #360]: #361 +#317 := [asserted]: #42 +#363 := [mp #317 #362]: #356 +#595 := [mp #363 #594]: #592 +#597 := [not-or-elim #595]: #588 +#802 := [mp~ #597 #801]: #799 +#1343 := (not #825) +#1350 := (not #799) +#1351 := (or #288 #1075 #1350 #429 #1113 #1343) +#1352 := [th-lemma]: #1351 +#1689 := [unit-resolution #1352 #802 #1447 #828]: #1688 +#2046 := [unit-resolution #1689 #1405 #1870]: #429 +#430 := (not #429) +#1156 := (or #430 #1153) +#1161 := (>= #1159 0::int) +#1163 := (not #1161) +#1162 := (not #1160) +#1164 := (or #1162 #1163) +#1165 := (not #1164) +#1168 := (or #429 #1165) +#1175 := (not #1168) +#1174 := (not #1156) +#1176 := (or #1174 #1175) +#1177 := (not #1176) +#412 := (* -1::int uf_9) +#832 := (= z3name!7 #412) +#833 := (or #429 #832) +#830 := (= z3name!7 uf_9) +#831 := (or #430 #830) +#834 := (and #831 #833) +#1180 := (iff #834 #1177) +#1171 := (and #1156 #1168) +#1178 := (iff #1171 #1177) +#1179 := [rewrite]: #1178 +#1172 := (iff #834 #1171) +#1169 := (iff #833 #1168) +#1166 := (iff #832 #1165) +#1167 := [rewrite]: #1166 +#1170 := [monotonicity #1167]: #1169 +#1157 := (iff #831 #1156) +#1154 := (iff #830 #1153) +#1155 := [rewrite]: #1154 +#1158 := [monotonicity #1155]: #1157 +#1173 := [monotonicity #1158 #1170]: #1172 +#1181 := [trans #1173 #1179]: #1180 +#835 := [intro-def]: #834 +#1182 := [mp #835 #1181]: #1177 +#1183 := [not-or-elim #1182]: #1156 +#2047 := [unit-resolution #1183 #2046]: #1153 +#1307 := (or #1152 #1147) +#1308 := [def-axiom]: #1307 +#2112 := [unit-resolution #1308 #2047]: #1147 +#2009 := (or #288 #382) +#1998 := (or #1036 #288) +#1045 := (+ uf_6 z3name!4) +#1047 := (>= #1045 0::int) +#1049 := (not #1047) +#1046 := (<= #1045 0::int) +#1048 := (not #1046) +#1050 := (or #1048 #1049) +#1460 := [hypothesis]: #1049 +#1279 := (or #1050 #1047) +#1280 := [def-axiom]: #1279 +#1461 := [unit-resolution #1280 #1460]: #1050 +#1464 := (or #1047 #289) +#1051 := (not #1050) +#1448 := [hypothesis]: #1037 +#1273 := (or #1038 #1035) +#1274 := [def-axiom]: #1273 +#1449 := [unit-resolution #1274 #1448]: #1038 +#1042 := (or #289 #1039) +#1054 := (or #288 #1051) +#1061 := (not #1054) +#1060 := (not #1042) +#1062 := (or #1060 #1061) +#1063 := (not #1062) +#754 := (= z3name!4 #271) +#755 := (or #288 #754) +#752 := (= z3name!4 uf_6) +#753 := (or #289 #752) +#756 := (and #753 #755) +#1066 := (iff #756 #1063) +#1057 := (and #1042 #1054) +#1064 := (iff #1057 #1063) +#1065 := [rewrite]: #1064 +#1058 := (iff #756 #1057) +#1055 := (iff #755 #1054) +#1052 := (iff #754 #1051) +#1053 := [rewrite]: #1052 +#1056 := [monotonicity #1053]: #1055 +#1043 := (iff #753 #1042) +#1040 := (iff #752 #1039) +#1041 := [rewrite]: #1040 +#1044 := [monotonicity #1041]: #1043 +#1059 := [monotonicity #1044 #1056]: #1058 +#1067 := [trans #1059 #1065]: #1066 +#757 := [intro-def]: #756 +#1068 := [mp #757 #1067]: #1063 +#1069 := [not-or-elim #1068]: #1042 +#1450 := [unit-resolution #1069 #1449]: #289 +#1070 := [not-or-elim #1068]: #1054 +#1451 := [unit-resolution #1070 #1450]: #1051 +#1452 := (or #1035 #1033) +#1453 := [th-lemma]: #1452 +#1454 := [unit-resolution #1453 #1448]: #1033 +#1455 := (or #1036 #288 #1049) +#1456 := [th-lemma]: #1455 +#1457 := [unit-resolution #1456 #1450 #1454]: #1049 +#1458 := [unit-resolution #1280 #1457 #1451]: false +#1459 := [lemma #1458]: #1035 +#1462 := (or #1047 #1037 #289) +#1463 := [th-lemma]: #1462 +#1465 := [unit-resolution #1463 #1459]: #1464 +#1466 := [unit-resolution #1465 #1460]: #289 +#1467 := [unit-resolution #1070 #1466 #1461]: false +#1468 := [lemma #1467]: #1047 +#1999 := [unit-resolution #1456 #1468]: #1998 +#2000 := [unit-resolution #1999 #1405]: #1036 +#1407 := [unit-resolution #1070 #1405]: #1051 +#1277 := (or #1050 #1046) +#1278 := [def-axiom]: #1277 +#1497 := [unit-resolution #1278 #1407]: #1046 +#2001 := (or #336 #1048 #1033 #382 #1350 #1075) +#2002 := [th-lemma]: #2001 +#2003 := [unit-resolution #2002 #1497 #2000 #1447 #802 #1428]: #336 +#2004 := (or #1087 #1075 #1048 #1033 #382 #1350) +#2005 := [th-lemma]: #2004 +#2006 := [unit-resolution #2005 #1497 #1447 #2000 #802 #1428]: #1087 +#2007 := [unit-resolution #1292 #2006]: #1088 +#2008 := [unit-resolution #1108 #2007 #2003]: false +#2010 := [lemma #2008]: #2009 +#2113 := [unit-resolution #2010 #1405]: #382 +#2114 := [unit-resolution #1145 #2113]: #1115 +#1295 := (or #1114 #1109) +#1296 := [def-axiom]: #1295 +#2115 := [unit-resolution #1296 #2114]: #1109 +decl z3name!2 :: int +#699 := z3name!2 +#708 := (* -1::int z3name!2) +#958 := (+ uf_4 #708) +#957 := (>= #958 0::int) +#959 := (<= #958 0::int) +#961 := (not #959) +#960 := (not #957) +#962 := (or #960 #961) +#963 := (not #962) +decl uf_5 :: int +#19 := uf_5 +#241 := (>= uf_5 0::int) +#242 := (not #241) +#1406 := [hypothesis]: #242 +#1579 := (or #1048 #241) +#516 := (>= #514 0::int) +#476 := (>= uf_10 0::int) +#477 := (not #476) +#1484 := (or #382 #241) +#1430 := (or #382 #241 #1075 #1037) +#1421 := [hypothesis]: #1035 +#1427 := [hypothesis]: #1073 +#763 := (+ uf_7 #760) +#766 := (+ uf_5 #763) +#773 := (>= #766 0::int) +#295 := (ite #288 uf_6 #271) +#306 := (* -1::int #295) +#307 := (+ uf_7 #306) +#308 := (+ uf_5 #307) +#577 := (>= #308 0::int) +#774 := (= #577 #773) +#767 := (~ #308 #766) +#764 := (~ #307 #763) +#761 := (~ #306 #760) +#758 := (~ #295 z3name!4) +#759 := [apply-def #757]: #758 +#762 := [monotonicity #759]: #761 +#765 := [monotonicity #762]: #764 +#768 := [monotonicity #765]: #767 +#775 := [monotonicity #768]: #774 +#579 := (not #577) +#576 := (<= #308 0::int) +#578 := (not #576) +#580 := (or #578 #579) +#581 := (not #580) +#309 := (= #308 0::int) +#582 := (iff #309 #581) +#583 := [rewrite]: #582 +#33 := (- uf_6) +#32 := (< uf_6 0::int) +#34 := (ite #32 #33 uf_6) +#35 := (- #34 uf_5) +#36 := (= uf_7 #35) +#314 := (iff #36 #309) +#274 := (ite #32 #271 uf_6) +#224 := (* -1::int uf_5) +#280 := (+ #224 #274) +#285 := (= uf_7 #280) +#312 := (iff #285 #309) +#300 := (+ #224 #295) +#303 := (= uf_7 #300) +#310 := (iff #303 #309) +#311 := [rewrite]: #310 +#304 := (iff #285 #303) +#301 := (= #280 #300) +#298 := (= #274 #295) +#292 := (ite #289 #271 uf_6) +#296 := (= #292 #295) +#297 := [rewrite]: #296 +#293 := (= #274 #292) +#290 := (iff #32 #289) +#291 := [rewrite]: #290 +#294 := [monotonicity #291]: #293 +#299 := [trans #294 #297]: #298 +#302 := [monotonicity #299]: #301 +#305 := [monotonicity #302]: #304 +#313 := [trans #305 #311]: #312 +#286 := (iff #36 #285) +#283 := (= #35 #280) +#277 := (- #274 uf_5) +#281 := (= #277 #280) +#282 := [rewrite]: #281 +#278 := (= #35 #277) +#275 := (= #34 #274) +#272 := (= #33 #271) +#273 := [rewrite]: #272 +#276 := [monotonicity #273]: #275 +#279 := [monotonicity #276]: #278 +#284 := [trans #279 #282]: #283 +#287 := [monotonicity #284]: #286 +#315 := [trans #287 #313]: #314 +#270 := [asserted]: #36 +#316 := [mp #270 #315]: #309 +#584 := [mp #316 #583]: #581 +#586 := [not-or-elim #584]: #577 +#776 := [mp~ #586 #775]: #773 +#1429 := [th-lemma #776 #1406 #1428 #1427 #802 #1421]: false +#1431 := [lemma #1429]: #1430 +#1485 := [unit-resolution #1431 #1447 #1459]: #1484 +#1486 := [unit-resolution #1485 #1406]: #382 +#1487 := [unit-resolution #1145 #1486]: #1115 +#1496 := [unit-resolution #1298 #1487]: #1111 +#1545 := [hypothesis]: #1046 +#1548 := (or #1048 #1113 #429) +#1546 := (or #1048 #1113 #429 #1343 #1075 #1350 #1037) +#1547 := [th-lemma]: #1546 +#1549 := [unit-resolution #1547 #1447 #802 #1459 #828]: #1548 +#1550 := [unit-resolution #1549 #1545 #1496]: #429 +#1551 := [unit-resolution #1183 #1550]: #1153 +#1552 := [unit-resolution #1308 #1551]: #1147 +#1543 := (or #477 #241 #1150) +#1488 := [unit-resolution #1296 #1487]: #1109 +#821 := (<= #818 0::int) +#822 := (= #598 #821) +#823 := [monotonicity #820]: #822 +#607 := [not-or-elim #606]: #598 +#824 := [mp~ #607 #823]: #821 +#841 := (+ uf_10 #838) +#844 := (+ uf_8 #841) +#847 := (<= #844 0::int) +#436 := (ite #429 uf_9 #412) +#447 := (* -1::int #436) +#448 := (+ uf_10 #447) +#449 := (+ uf_8 #448) +#609 := (<= #449 0::int) +#848 := (= #609 #847) +#845 := (~ #449 #844) +#842 := (~ #448 #841) +#839 := (~ #447 #838) +#836 := (~ #436 z3name!7) +#837 := [apply-def #835]: #836 +#840 := [monotonicity #837]: #839 +#843 := [monotonicity #840]: #842 +#846 := [monotonicity #843]: #845 +#849 := [monotonicity #846]: #848 +#610 := (>= #449 0::int) +#612 := (not #610) +#611 := (not #609) +#613 := (or #611 #612) +#614 := (not #613) +#450 := (= #449 0::int) +#615 := (iff #450 #614) +#616 := [rewrite]: #615 +#51 := (- uf_9) +#50 := (< uf_9 0::int) +#52 := (ite #50 #51 uf_9) +#53 := (- #52 uf_8) +#54 := (= uf_10 #53) +#455 := (iff #54 #450) +#415 := (ite #50 #412 uf_9) +#421 := (+ #365 #415) +#426 := (= uf_10 #421) +#453 := (iff #426 #450) +#441 := (+ #365 #436) +#444 := (= uf_10 #441) +#451 := (iff #444 #450) +#452 := [rewrite]: #451 +#445 := (iff #426 #444) +#442 := (= #421 #441) +#439 := (= #415 #436) +#433 := (ite #430 #412 uf_9) +#437 := (= #433 #436) +#438 := [rewrite]: #437 +#434 := (= #415 #433) +#431 := (iff #50 #430) +#432 := [rewrite]: #431 +#435 := [monotonicity #432]: #434 +#440 := [trans #435 #438]: #439 +#443 := [monotonicity #440]: #442 +#446 := [monotonicity #443]: #445 +#454 := [trans #446 #452]: #453 +#427 := (iff #54 #426) +#424 := (= #53 #421) +#418 := (- #415 uf_8) +#422 := (= #418 #421) +#423 := [rewrite]: #422 +#419 := (= #53 #418) +#416 := (= #52 #415) +#413 := (= #51 #412) +#414 := [rewrite]: #413 +#417 := [monotonicity #414]: #416 +#420 := [monotonicity #417]: #419 +#425 := [trans #420 #423]: #424 +#428 := [monotonicity #425]: #427 +#456 := [trans #428 #454]: #455 +#411 := [asserted]: #54 +#457 := [mp #411 #456]: #450 +#617 := [mp #457 #616]: #614 +#618 := [not-or-elim #617]: #609 +#850 := [mp~ #618 #849]: #847 +#1540 := [hypothesis]: #1147 +#1541 := [hypothesis]: #476 +#1542 := [th-lemma #1468 #1406 #1541 #1540 #850 #824 #1488 #776 #1459]: false +#1544 := [lemma #1542]: #1543 +#1553 := [unit-resolution #1544 #1552 #1406]: #477 +#851 := (>= #844 0::int) +#852 := (= #610 #851) +#853 := [monotonicity #846]: #852 +#619 := [not-or-elim #617]: #610 +#854 := [mp~ #619 #853]: #851 +#1309 := (or #1152 #1149) +#1310 := [def-axiom]: #1309 +#1554 := [unit-resolution #1310 #1551]: #1149 +#769 := (<= #766 0::int) +#770 := (= #576 #769) +#771 := [monotonicity #768]: #770 +#585 := [not-or-elim #584]: #576 +#772 := [mp~ #585 #771]: #769 +decl z3name!3 :: int +#725 := z3name!3 +#1007 := (+ uf_5 z3name!3) +#1009 := (>= #1007 0::int) +#1011 := (not #1009) +#1398 := [hypothesis]: #1011 +#734 := (* -1::int z3name!3) +#996 := (+ uf_5 #734) +#997 := (<= #996 0::int) +#999 := (not #997) +#995 := (>= #996 0::int) +#998 := (not #995) +#1000 := (or #998 #999) +#1001 := (not #1000) +#1008 := (<= #1007 0::int) +#1010 := (not #1008) +#1012 := (or #1010 #1011) +#1267 := (or #1012 #1009) +#1268 := [def-axiom]: #1267 +#1399 := [unit-resolution #1268 #1398]: #1012 +#1013 := (not #1012) +#1016 := (or #241 #1013) +#1023 := (not #1016) +#1004 := (or #242 #1001) +#1022 := (not #1004) +#1024 := (or #1022 #1023) +#1025 := (not #1024) +#728 := (= z3name!3 #224) +#729 := (or #241 #728) +#726 := (= z3name!3 uf_5) +#727 := (or #242 #726) +#730 := (and #727 #729) +#1028 := (iff #730 #1025) +#1019 := (and #1004 #1016) +#1026 := (iff #1019 #1025) +#1027 := [rewrite]: #1026 +#1020 := (iff #730 #1019) +#1017 := (iff #729 #1016) +#1014 := (iff #728 #1013) +#1015 := [rewrite]: #1014 +#1018 := [monotonicity #1015]: #1017 +#1005 := (iff #727 #1004) +#1002 := (iff #726 #1001) +#1003 := [rewrite]: #1002 +#1006 := [monotonicity #1003]: #1005 +#1021 := [monotonicity #1006 #1018]: #1020 +#1029 := [trans #1021 #1027]: #1028 +#731 := [intro-def]: #730 +#1030 := [mp #731 #1029]: #1025 +#1032 := [not-or-elim #1030]: #1016 +#1400 := [unit-resolution #1032 #1399]: #241 +#1031 := [not-or-elim #1030]: #1004 +#1401 := [unit-resolution #1031 #1400]: #1001 +#1261 := (or #1000 #997) +#1262 := [def-axiom]: #1261 +#1402 := [unit-resolution #1262 #1401]: #997 +#1403 := [th-lemma #1400 #1402 #1398]: false +#1404 := [lemma #1403]: #1009 +#737 := (+ uf_6 #734) +#740 := (+ uf_4 #737) +#747 := (>= #740 0::int) +#248 := (ite #241 uf_5 #224) +#259 := (* -1::int #248) +#260 := (+ uf_6 #259) +#261 := (+ uf_4 #260) +#566 := (>= #261 0::int) +#748 := (= #566 #747) +#741 := (~ #261 #740) +#738 := (~ #260 #737) +#735 := (~ #259 #734) +#732 := (~ #248 z3name!3) +#733 := [apply-def #731]: #732 +#736 := [monotonicity #733]: #735 +#739 := [monotonicity #736]: #738 +#742 := [monotonicity #739]: #741 +#749 := [monotonicity #742]: #748 +#568 := (not #566) +#565 := (<= #261 0::int) +#567 := (not #565) +#569 := (or #567 #568) +#570 := (not #569) +#262 := (= #261 0::int) +#571 := (iff #262 #570) +#572 := [rewrite]: #571 +#27 := (- uf_5) +#26 := (< uf_5 0::int) +#28 := (ite #26 #27 uf_5) +#29 := (- #28 uf_4) +#30 := (= uf_6 #29) +#267 := (iff #30 #262) +#227 := (ite #26 #224 uf_5) +#177 := (* -1::int uf_4) +#233 := (+ #177 #227) +#238 := (= uf_6 #233) +#265 := (iff #238 #262) +#253 := (+ #177 #248) +#256 := (= uf_6 #253) +#263 := (iff #256 #262) +#264 := [rewrite]: #263 +#257 := (iff #238 #256) +#254 := (= #233 #253) +#251 := (= #227 #248) +#245 := (ite #242 #224 uf_5) +#249 := (= #245 #248) +#250 := [rewrite]: #249 +#246 := (= #227 #245) +#243 := (iff #26 #242) +#244 := [rewrite]: #243 +#247 := [monotonicity #244]: #246 +#252 := [trans #247 #250]: #251 +#255 := [monotonicity #252]: #254 +#258 := [monotonicity #255]: #257 +#266 := [trans #258 #264]: #265 +#239 := (iff #30 #238) +#236 := (= #29 #233) +#230 := (- #227 uf_4) +#234 := (= #230 #233) +#235 := [rewrite]: #234 +#231 := (= #29 #230) +#228 := (= #28 #227) +#225 := (= #27 #224) +#226 := [rewrite]: #225 +#229 := [monotonicity #226]: #228 +#232 := [monotonicity #229]: #231 +#237 := [trans #232 #235]: #236 +#240 := [monotonicity #237]: #239 +#268 := [trans #240 #266]: #267 +#223 := [asserted]: #30 +#269 := [mp #223 #268]: #262 +#573 := [mp #269 #572]: #570 +#575 := [not-or-elim #573]: #566 +#750 := [mp~ #575 #749]: #747 +#1364 := (not #747) +#1357 := (not #769) +#1337 := (not #851) +#1555 := (or #194 #476 #1151 #1337 #1343 #1113 #1048 #1357 #1364 #1011) +#1556 := [th-lemma]: #1555 +#1557 := [unit-resolution #1556 #1545 #750 #1404 #772 #1496 #828 #1554 #854 #1553]: #194 +#195 := (not #194) +#966 := (or #195 #963) +#969 := (+ uf_4 z3name!2) +#971 := (>= #969 0::int) +#973 := (not #971) +#970 := (<= #969 0::int) +#972 := (not #970) +#974 := (or #972 #973) +#975 := (not #974) +#978 := (or #194 #975) +#985 := (not #978) +#984 := (not #966) +#986 := (or #984 #985) +#987 := (not #986) +#702 := (= z3name!2 #177) +#703 := (or #194 #702) +#700 := (= z3name!2 uf_4) +#701 := (or #195 #700) +#704 := (and #701 #703) +#990 := (iff #704 #987) +#981 := (and #966 #978) +#988 := (iff #981 #987) +#989 := [rewrite]: #988 +#982 := (iff #704 #981) +#979 := (iff #703 #978) +#976 := (iff #702 #975) +#977 := [rewrite]: #976 +#980 := [monotonicity #977]: #979 +#967 := (iff #701 #966) +#964 := (iff #700 #963) +#965 := [rewrite]: #964 +#968 := [monotonicity #965]: #967 +#983 := [monotonicity #968 #980]: #982 +#991 := [trans #983 #989]: #990 +#705 := [intro-def]: #704 +#992 := [mp #705 #991]: #987 +#993 := [not-or-elim #992]: #966 +#1558 := [unit-resolution #993 #1557]: #963 +#1249 := (or #962 #959) +#1250 := [def-axiom]: #1249 +#1559 := [unit-resolution #1250 #1558]: #959 +decl z3name!8 :: int +#855 := z3name!8 +#864 := (* -1::int z3name!8) +#867 := (+ uf_11 #864) +#870 := (+ uf_9 #867) +#873 := (<= #870 0::int) +#483 := (ite #476 uf_10 #459) +#494 := (* -1::int #483) +#495 := (+ uf_11 #494) +#496 := (+ uf_9 #495) +#620 := (<= #496 0::int) +#874 := (= #620 #873) +#871 := (~ #496 #870) +#868 := (~ #495 #867) +#865 := (~ #494 #864) +#862 := (~ #483 z3name!8) +#858 := (= z3name!8 #459) +#859 := (or #476 #858) +#856 := (= z3name!8 uf_10) +#857 := (or #477 #856) +#860 := (and #857 #859) +#861 := [intro-def]: #860 +#863 := [apply-def #861]: #862 +#866 := [monotonicity #863]: #865 +#869 := [monotonicity #866]: #868 +#872 := [monotonicity #869]: #871 +#875 := [monotonicity #872]: #874 +#621 := (>= #496 0::int) +#623 := (not #621) +#622 := (not #620) +#624 := (or #622 #623) +#625 := (not #624) +#497 := (= #496 0::int) +#626 := (iff #497 #625) +#627 := [rewrite]: #626 +#57 := (- uf_10) +#56 := (< uf_10 0::int) +#58 := (ite #56 #57 uf_10) +#59 := (- #58 uf_9) +#60 := (= uf_11 #59) +#502 := (iff #60 #497) +#462 := (ite #56 #459 uf_10) +#468 := (+ #412 #462) +#473 := (= uf_11 #468) +#500 := (iff #473 #497) +#488 := (+ #412 #483) +#491 := (= uf_11 #488) +#498 := (iff #491 #497) +#499 := [rewrite]: #498 +#492 := (iff #473 #491) +#489 := (= #468 #488) +#486 := (= #462 #483) +#480 := (ite #477 #459 uf_10) +#484 := (= #480 #483) +#485 := [rewrite]: #484 +#481 := (= #462 #480) +#478 := (iff #56 #477) +#479 := [rewrite]: #478 +#482 := [monotonicity #479]: #481 +#487 := [trans #482 #485]: #486 +#490 := [monotonicity #487]: #489 +#493 := [monotonicity #490]: #492 +#501 := [trans #493 #499]: #500 +#474 := (iff #60 #473) +#471 := (= #59 #468) +#465 := (- #462 uf_9) +#469 := (= #465 #468) +#470 := [rewrite]: #469 +#466 := (= #59 #465) +#463 := (= #58 #462) +#460 := (= #57 #459) +#461 := [rewrite]: #460 +#464 := [monotonicity #461]: #463 +#467 := [monotonicity #464]: #466 +#472 := [trans #467 #470]: #471 +#475 := [monotonicity #472]: #474 +#503 := [trans #475 #501]: #502 +#458 := [asserted]: #60 +#504 := [mp #458 #503]: #497 +#628 := [mp #504 #627]: #625 +#629 := [not-or-elim #628]: #620 +#876 := [mp~ #629 #875]: #873 +#1197 := (+ uf_10 z3name!8) +#1198 := (<= #1197 0::int) +#1199 := (>= #1197 0::int) +#1201 := (not #1199) +#1200 := (not #1198) +#1202 := (or #1200 #1201) +#1203 := (not #1202) +#1206 := (or #476 #1203) +#1213 := (not #1206) +#1186 := (+ uf_10 #864) +#1187 := (<= #1186 0::int) +#1189 := (not #1187) +#1185 := (>= #1186 0::int) +#1188 := (not #1185) +#1190 := (or #1188 #1189) +#1191 := (not #1190) +#1194 := (or #477 #1191) +#1212 := (not #1194) +#1214 := (or #1212 #1213) +#1215 := (not #1214) +#1218 := (iff #860 #1215) +#1209 := (and #1194 #1206) +#1216 := (iff #1209 #1215) +#1217 := [rewrite]: #1216 +#1210 := (iff #860 #1209) +#1207 := (iff #859 #1206) +#1204 := (iff #858 #1203) +#1205 := [rewrite]: #1204 +#1208 := [monotonicity #1205]: #1207 +#1195 := (iff #857 #1194) +#1192 := (iff #856 #1191) +#1193 := [rewrite]: #1192 +#1196 := [monotonicity #1193]: #1195 +#1211 := [monotonicity #1196 #1208]: #1210 +#1219 := [trans #1211 #1217]: #1218 +#1220 := [mp #861 #1219]: #1215 +#1222 := [not-or-elim #1220]: #1206 +#1560 := [unit-resolution #1222 #1553]: #1203 +#1325 := (or #1202 #1198) +#1326 := [def-axiom]: #1325 +#1561 := [unit-resolution #1326 #1560]: #1198 +#711 := (+ uf_5 #708) +#714 := (+ uf_1 #711) +#721 := (>= #714 0::int) +#201 := (ite #194 uf_4 #177) +#212 := (* -1::int #201) +#213 := (+ uf_5 #212) +#214 := (+ uf_1 #213) +#555 := (>= #214 0::int) +#722 := (= #555 #721) +#715 := (~ #214 #714) +#712 := (~ #213 #711) +#709 := (~ #212 #708) +#706 := (~ #201 z3name!2) +#707 := [apply-def #705]: #706 +#710 := [monotonicity #707]: #709 +#713 := [monotonicity #710]: #712 +#716 := [monotonicity #713]: #715 +#723 := [monotonicity #716]: #722 +#557 := (not #555) +#554 := (<= #214 0::int) +#556 := (not #554) +#558 := (or #556 #557) +#559 := (not #558) +#215 := (= #214 0::int) +#560 := (iff #215 #559) +#561 := [rewrite]: #560 +#21 := (- uf_4) +#20 := (< uf_4 0::int) +#22 := (ite #20 #21 uf_4) +#23 := (- #22 uf_1) +#24 := (= uf_5 #23) +#220 := (iff #24 #215) +#180 := (ite #20 #177 uf_4) +#186 := (+ #130 #180) +#191 := (= uf_5 #186) +#218 := (iff #191 #215) +#206 := (+ #130 #201) +#209 := (= uf_5 #206) +#216 := (iff #209 #215) +#217 := [rewrite]: #216 +#210 := (iff #191 #209) +#207 := (= #186 #206) +#204 := (= #180 #201) +#198 := (ite #195 #177 uf_4) +#202 := (= #198 #201) +#203 := [rewrite]: #202 +#199 := (= #180 #198) +#196 := (iff #20 #195) +#197 := [rewrite]: #196 +#200 := [monotonicity #197]: #199 +#205 := [trans #200 #203]: #204 +#208 := [monotonicity #205]: #207 +#211 := [monotonicity #208]: #210 +#219 := [trans #211 #217]: #218 +#192 := (iff #24 #191) +#189 := (= #23 #186) +#183 := (- #180 uf_1) +#187 := (= #183 #186) +#188 := [rewrite]: #187 +#184 := (= #23 #183) +#181 := (= #22 #180) +#178 := (= #21 #177) +#179 := [rewrite]: #178 +#182 := [monotonicity #179]: #181 +#185 := [monotonicity #182]: #184 +#190 := [trans #185 #188]: #189 +#193 := [monotonicity #190]: #192 +#221 := [trans #193 #219]: #220 +#176 := [asserted]: #24 +#222 := [mp #176 #221]: #215 +#562 := [mp #222 #561]: #559 +#564 := [not-or-elim #562]: #555 +#724 := [mp~ #564 #723]: #721 +#685 := (+ uf_4 #682) +#688 := (+ uf_2 #685) +#695 := (>= #688 0::int) +#154 := (ite #147 uf_1 #130) +#165 := (* -1::int #154) +#166 := (+ uf_4 #165) +#167 := (+ uf_2 #166) +#544 := (>= #167 0::int) +#696 := (= #544 #695) +#689 := (~ #167 #688) +#686 := (~ #166 #685) +#683 := (~ #165 #682) +#680 := (~ #154 z3name!1) +#681 := [apply-def #679]: #680 +#684 := [monotonicity #681]: #683 +#687 := [monotonicity #684]: #686 +#690 := [monotonicity #687]: #689 +#697 := [monotonicity #690]: #696 +#546 := (not #544) +#543 := (<= #167 0::int) +#545 := (not #543) +#547 := (or #545 #546) +#548 := (not #547) +#168 := (= #167 0::int) +#549 := (iff #168 #548) +#550 := [rewrite]: #549 +#15 := (- uf_1) +#14 := (< uf_1 0::int) +#16 := (ite #14 #15 uf_1) +#17 := (- #16 uf_2) +#18 := (= uf_4 #17) +#173 := (iff #18 #168) +#133 := (ite #14 #130 uf_1) +#139 := (+ #82 #133) +#144 := (= uf_4 #139) +#171 := (iff #144 #168) +#159 := (+ #82 #154) +#162 := (= uf_4 #159) +#169 := (iff #162 #168) +#170 := [rewrite]: #169 +#163 := (iff #144 #162) +#160 := (= #139 #159) +#157 := (= #133 #154) +#151 := (ite #148 #130 uf_1) +#155 := (= #151 #154) +#156 := [rewrite]: #155 +#152 := (= #133 #151) +#149 := (iff #14 #148) +#150 := [rewrite]: #149 +#153 := [monotonicity #150]: #152 +#158 := [trans #153 #156]: #157 +#161 := [monotonicity #158]: #160 +#164 := [monotonicity #161]: #163 +#172 := [trans #164 #170]: #171 +#145 := (iff #18 #144) +#142 := (= #17 #139) +#136 := (- #133 uf_2) +#140 := (= #136 #139) +#141 := [rewrite]: #140 +#137 := (= #17 #136) +#134 := (= #16 #133) +#131 := (= #15 #130) +#132 := [rewrite]: #131 +#135 := [monotonicity #132]: #134 +#138 := [monotonicity #135]: #137 +#143 := [trans #138 #141]: #142 +#146 := [monotonicity #143]: #145 +#174 := [trans #146 #172]: #173 +#129 := [asserted]: #18 +#175 := [mp #129 #174]: #168 +#551 := [mp #175 #550]: #548 +#553 := [not-or-elim #551]: #544 +#698 := [mp~ #553 #697]: #695 +#1373 := (not #721) +#1562 := (or #147 #1373 #961 #241 #195) +#1563 := [th-lemma]: #1562 +#1564 := [unit-resolution #1563 #1559 #1557 #724 #1406]: #147 +#1565 := [unit-resolution #955 #1564]: #925 +#1566 := [unit-resolution #1238 #1565]: #921 +#1372 := (not #873) +#1371 := (not #695) +#1498 := (or #516 #923 #1373 #1371 #1372 #1343 #1200 #1075 #1350 #1113 #961 #1151 #1337 #1048 #1357) +#1499 := [th-lemma]: #1498 +#1567 := [unit-resolution #1499 #1566 #698 #724 #1545 #772 #1447 #802 #1496 #828 #1554 #854 #1561 #876 #1559]: #516 +#1247 := (or #962 #957) +#1248 := [def-axiom]: #1247 +#1568 := [unit-resolution #1248 #1558]: #957 +#877 := (>= #870 0::int) +#878 := (= #621 #877) +#879 := [monotonicity #872]: #878 +#630 := [not-or-elim #628]: #621 +#880 := [mp~ #630 #879]: #877 +#1327 := (or #1202 #1199) +#1328 := [def-axiom]: #1327 +#1569 := [unit-resolution #1328 #1560]: #1199 +#795 := (<= #792 0::int) +#796 := (= #587 #795) +#797 := [monotonicity #794]: #796 +#596 := [not-or-elim #595]: #587 +#798 := [mp~ #596 #797]: #795 +#1503 := (or #335 #1049 #241) +#1425 := (or #335 #1049 #241 #1037) +#1422 := [hypothesis]: #336 +#1423 := [hypothesis]: #1047 +#1424 := [th-lemma #1423 #1422 #776 #1406 #1421]: false +#1426 := [lemma #1424]: #1425 +#1504 := [unit-resolution #1426 #1459]: #1503 +#1505 := [unit-resolution #1504 #1406 #1468]: #335 +#1506 := [unit-resolution #1107 #1505]: #1077 +#1283 := (or #1076 #1071) +#1284 := [def-axiom]: #1283 +#1507 := [unit-resolution #1284 #1506]: #1071 +#717 := (<= #714 0::int) +#718 := (= #554 #717) +#719 := [monotonicity #716]: #718 +#563 := [not-or-elim #562]: #554 +#720 := [mp~ #563 #719]: #717 +#691 := (<= #688 0::int) +#692 := (= #543 #691) +#693 := [monotonicity #690]: #692 +#552 := [not-or-elim #551]: #543 +#694 := [mp~ #552 #693]: #691 +#1235 := (or #924 #919) +#1236 := [def-axiom]: #1235 +#1570 := [unit-resolution #1236 #1565]: #919 +#1409 := (not #773) +#1489 := (not #847) +#1358 := (not #795) +#1365 := (not #821) +#1511 := (not #877) +#1510 := (not #691) +#1509 := (not #717) +#1512 := (or #515 #922 #1509 #1510 #1511 #1365 #1201 #1074 #1358 #1112 #960 #1150 #1489 #1049 #1409) +#1513 := [th-lemma]: #1512 +#1571 := [unit-resolution #1513 #1570 #694 #720 #1468 #776 #1507 #798 #1488 #824 #1552 #850 #1569 #880 #1568]: #515 +#506 := (<= #508 0::int) +#659 := (+ uf_3 #656) +#662 := (+ uf_1 #659) +#665 := (<= #662 0::int) +#107 := (ite #100 uf_2 #82) +#118 := (* -1::int #107) +#119 := (+ uf_3 #118) +#120 := (+ uf_1 #119) +#532 := (<= #120 0::int) +#666 := (= #532 #665) +#663 := (~ #120 #662) +#660 := (~ #119 #659) +#657 := (~ #118 #656) +#654 := (~ #107 z3name!0) +#655 := [apply-def #653]: #654 +#658 := [monotonicity #655]: #657 +#661 := [monotonicity #658]: #660 +#664 := [monotonicity #661]: #663 +#667 := [monotonicity #664]: #666 +#533 := (>= #120 0::int) +#535 := (not #533) +#534 := (not #532) +#536 := (or #534 #535) +#537 := (not #536) +#121 := (= #120 0::int) +#538 := (iff #121 #537) +#539 := [rewrite]: #538 +#8 := (- uf_2) +#7 := (< uf_2 0::int) +#9 := (ite #7 #8 uf_2) +#11 := (- #9 uf_3) +#12 := (= uf_1 #11) +#126 := (iff #12 #121) +#85 := (ite #7 #82 uf_2) +#91 := (* -1::int uf_3) +#92 := (+ #91 #85) +#97 := (= uf_1 #92) +#124 := (iff #97 #121) +#112 := (+ #91 #107) +#115 := (= uf_1 #112) +#122 := (iff #115 #121) +#123 := [rewrite]: #122 +#116 := (iff #97 #115) +#113 := (= #92 #112) +#110 := (= #85 #107) +#104 := (ite #101 #82 uf_2) +#108 := (= #104 #107) +#109 := [rewrite]: #108 +#105 := (= #85 #104) +#102 := (iff #7 #101) +#103 := [rewrite]: #102 +#106 := [monotonicity #103]: #105 +#111 := [trans #106 #109]: #110 +#114 := [monotonicity #111]: #113 +#117 := [monotonicity #114]: #116 +#125 := [trans #117 #123]: #124 +#98 := (iff #12 #97) +#95 := (= #11 #92) +#88 := (- #85 uf_3) +#93 := (= #88 #92) +#94 := [rewrite]: #93 +#89 := (= #11 #88) +#86 := (= #9 #85) +#83 := (= #8 #82) +#84 := [rewrite]: #83 +#87 := [monotonicity #84]: #86 +#90 := [monotonicity #87]: #89 +#96 := [trans #90 #94]: #95 +#99 := [monotonicity #96]: #98 +#127 := [trans #99 #125]: #126 +#80 := [asserted]: #12 +#128 := [mp #80 #127]: #121 +#540 := [mp #128 #539]: #537 +#541 := [not-or-elim #540]: #532 +#668 := [mp~ #541 #667]: #665 +#1515 := (or #100 #241 #923 #1373 #1371 #961) +#1516 := [th-lemma]: #1515 +#1572 := [unit-resolution #1516 #1566 #698 #1559 #724 #1406]: #100 +#1573 := [unit-resolution #917 #1572]: #887 +#1223 := (or #886 #881) +#1224 := [def-axiom]: #1223 +#1574 := [unit-resolution #1224 #1573]: #881 +#1528 := (not #665) +#1529 := (or #506 #884 #1528 #1364 #1011 #1343 #1113 #1151 #1337 #1048 #1357 #922 #1510) +#1530 := [th-lemma]: #1529 +#1575 := [unit-resolution #1530 #1574 #668 #694 #1404 #750 #1545 #772 #1496 #828 #1554 #854 #1570]: #506 +#743 := (<= #740 0::int) +#744 := (= #565 #743) +#745 := [monotonicity #742]: #744 +#574 := [not-or-elim #573]: #565 +#746 := [mp~ #574 #745]: #743 +#1520 := [unit-resolution #1032 #1406]: #1013 +#1265 := (or #1012 #1008) +#1266 := [def-axiom]: #1265 +#1521 := [unit-resolution #1266 #1520]: #1008 +#669 := (>= #662 0::int) +#670 := (= #533 #669) +#671 := [monotonicity #664]: #670 +#542 := [not-or-elim #540]: #533 +#672 := [mp~ #542 #671]: #669 +#1576 := [unit-resolution #1226 #1573]: #883 +#1523 := (not #743) +#1522 := (not #669) +#1524 := (or #509 #885 #1522 #1523 #1010 #1365 #1112 #1150 #1489 #1049 #1409 #923 #1371) +#1525 := [th-lemma]: #1524 +#1577 := [unit-resolution #1525 #1576 #672 #698 #1521 #746 #1468 #776 #1488 #824 #1552 #850 #1566]: #509 +#634 := (not #516) +#633 := (not #515) +#632 := (not #509) +#631 := (not #506) +#635 := (or #631 #632 #633 #634) +#523 := (and #506 #509 #515 #516) +#528 := (not #523) +#644 := (iff #528 #635) +#636 := (not #635) +#639 := (not #636) +#642 := (iff #639 #635) +#643 := [rewrite]: #642 +#640 := (iff #528 #639) +#637 := (iff #523 #636) +#638 := [rewrite]: #637 +#641 := [monotonicity #638]: #640 +#645 := [trans #641 #643]: #644 +#62 := (= uf_2 uf_11) +#61 := (= uf_3 uf_10) +#63 := (and #61 #62) +#64 := (not #63) +#529 := (iff #64 #528) +#526 := (iff #63 #523) +#517 := (and #515 #516) +#510 := (and #506 #509) +#520 := (and #510 #517) +#524 := (iff #520 #523) +#525 := [rewrite]: #524 +#521 := (iff #63 #520) +#518 := (iff #62 #517) +#519 := [rewrite]: #518 +#511 := (iff #61 #510) +#512 := [rewrite]: #511 +#522 := [monotonicity #512 #519]: #521 +#527 := [trans #522 #525]: #526 +#530 := [monotonicity #527]: #529 +#505 := [asserted]: #64 +#531 := [mp #505 #530]: #528 +#646 := [mp #531 #645]: #635 +#1578 := [unit-resolution #646 #1577 #1575 #1571 #1567]: false +#1580 := [lemma #1578]: #1579 +#1657 := [unit-resolution #1580 #1406]: #1048 +#1625 := (or #194 #241) +#1535 := [hypothesis]: #195 +#1538 := (or #194 #960) +#1432 := [hypothesis]: #973 +#1255 := (or #974 #971) +#1256 := [def-axiom]: #1255 +#1433 := [unit-resolution #1256 #1432]: #974 +#994 := [not-or-elim #992]: #978 +#1434 := [unit-resolution #994 #1433]: #194 +#1435 := [unit-resolution #993 #1434]: #963 +#1436 := (or #971 #195 #961) +#1437 := [th-lemma]: #1436 +#1438 := [unit-resolution #1437 #1434 #1432]: #961 +#1439 := [unit-resolution #1250 #1438 #1435]: false +#1440 := [lemma #1439]: #971 +#1536 := [hypothesis]: #957 +#1537 := [th-lemma #1536 #1535 #1440]: false +#1539 := [lemma #1537]: #1538 +#1581 := [unit-resolution #1539 #1535]: #960 +#1582 := (or #959 #957) +#1583 := [th-lemma]: #1582 +#1584 := [unit-resolution #1583 #1581]: #959 +#1585 := (or #147 #1373 #241 #194 #973) +#1586 := [th-lemma]: #1585 +#1587 := [unit-resolution #1586 #1535 #1440 #724 #1406]: #147 +#1588 := [unit-resolution #955 #1587]: #925 +#1589 := [unit-resolution #1238 #1588]: #921 +#1590 := [unit-resolution #1516 #1589 #698 #1584 #724 #1406]: #100 +#1591 := [unit-resolution #917 #1590]: #887 +#1592 := [unit-resolution #1224 #1591]: #881 +#1593 := (or #430 #1365 #1074 #1358 #1112 #194 #1364 #1011 #241) +#1594 := [th-lemma]: #1593 +#1595 := [unit-resolution #1594 #1535 #1404 #750 #1507 #798 #1488 #824 #1406]: #430 +#1184 := [not-or-elim #1182]: #1168 +#1596 := [unit-resolution #1184 #1595]: #1165 +#1315 := (or #1164 #1161) +#1316 := [def-axiom]: #1315 +#1597 := [unit-resolution #1316 #1596]: #1161 +#1533 := (or #288 #241) +#1471 := (or #194 #288 #241) +#1469 := (or #194 #288 #241 #1364 #1011) +#1470 := [th-lemma]: #1469 +#1472 := [unit-resolution #1470 #1404 #750]: #1471 +#1473 := [unit-resolution #1472 #1405 #1406]: #194 +#1474 := [unit-resolution #993 #1473]: #963 +#1475 := [unit-resolution #1250 #1474]: #959 +#1476 := (or #147 #1373 #1364 #1011 #961 #241 #288) +#1477 := [th-lemma]: #1476 +#1478 := [unit-resolution #1477 #1475 #724 #1406 #1404 #750 #1405]: #147 +#1479 := [unit-resolution #955 #1478]: #925 +#1480 := [unit-resolution #1238 #1479]: #921 +#1419 := (or #288 #241 #429) +#1333 := [hypothesis]: #430 +#1408 := [unit-resolution #1280 #1407]: #1047 +#1410 := (or #335 #1049 #1409 #288 #241) +#1411 := [th-lemma]: #1410 +#1412 := [unit-resolution #1411 #1405 #1408 #776 #1406]: #335 +#1413 := [unit-resolution #1107 #1412]: #1077 +#1414 := [unit-resolution #1286 #1413]: #1073 +#1415 := [unit-resolution #1352 #1414 #802 #1405 #828 #1333]: #1113 +#1416 := [unit-resolution #1298 #1415]: #1114 +#1417 := [unit-resolution #1145 #1416]: #383 +#1418 := [th-lemma #1414 #802 #1405 #1408 #776 #1406 #1417]: false +#1420 := [lemma #1418]: #1419 +#1481 := [unit-resolution #1420 #1405 #1406]: #429 +#1482 := [unit-resolution #1183 #1481]: #1153 +#1483 := [unit-resolution #1308 #1482]: #1147 +#1490 := (or #477 #1150 #1489 #1365 #1112 #1049 #241 #1409 #288) +#1491 := [th-lemma]: #1490 +#1492 := [unit-resolution #1491 #1405 #1468 #776 #1488 #824 #1483 #850 #1406]: #477 +#1493 := [unit-resolution #1222 #1492]: #1203 +#1494 := [unit-resolution #1326 #1493]: #1198 +#1495 := [unit-resolution #1310 #1482]: #1149 +#1500 := [unit-resolution #1499 #1475 #698 #724 #1497 #772 #1447 #802 #1496 #828 #1495 #854 #1494 #876 #1480]: #516 +#1501 := [unit-resolution #1236 #1479]: #919 +#1502 := [unit-resolution #1328 #1493]: #1199 +#1508 := [unit-resolution #1248 #1474]: #957 +#1514 := [unit-resolution #1513 #1508 #694 #720 #1468 #776 #1507 #798 #1488 #824 #1483 #850 #1502 #880 #1501]: #515 +#1517 := [unit-resolution #1516 #1480 #698 #1475 #724 #1406]: #100 +#1518 := [unit-resolution #917 #1517]: #887 +#1519 := [unit-resolution #1226 #1518]: #883 +#1526 := [unit-resolution #1525 #1480 #672 #698 #1521 #746 #1468 #776 #1488 #824 #1483 #850 #1519]: #509 +#1527 := [unit-resolution #1224 #1518]: #881 +#1531 := [unit-resolution #1530 #1501 #668 #694 #1404 #750 #1497 #772 #1496 #828 #1495 #854 #1527]: #506 +#1532 := [unit-resolution #646 #1531 #1526 #1514 #1500]: false +#1534 := [lemma #1532]: #1533 +#1598 := [unit-resolution #1534 #1406]: #288 +#1599 := [unit-resolution #1069 #1598]: #1039 +#1271 := (or #1038 #1033) +#1272 := [def-axiom]: #1271 +#1600 := [unit-resolution #1272 #1599]: #1033 +#1601 := [unit-resolution #1236 #1588]: #919 +#1602 := (or #506 #884 #1528 #1364 #1011 #1365 #1112 #1337 #1357 #922 #1510 #1036 #1163 #1074 #1358) +#1603 := [th-lemma]: #1602 +#1604 := [unit-resolution #1603 #1601 #668 #694 #1404 #750 #1600 #772 #1507 #798 #1488 #824 #1597 #854 #1592]: #506 +#1605 := [unit-resolution #1226 #1591]: #883 +#1313 := (or #1164 #1160) +#1314 := [def-axiom]: #1313 +#1606 := [unit-resolution #1314 #1596]: #1160 +#1607 := (or #509 #885 #1522 #1523 #1010 #1343 #1113 #1489 #1409 #923 #1371 #1037 #1162 #1075 #1350) +#1608 := [th-lemma]: #1607 +#1609 := [unit-resolution #1608 #1589 #672 #698 #1521 #746 #1459 #776 #1447 #802 #1496 #828 #1606 #850 #1605]: #509 +#1610 := (or #476 #1036 #1337 #1365 #1112 #1357 #194 #1364 #1011 #1163 #1074 #1358) +#1611 := [th-lemma]: #1610 +#1612 := [unit-resolution #1611 #1597 #750 #1600 #772 #1507 #798 #1488 #824 #1404 #854 #1535]: #476 +#1221 := [not-or-elim #1220]: #1194 +#1613 := [unit-resolution #1221 #1612]: #1191 +#1319 := (or #1190 #1185) +#1320 := [def-axiom]: #1319 +#1614 := [unit-resolution #1320 #1613]: #1185 +#1615 := (or #516 #923 #1373 #1371 #1372 #1075 #1350 #1489 #1409 #1037 #973 #1162 #1188 #1343 #1113 #1523 #1010) +#1616 := [th-lemma]: #1615 +#1617 := [unit-resolution #1616 #1606 #1440 #724 #1521 #746 #1459 #776 #1447 #802 #1496 #828 #698 #850 #1614 #876 #1589]: #516 +#1321 := (or #1190 #1187) +#1322 := [def-axiom]: #1321 +#1618 := [unit-resolution #1322 #1613]: #1187 +#1619 := [unit-resolution #994 #1535]: #975 +#1253 := (or #974 #970) +#1254 := [def-axiom]: #1253 +#1620 := [unit-resolution #1254 #1619]: #970 +#1621 := (or #515 #922 #1509 #1510 #1511 #1074 #1358 #1337 #1357 #1036 #972 #1163 #1189 #1365 #1112 #1364 #1011) +#1622 := [th-lemma]: #1621 +#1623 := [unit-resolution #1622 #1620 #694 #720 #1404 #750 #1600 #772 #1507 #798 #1488 #824 #1597 #854 #1618 #880 #1601]: #515 +#1624 := [unit-resolution #646 #1623 #1617 #1609 #1604]: false +#1626 := [lemma #1624]: #1625 +#1658 := [unit-resolution #1626 #1406]: #194 +#1659 := [unit-resolution #993 #1658]: #963 +#1660 := [unit-resolution #1250 #1659]: #959 +#1661 := [unit-resolution #1563 #1660 #1658 #724 #1406]: #147 +#1662 := [unit-resolution #955 #1661]: #925 +#1663 := [unit-resolution #1238 #1662]: #921 +#1664 := [unit-resolution #1516 #1663 #698 #1660 #724 #1406]: #100 +#1665 := [unit-resolution #917 #1664]: #887 +#1666 := [unit-resolution #1226 #1665]: #883 +#1667 := [unit-resolution #1224 #1665]: #881 +#1668 := [unit-resolution #1236 #1662]: #919 +#1669 := [unit-resolution #1248 #1659]: #957 +#1655 := (or #429 #1113 #1010 #960 #1036 #1074 #1112 #922 #923 #884 #885) +#1632 := [hypothesis]: #919 +#1636 := [hypothesis]: #881 +#1638 := [hypothesis]: #1071 +#1639 := [hypothesis]: #1033 +#1334 := [unit-resolution #1184 #1333]: #1165 +#1335 := [unit-resolution #1316 #1334]: #1161 +#1640 := [unit-resolution #1603 #1335 #668 #694 #1404 #750 #1639 #772 #1638 #798 #1637 #824 #1632 #854 #1636]: #506 +#1641 := [hypothesis]: #883 +#1642 := [hypothesis]: #921 +#1643 := [hypothesis]: #1111 +#1644 := [hypothesis]: #1008 +#1631 := [unit-resolution #1314 #1334]: #1160 +#1645 := [unit-resolution #1608 #1631 #672 #698 #1644 #746 #1459 #776 #1447 #802 #1643 #828 #1642 #850 #1641]: #509 +#1634 := (or #1202 #922 #960 #632 #631 #429) +#1627 := [hypothesis]: #506 +#1628 := [hypothesis]: #509 +#1384 := [hypothesis]: #1203 +#1396 := (or #1202 #516 #429) +#1331 := [hypothesis]: #634 +#1385 := [unit-resolution #1326 #1384]: #1198 +#1382 := (or #1189 #1200 #516 #429) +#1332 := [hypothesis]: #1198 +#1336 := [hypothesis]: #1187 +#1338 := (or #382 #1189 #1337 #429 #1163 #1200) +#1339 := [th-lemma]: #1338 +#1340 := [unit-resolution #1339 #1336 #1335 #854 #1333 #1332]: #382 +#1341 := [unit-resolution #1145 #1340]: #1115 +#1342 := [unit-resolution #1298 #1341]: #1111 +#1344 := (or #335 #1113 #429 #1343 #1189 #1337 #1163 #1200) +#1345 := [th-lemma]: #1344 +#1346 := [unit-resolution #1345 #1342 #828 #1333 #1335 #854 #1336 #1332]: #335 +#1347 := [unit-resolution #1107 #1346]: #1077 +#1348 := [unit-resolution #1284 #1347]: #1071 +#1349 := [unit-resolution #1286 #1347]: #1073 +#1353 := [unit-resolution #1352 #1349 #802 #1342 #828 #1333]: #288 +#1354 := [unit-resolution #1069 #1353]: #1039 +#1355 := [unit-resolution #1272 #1354]: #1033 +#1356 := [unit-resolution #1296 #1341]: #1109 +#1359 := (or #242 #1036 #1357 #429 #1189 #1337 #1163 #1200 #1074 #1358) +#1360 := [th-lemma]: #1359 +#1361 := [unit-resolution #1360 #1355 #772 #1348 #798 #1333 #1335 #854 #1336 #1332]: #242 +#1362 := [unit-resolution #1032 #1361]: #1013 +#1363 := [unit-resolution #1268 #1362]: #1009 +#1366 := (or #194 #1011 #1364 #1074 #1358 #1112 #1365 #1036 #1357 #1189 #1337 #1163 #1200) +#1367 := [th-lemma]: #1366 +#1368 := [unit-resolution #1367 #1363 #750 #1355 #772 #1348 #798 #1356 #824 #1335 #854 #1336 #1332]: #194 +#1369 := [unit-resolution #993 #1368]: #963 +#1370 := [unit-resolution #1250 #1369]: #959 +#1374 := (or #923 #1371 #516 #1372 #1200 #961 #1373 #1036 #1357 #1337 #1163 #1074 #1358) +#1375 := [th-lemma]: #1374 +#1376 := [unit-resolution #1375 #1370 #698 #724 #1355 #772 #1348 #798 #1335 #854 #1332 #876 #1331]: #923 +#1377 := (or #147 #195 #961 #1373 #1036 #1357 #429 #1189 #1337 #1163 #1200 #1074 #1358) +#1378 := [th-lemma]: #1377 +#1379 := [unit-resolution #1378 #1368 #1370 #724 #1355 #772 #1348 #798 #1333 #1335 #854 #1336 #1332]: #147 +#1380 := [unit-resolution #955 #1379]: #925 +#1381 := [unit-resolution #1238 #1380 #1376]: false +#1383 := [lemma #1381]: #1382 +#1386 := [unit-resolution #1383 #1385 #1331 #1333]: #1189 +#1387 := [unit-resolution #1322 #1386]: #1190 +#1388 := [unit-resolution #1328 #1384]: #1199 +#1389 := (or #1187 #1185) +#1390 := [th-lemma]: #1389 +#1391 := [unit-resolution #1390 #1386]: #1185 +#1392 := (or #476 #1188 #1201) +#1393 := [th-lemma]: #1392 +#1394 := [unit-resolution #1393 #1391 #1388]: #476 +#1395 := [unit-resolution #1221 #1394 #1387]: false +#1397 := [lemma #1395]: #1396 +#1629 := [unit-resolution #1397 #1384 #1333]: #516 +#1630 := [unit-resolution #646 #1629 #1628 #1627]: #633 +#1633 := [th-lemma #1632 #720 #694 #880 #1447 #802 #850 #776 #1459 #1631 #1536 #1388 #1630]: false +#1635 := [lemma #1633]: #1634 +#1646 := [unit-resolution #1635 #1645 #1536 #1632 #1640 #1333]: #1202 +#1647 := [unit-resolution #1222 #1646]: #476 +#1648 := [unit-resolution #1221 #1647]: #1191 +#1649 := [unit-resolution #1322 #1648]: #1187 +#1650 := [unit-resolution #1320 #1648]: #1185 +#1651 := [unit-resolution #1616 #1650 #1440 #724 #1644 #746 #1459 #776 #1447 #802 #1643 #828 #698 #850 #1631 #876 #1642]: #516 +#1652 := [unit-resolution #646 #1651 #1645 #1640]: #633 +#1653 := [unit-resolution #1622 #1652 #694 #720 #1404 #750 #1639 #772 #1638 #798 #1637 #824 #1335 #854 #1649 #880 #1632]: #972 +#1654 := [th-lemma #1459 #1647 #850 #828 #1643 #776 #746 #1644 #1631 #1447 #802 #1536 #1653]: false +#1656 := [lemma #1654]: #1655 +#1670 := [unit-resolution #1656 #1496 #1521 #1669 #1600 #1507 #1488 #1668 #1663 #1667 #1666]: #429 +#1671 := [th-lemma #1600 #1670 #824 #1507 #798 #1488 #1657]: false +#1672 := [lemma #1671]: #241 +#1683 := [unit-resolution #1031 #1672]: #1001 +#1703 := [unit-resolution #1262 #1683]: #997 +#1920 := (or #194 #242 #1364 #999 #288) +#1921 := [th-lemma]: #1920 +#1922 := [unit-resolution #1921 #1405 #1703 #750 #1672]: #194 +#1923 := [unit-resolution #993 #1922]: #963 +#1924 := [unit-resolution #1248 #1923]: #957 +#1925 := [unit-resolution #1250 #1923]: #959 +#1843 := (or #288 #961 #147) +#1763 := [hypothesis]: #148 +#1828 := [hypothesis]: #959 +#1842 := [th-lemma #724 #750 #1703 #1828 #1405 #1763]: false +#1844 := [lemma #1842]: #1843 +#1926 := [unit-resolution #1844 #1925 #1405]: #147 +#1927 := [unit-resolution #955 #1926]: #925 +#1928 := [unit-resolution #1236 #1927]: #919 +#2116 := [unit-resolution #1310 #2047]: #1149 +#2084 := (or #288 #516) +#2050 := (or #288 #961 #516) +#2037 := [hypothesis]: #1087 +#2038 := [unit-resolution #1292 #2037]: #1088 +#2041 := (or #1085 #336) +#2039 := (or #1085 #1075 #336) +#2040 := [th-lemma]: #2039 +#2042 := [unit-resolution #2040 #1447]: #2041 +#2043 := [unit-resolution #2042 #2037]: #336 +#2044 := [unit-resolution #1108 #2043 #2038]: false +#2045 := [lemma #2044]: #1085 +#2035 := (or #1087 #1150 #961 #1048 #516) +#1845 := [hypothesis]: #1085 +#1874 := [hypothesis]: #477 +#1901 := (or #335 #476) +#1895 := [unit-resolution #1222 #1874]: #1203 +#1896 := [unit-resolution #1326 #1895]: #1198 +#1893 := (or #429 #1200) +#1880 := (or #335 #1113 #429 #1163 #1200) +#1857 := [hypothesis]: #1189 +#1858 := [unit-resolution #1322 #1857]: #1190 +#1859 := [unit-resolution #1221 #1858]: #477 +#1860 := [unit-resolution #1222 #1859]: #1203 +#1861 := [unit-resolution #1390 #1857]: #1185 +#1862 := [unit-resolution #1393 #1859 #1861]: #1201 +#1863 := [unit-resolution #1328 #1862 #1860]: false +#1864 := [lemma #1863]: #1187 +#1878 := (or #335 #1113 #429 #1189 #1163 #1200) +#1879 := [unit-resolution #1345 #828 #854]: #1878 +#1881 := [unit-resolution #1879 #1864]: #1880 +#1882 := [unit-resolution #1881 #1335 #1870 #1333 #1332]: #335 +#1883 := [unit-resolution #1107 #1882]: #1077 +#1884 := [unit-resolution #1689 #1333 #1870]: #288 +#1885 := [unit-resolution #1069 #1884]: #1039 +#1886 := [unit-resolution #1272 #1885]: #1033 +#1889 := (or #1036 #429 #1163 #1200 #1074) +#1887 := (or #242 #1036 #429 #1189 #1163 #1200 #1074) +#1888 := [unit-resolution #1360 #772 #798 #854]: #1887 +#1890 := [unit-resolution #1888 #1672 #1864]: #1889 +#1891 := [unit-resolution #1890 #1886 #1332 #1333 #1335]: #1074 +#1892 := [unit-resolution #1284 #1891 #1883]: false +#1894 := [lemma #1892]: #1893 +#1897 := [unit-resolution #1894 #1896]: #429 +#1898 := [unit-resolution #1183 #1897]: #1153 +#1899 := [unit-resolution #1310 #1898]: #1149 +#1900 := [th-lemma #854 #1899 #1870 #828 #1422 #1874]: false +#1902 := [lemma #1900]: #1901 +#1950 := [unit-resolution #1902 #1874]: #335 +#1951 := [unit-resolution #1107 #1950]: #1077 +#1952 := [unit-resolution #1284 #1951]: #1071 +#1953 := [unit-resolution #1328 #1895]: #1199 +#1876 := (or #1109 #476) +#1673 := [hypothesis]: #1112 +#1760 := (or #429 #1109) +#1674 := [unit-resolution #1296 #1673]: #1114 +#1675 := [unit-resolution #1145 #1674]: #383 +#1676 := [unit-resolution #1146 #1675]: #1127 +#1677 := [unit-resolution #1304 #1676]: #1123 +#1687 := [unit-resolution #1686 #1673]: #1111 +#1743 := [unit-resolution #1689 #1333 #1687]: #288 +#1744 := [unit-resolution #1069 #1743]: #1039 +#1745 := [unit-resolution #1272 #1744]: #1033 +#1678 := (or #335 #1343 #429 #382 #1125) +#1679 := [th-lemma]: #1678 +#1746 := [unit-resolution #1679 #1333 #1675 #828 #1677]: #335 +#1747 := [unit-resolution #1107 #1746]: #1077 +#1748 := [unit-resolution #1284 #1747]: #1071 +#1259 := (or #1000 #995) +#1260 := [def-axiom]: #1259 +#1684 := [unit-resolution #1260 #1683]: #995 +#1693 := (or #147 #1373 #1343 #1074 #1358 #1523 #429 #973 #998 #1036 #1357 #1125) +#1694 := [th-lemma]: #1693 +#1749 := [unit-resolution #1694 #1745 #724 #1684 #746 #1440 #772 #1748 #798 #1677 #828 #1333]: #147 +#1750 := [unit-resolution #955 #1749]: #925 +#1751 := [unit-resolution #1238 #1750]: #921 +#1714 := (or #100 #923 #1373 #1371 #1343 #1523 #1074 #1358 #973 #429 #382 #1036 #1357 #998 #1125) +#1715 := [th-lemma]: #1714 +#1752 := [unit-resolution #1715 #1751 #698 #1440 #724 #1684 #746 #1675 #772 #1748 #798 #1745 #1677 #828 #1333]: #100 +#1753 := [unit-resolution #1236 #1750]: #919 +#1727 := (or #1109 #429 #972) +#1680 := [unit-resolution #1679 #1675 #1677 #828 #1333]: #335 +#1681 := [unit-resolution #1107 #1680]: #1077 +#1682 := [unit-resolution #1284 #1681]: #1071 +#1690 := [unit-resolution #1689 #1687 #1333]: #288 +#1691 := [unit-resolution #1069 #1690]: #1039 +#1692 := [unit-resolution #1272 #1691]: #1033 +#1695 := [unit-resolution #1694 #1692 #724 #1684 #746 #1440 #772 #1682 #798 #1677 #828 #1333]: #147 +#1696 := [unit-resolution #955 #1695]: #925 +#1697 := [unit-resolution #1236 #1696]: #919 +#1698 := (or #476 #429 #1337 #1163 #382) +#1699 := [th-lemma]: #1698 +#1700 := [unit-resolution #1699 #1675 #1335 #854 #1333]: #476 +#1701 := [unit-resolution #1221 #1700]: #1191 +#1702 := [unit-resolution #1322 #1701]: #1187 +#1704 := [hypothesis]: #970 +#1301 := (or #1126 #1122) +#1302 := [def-axiom]: #1301 +#1705 := [unit-resolution #1302 #1676]: #1122 +#1706 := (or #515 #922 #1509 #1510 #1511 #1075 #1350 #1337 #1409 #1037 #1163 #1365 #1364 #972 #999 #1124 #1189) +#1707 := [th-lemma]: #1706 +#1708 := [unit-resolution #1707 #1705 #1704 #720 #1703 #750 #1459 #776 #1447 #802 #694 #824 #1335 #854 #1702 #880 #1697]: #515 +#1709 := [unit-resolution #1238 #1696]: #921 +#1710 := [unit-resolution #1320 #1701]: #1185 +#1711 := (or #516 #923 #1373 #1371 #1372 #1074 #1358 #1489 #1357 #1036 #1162 #1343 #1523 #973 #998 #1125 #1188) +#1712 := [th-lemma]: #1711 +#1713 := [unit-resolution #1712 #1692 #1440 #724 #1684 #746 #698 #772 #1682 #798 #1677 #828 #1631 #850 #1710 #876 #1709]: #516 +#1716 := [unit-resolution #1715 #1709 #698 #1440 #724 #1684 #746 #1692 #772 #1682 #798 #1675 #1677 #828 #1333]: #100 +#1717 := [unit-resolution #917 #1716]: #887 +#1718 := [unit-resolution #1226 #1717]: #883 +#1719 := (or #509 #885 #1522 #1523 #1343 #1489 #1357 #923 #1371 #1036 #1162 #998 #1125) +#1720 := [th-lemma]: #1719 +#1721 := [unit-resolution #1720 #1709 #672 #698 #1684 #746 #1692 #772 #1677 #828 #1631 #850 #1718]: #509 +#1722 := [unit-resolution #1224 #1717]: #881 +#1723 := (or #506 #884 #1528 #1364 #1365 #1337 #1409 #922 #1510 #1037 #1163 #999 #1124) +#1724 := [th-lemma]: #1723 +#1725 := [unit-resolution #1724 #1697 #668 #694 #1703 #750 #1459 #776 #1705 #824 #1335 #854 #1722]: #506 +#1726 := [unit-resolution #646 #1725 #1721 #1713 #1708]: false +#1728 := [lemma #1726]: #1727 +#1754 := [unit-resolution #1728 #1333 #1673]: #972 +#1755 := [unit-resolution #1254 #1754]: #974 +#1756 := [unit-resolution #994 #1755]: #194 +#1757 := [unit-resolution #993 #1756]: #963 +#1758 := [unit-resolution #1248 #1757]: #957 +#1759 := [th-lemma #1758 #1753 #720 #694 #1675 #1459 #776 #1447 #802 #1752]: false +#1761 := [lemma #1759]: #1760 +#1871 := [unit-resolution #1761 #1673]: #429 +#1872 := [unit-resolution #1183 #1871]: #1153 +#1873 := [unit-resolution #1310 #1872]: #1149 +#1875 := [th-lemma #1675 #1874 #854 #1873 #1871]: false +#1877 := [lemma #1875]: #1876 +#1954 := [unit-resolution #1877 #1874]: #1109 +#1948 := (or #288 #1112 #1200 #1201 #1074) +#1917 := [unit-resolution #1894 #1332]: #429 +#1918 := [unit-resolution #1183 #1917]: #1153 +#1919 := [unit-resolution #1308 #1918]: #1147 +#1929 := [unit-resolution #1310 #1918]: #1149 +#1930 := [unit-resolution #1238 #1927]: #921 +#1931 := [hypothesis]: #1199 +#1932 := (or #515 #922 #1201 #1074 #1112 #960 #1150) +#1933 := [unit-resolution #1513 #694 #720 #1468 #776 #798 #824 #850 #880]: #1932 +#1934 := [unit-resolution #1933 #1928 #1931 #1637 #1638 #1919 #1924]: #515 +#1935 := (or #516 #923 #1200 #1113 #961 #1151 #1048) +#1936 := [unit-resolution #1499 #698 #724 #772 #1447 #802 #828 #854 #876]: #1935 +#1937 := [unit-resolution #1936 #1930 #1870 #1332 #1929 #1497 #1925]: #516 +#1915 := (or #898 #634 #633 #923 #961 #1048 #1151 #922 #960 #1112 #1150) +#1903 := [hypothesis]: #515 +#1904 := [hypothesis]: #516 +#1905 := [hypothesis]: #899 +#1906 := [unit-resolution #1232 #1905]: #895 +#1907 := (or #509 #1522 #1523 #897 #998 #1489 #1150 #960 #1509 #1112 #1365 #1049 #922 #1510 #1409) +#1908 := [th-lemma]: #1907 +#1909 := [unit-resolution #1908 #1906 #1632 #694 #1536 #720 #1684 #746 #1468 #776 #1637 #824 #1540 #850 #672]: #509 +#1774 := [hypothesis]: #1149 +#1229 := (or #898 #894) +#1230 := [def-axiom]: #1229 +#1910 := [unit-resolution #1230 #1905]: #894 +#1911 := (or #506 #1528 #1364 #896 #999 #1337 #1151 #961 #1373 #1113 #1343 #1048 #923 #1371 #1357) +#1912 := [th-lemma]: #1911 +#1913 := [unit-resolution #1912 #1910 #1642 #698 #1828 #724 #1703 #750 #1545 #772 #1870 #828 #1774 #854 #668]: #506 +#1914 := [unit-resolution #646 #1913 #1909 #1904 #1903]: false +#1916 := [lemma #1914]: #1915 +#1938 := [unit-resolution #1916 #1937 #1934 #1930 #1925 #1497 #1929 #1928 #1924 #1637 #1919]: #898 +#1939 := [unit-resolution #918 #1938]: #100 +#1940 := [unit-resolution #917 #1939]: #887 +#1941 := [unit-resolution #1224 #1940]: #881 +#1942 := (or #506 #884 #1113 #1151 #1048 #922) +#1943 := [unit-resolution #1530 #668 #694 #1404 #750 #772 #828 #854]: #1942 +#1944 := [unit-resolution #1943 #1941 #1497 #1870 #1929 #1928]: #506 +#1945 := [unit-resolution #646 #1944 #1937 #1934]: #632 +#1946 := [unit-resolution #1908 #1945 #1928 #694 #1924 #720 #1684 #746 #1468 #776 #1637 #824 #1919 #850 #672]: #897 +#1947 := [th-lemma #1946 #1939 #1742]: false +#1949 := [lemma #1947]: #1948 +#1955 := [unit-resolution #1949 #1954 #1896 #1953 #1952]: #288 +#1956 := [unit-resolution #1069 #1955]: #1039 +#1957 := [unit-resolution #1272 #1956]: #1033 +#1958 := [unit-resolution #1735 #1954]: #382 +#1959 := (or #1123 #383 #1113) +#1960 := [th-lemma]: #1959 +#1961 := [unit-resolution #1960 #1958 #1870]: #1123 +#1962 := [unit-resolution #1308 #1898]: #1147 +#1965 := (or #1160 #1112 #1074 #289 #1150) +#1963 := (or #1160 #1365 #1112 #1074 #1358 #289 #1150) +#1964 := [th-lemma]: #1963 +#1966 := [unit-resolution #1964 #798 #824]: #1965 +#1967 := [unit-resolution #1966 #1955 #1954 #1962 #1952]: #1160 +#1970 := (or #1162 #1151 #1036 #1125 #147 #1074) +#1968 := (or #1162 #1151 #1343 #1523 #998 #1036 #1357 #1125 #973 #147 #1373 #1074 #1358) +#1969 := [th-lemma]: #1968 +#1971 := [unit-resolution #1969 #724 #1684 #746 #1440 #772 #798 #828]: #1970 +#1972 := [unit-resolution #1971 #1967 #1952 #1961 #1899 #1957]: #147 +#1973 := [unit-resolution #955 #1972]: #925 +#1974 := [unit-resolution #1236 #1973]: #919 +#1975 := (or #1161 #1151 #430) +#1976 := [th-lemma]: #1975 +#1977 := [unit-resolution #1976 #1899 #1897]: #1161 +#1978 := (or #476 #1036 #1112 #194 #1163 #1074) +#1979 := [unit-resolution #1611 #750 #772 #798 #824 #1404 #854]: #1978 +#1980 := [unit-resolution #1979 #1957 #1874 #1954 #1952 #1977]: #194 +#1981 := [unit-resolution #993 #1980]: #963 +#1982 := [unit-resolution #1248 #1981]: #957 +#1983 := [unit-resolution #1933 #1974 #1953 #1954 #1952 #1962 #1982]: #515 +#1984 := [unit-resolution #1238 #1973]: #921 +#1985 := [unit-resolution #1250 #1981]: #959 +#1849 := (or #923 #516 #1200 #961 #1036 #1163 #1074) +#1850 := [unit-resolution #1375 #698 #724 #772 #798 #854 #876]: #1849 +#1986 := [unit-resolution #1850 #1985 #1896 #1952 #1977 #1957 #1984]: #516 +#1987 := (or #509 #923 #1036 #1162 #1125) +#1988 := [unit-resolution #1720 #672 #698 #1684 #746 #1742 #772 #828 #850]: #1987 +#1989 := [unit-resolution #1988 #1984 #1961 #1967 #1957]: #509 +#1990 := [unit-resolution #646 #1989 #1986 #1983]: #631 +#1991 := (or #506 #884 #1112 #922 #1036 #1163 #1074) +#1992 := [unit-resolution #1603 #668 #694 #1404 #750 #772 #798 #824 #854]: #1991 +#1993 := [unit-resolution #1992 #1990 #1977 #1954 #1952 #1957 #1974]: #884 +#1994 := [unit-resolution #1224 #1993]: #886 +#1995 := [unit-resolution #917 #1994]: #101 +#1996 := [th-lemma #746 #1684 #1957 #1874 #854 #1899 #1870 #828 #1984 #1995 #698 #772 #1972]: false +#1997 := [lemma #1996]: #476 +#2014 := [unit-resolution #1221 #1997]: #1191 +#2015 := [unit-resolution #1320 #2014]: #1185 +#2034 := [th-lemma #876 #850 #1540 #2015 #802 #2033 #698 #772 #1828 #724 #1545 #1845 #1331]: false +#2036 := [lemma #2034]: #2035 +#2048 := [unit-resolution #2036 #1497 #2045 #1828 #1331]: #1150 +#2049 := [unit-resolution #1308 #2048 #2047]: false +#2051 := [lemma #2049]: #2050 +#2082 := [unit-resolution #2051 #1405 #1331]: #961 +#2083 := [unit-resolution #1250 #1923 #2082]: false +#2085 := [lemma #2083]: #2084 +#2089 := [unit-resolution #2085 #1331]: #288 +#2090 := [unit-resolution #1069 #2089]: #1039 +#2091 := [unit-resolution #1272 #2090]: #1033 +#2065 := [hypothesis]: #935 +#2066 := [unit-resolution #1244 #2065]: #936 +#2067 := [unit-resolution #956 #2066]: #147 +#2068 := [th-lemma #2065 #2033 #2067]: false +#2069 := [lemma #2068]: #933 +#2100 := (or #429 #516) +#2063 := (or #429 #1086 #516) +#2052 := [unit-resolution #1761 #1333]: #1109 +#2053 := [unit-resolution #1735 #2052]: #382 +#2054 := [hypothesis]: #1084 +#2055 := (or #1200 #516 #429) +#2056 := [unit-resolution #1383 #1864]: #2055 +#2057 := [unit-resolution #2056 #1333 #1331]: #1200 +#2060 := (or #1086 #383 #1113 #1188 #1162 #1198) +#2058 := (or #1086 #383 #1113 #1343 #1188 #1489 #1162 #1198 #1075) +#2059 := [th-lemma]: #2058 +#2061 := [unit-resolution #2059 #1447 #828 #850]: #2060 +#2062 := [unit-resolution #2061 #1631 #2057 #2015 #1870 #2054 #2053]: false +#2064 := [lemma #2062]: #2063 +#2086 := [unit-resolution #2064 #1333 #1331]: #1086 +#2087 := [unit-resolution #1290 #2086]: #1088 +#2088 := [unit-resolution #1108 #2087]: #335 +#2080 := (or #1109 #516) +#2070 := [unit-resolution #1308 #1872]: #1147 +#2020 := (or #194 #1150 #516 #1125 #1151 #1124) +#1762 := [hypothesis]: #1122 +#1775 := [hypothesis]: #1123 +#1803 := (or #194 #1151 #1150 #1125 #147 #1124) +#1764 := [unit-resolution #956 #1763]: #937 +#1765 := [unit-resolution #1244 #1764]: #933 +#1766 := (or #509 #885 #1522 #1364 #1365 #1489 #999 #1124 #1371 #1037 #1409 #935 #1150 #972 #1509 #1075 #1350) +#1767 := [th-lemma]: #1766 +#1768 := [unit-resolution #1767 #1620 #1765 #698 #672 #720 #1703 #750 #1459 #776 #1447 #802 #1762 #824 #1540 #850 #1742]: #509 +#1769 := (or #100 #1371 #935 #194 #147) +#1770 := [th-lemma]: #1769 +#1771 := [unit-resolution #1770 #1535 #1765 #698 #1763]: #100 +#1772 := [unit-resolution #917 #1771]: #887 +#1773 := [unit-resolution #1224 #1772]: #881 +#1776 := (or #335 #194 #1364 #1037 #1409 #999) +#1777 := [th-lemma]: #1776 +#1778 := [unit-resolution #1777 #1535 #750 #1459 #776 #1703]: #335 +#1779 := [unit-resolution #1107 #1778]: #1077 +#1780 := [unit-resolution #1284 #1779]: #1071 +#1241 := (or #936 #932) +#1242 := [def-axiom]: #1241 +#1781 := [unit-resolution #1242 #1764]: #932 +#1782 := (or #288 #1364 #999 #973 #147 #1373 #194) +#1783 := [th-lemma]: #1782 +#1784 := [unit-resolution #1783 #1535 #1440 #724 #1703 #750 #1763]: #288 +#1785 := [unit-resolution #1069 #1784]: #1039 +#1786 := [unit-resolution #1272 #1785]: #1033 +#1787 := (or #506 #884 #1528 #1523 #1343 #1337 #998 #1125 #1510 #1036 #1357 #934 #1151 #973 #1373 #1074 #1358) +#1788 := [th-lemma]: #1787 +#1789 := [unit-resolution #1788 #1786 #1781 #694 #1440 #724 #1684 #746 #668 #772 #1780 #798 #1775 #828 #1774 #854 #1773]: #506 +#1790 := (or #476 #1337 #1343 #1523 #1036 #1357 #998 #1125 #973 #147 #1373 #1074 #1358 #1151 #194) +#1791 := [th-lemma]: #1790 +#1792 := [unit-resolution #1791 #1535 #1440 #724 #1684 #746 #1786 #772 #1780 #798 #1775 #828 #1774 #854 #1763]: #476 +#1793 := [unit-resolution #1221 #1792]: #1191 +#1794 := [unit-resolution #1320 #1793]: #1185 +#1795 := (or #516 #1372 #1489 #1409 #1037 #1188 #1371 #935 #972 #1509 #1075 #1350 #1150) +#1796 := [th-lemma]: #1795 +#1797 := [unit-resolution #1796 #1620 #698 #720 #1459 #776 #1447 #802 #1540 #850 #1794 #876 #1765]: #516 +#1798 := [unit-resolution #1322 #1793]: #1187 +#1799 := (or #515 #1511 #1337 #1357 #1036 #1189 #1510 #934 #973 #1373 #1074 #1358 #1151) +#1800 := [th-lemma]: #1799 +#1801 := [unit-resolution #1800 #1786 #1440 #724 #694 #772 #1780 #798 #1774 #854 #1798 #880 #1781]: #515 +#1802 := [unit-resolution #646 #1801 #1797 #1789 #1768]: false +#1804 := [lemma #1802]: #1803 +#2011 := [unit-resolution #1804 #1535 #1540 #1775 #1774 #1762]: #147 +#2012 := [unit-resolution #955 #2011]: #925 +#2013 := [unit-resolution #1238 #2012]: #921 +#2016 := (or #516 #1188 #935 #972 #1150) +#2017 := [unit-resolution #1796 #698 #720 #1459 #776 #1447 #802 #850 #876]: #2016 +#2018 := [unit-resolution #2017 #1620 #2015 #1540 #1331]: #935 +#2019 := [th-lemma #2018 #2013 #2011]: false +#2021 := [lemma #2019]: #2020 +#2071 := [unit-resolution #2021 #2070 #1331 #1677 #1873 #1705]: #194 +#2072 := [unit-resolution #993 #2071]: #963 +#2073 := [unit-resolution #2010 #1675]: #288 +#2074 := [unit-resolution #1069 #2073]: #1039 +#2075 := [unit-resolution #1272 #2074]: #1033 +#2076 := (or #516 #1036 #1188 #935 #1150 #960 #1087) +#1823 := (or #516 #1372 #1489 #1357 #1036 #1188 #1371 #935 #1509 #1350 #1150 #960 #1523 #998 #1087) +#1824 := [th-lemma]: #1823 +#2077 := [unit-resolution #1824 #720 #1684 #746 #698 #772 #802 #850 #876]: #2076 +#2078 := [unit-resolution #2077 #2075 #2015 #2045 #2069 #1331 #2070]: #960 +#2079 := [unit-resolution #1248 #2078 #2072]: false +#2081 := [lemma #2079]: #2080 +#2092 := [unit-resolution #2081 #1331]: #1109 +#2093 := [unit-resolution #1735 #2092]: #382 +#2094 := [unit-resolution #1960 #2093 #1870]: #1123 +#2095 := (or #516 #923 #1074 #1036 #1162 #1125 #1188) +#2096 := [unit-resolution #1712 #1440 #724 #1684 #746 #698 #772 #798 #828 #850 #876]: #2095 +#2097 := [unit-resolution #2096 #1631 #2015 #2094 #1331 #2091 #2033]: #1074 +#2098 := [unit-resolution #1284 #2097]: #1076 +#2099 := [unit-resolution #1107 #2098 #2088]: false +#2101 := [lemma #2099]: #2100 +#2102 := [unit-resolution #2101 #1331]: #429 +#2103 := [unit-resolution #1183 #2102]: #1153 +#2104 := [unit-resolution #1308 #2103]: #1147 +#2105 := [unit-resolution #2077 #2104 #2015 #2045 #2069 #1331 #2091]: #960 +#2106 := [unit-resolution #1248 #2105]: #962 +#2107 := [unit-resolution #2017 #2104 #2015 #2069 #1331]: #972 +#2108 := [unit-resolution #1254 #2107]: #974 +#2109 := [unit-resolution #994 #2108]: #194 +#2110 := [unit-resolution #993 #2109 #2106]: false +#2111 := [lemma #2110]: #516 +#2127 := (or #1199 #1189 #477) +#2128 := [th-lemma]: #2127 +#2129 := [unit-resolution #2128 #1864 #1997]: #1199 +#2125 := (or #335 #288) +#1806 := [unit-resolution #1108 #1422]: #1089 +#1829 := [unit-resolution #1290 #1806]: #1084 +#2117 := (or #515 #1511 #1337 #1151 #1189 #1358 #922 #1510 #1409 #960 #1509 #1049 #1086) +#2118 := [th-lemma]: #2117 +#2119 := [unit-resolution #2118 #1829 #1924 #720 #1468 #776 #694 #798 #2116 #854 #1864 #880 #1928]: #515 +#2120 := (or #101 #922 #1510 #1409 #960 #1509 #1049 #335 #288) +#2121 := [th-lemma]: #2120 +#2122 := [unit-resolution #2121 #1422 #694 #1924 #720 #1405 #1468 #776 #1928]: #101 +#2123 := [unit-resolution #918 #2122]: #899 +#2124 := [unit-resolution #1916 #2123 #2119 #2111 #1497 #1925 #2033 #2116 #1928 #1924 #2115 #2112]: false +#2126 := [lemma #2124]: #2125 +#2130 := [unit-resolution #2126 #1405]: #335 +#2131 := [unit-resolution #1107 #2130]: #1077 +#2132 := [unit-resolution #1284 #2131]: #1071 +#2133 := [unit-resolution #1933 #2132 #2129 #2115 #1928 #2112 #1924]: #515 +#2134 := [unit-resolution #1916 #2133 #2111 #1497 #1925 #2033 #2116 #1928 #1924 #2115 #2112]: #898 +#2135 := [unit-resolution #918 #2134]: #100 +#2136 := [unit-resolution #917 #2135]: #887 +#2137 := [unit-resolution #1224 #2136]: #881 +#2138 := [unit-resolution #1943 #2137 #1497 #1870 #2116 #1928]: #506 +#2139 := [unit-resolution #646 #2138 #2111 #2133]: #632 +#2140 := [unit-resolution #1908 #2139 #1928 #694 #1924 #720 #1684 #746 #1468 #776 #2115 #824 #2112 #850 #672]: #897 +#2141 := [th-lemma #2140 #2135 #1742]: false +#2142 := [lemma #2141]: #288 +#2143 := [unit-resolution #1069 #2142]: #1039 +#2144 := [unit-resolution #1272 #2143]: #1033 +#2145 := [hypothesis]: #1150 +#2146 := [unit-resolution #1308 #2145]: #1152 +#2147 := [unit-resolution #1183 #2146]: #430 +#2148 := [unit-resolution #1184 #2147]: #1165 +#2149 := [unit-resolution #1314 #2148]: #1160 +#2150 := [unit-resolution #1761 #2147]: #1109 +#2151 := [unit-resolution #1735 #2150]: #382 +#2152 := [unit-resolution #1960 #2151 #1870]: #1123 +#2153 := [unit-resolution #1988 #2152 #2149 #2033 #2144]: #509 +#2154 := (or #1149 #1147) +#2155 := [th-lemma]: #2154 +#2156 := [unit-resolution #2155 #2145]: #1149 +#2157 := [unit-resolution #1894 #2147]: #1200 +#2158 := [unit-resolution #2061 #2149 #2015 #1870 #2157 #2151]: #1086 +#2159 := [unit-resolution #1290 #2158]: #1088 +#2160 := [unit-resolution #1108 #2159]: #335 +#2161 := [unit-resolution #1107 #2160]: #1077 +#2162 := [unit-resolution #1284 #2161]: #1071 +#2163 := [unit-resolution #1971 #2162 #2149 #2152 #2156 #2144]: #147 +#2164 := [unit-resolution #955 #2163]: #925 +#2165 := [unit-resolution #1236 #2164]: #919 +#2166 := [unit-resolution #1316 #2148]: #1161 +#2167 := (or #100 #923 #1371 #1357 #1523 #998 #1036 #383 #429 #1343 #1113 #973 #1373 #1074 #1358) +#2168 := [th-lemma]: #2167 +#2169 := [unit-resolution #2168 #2162 #698 #1440 #724 #1684 #746 #2144 #772 #2033 #798 #2151 #1870 #828 #2147]: #100 +#2170 := [unit-resolution #917 #2169]: #887 +#2171 := [unit-resolution #1224 #2170]: #881 +#2172 := [unit-resolution #1992 #2171 #2166 #2150 #2162 #2144 #2165]: #506 +#2173 := (or #195 #1357 #1523 #998 #1036 #383 #429 #1343 #1113) +#2174 := [th-lemma]: #2173 +#2175 := [unit-resolution #2174 #2151 #746 #2144 #772 #1684 #1870 #828 #2147]: #195 +#2176 := [unit-resolution #994 #2175]: #975 +#2177 := [unit-resolution #1254 #2176]: #970 +#2178 := (or #515 #922 #1074 #1036 #972 #1163 #1112) +#2179 := [unit-resolution #1622 #694 #720 #1404 #750 #772 #1864 #798 #824 #854 #880]: #2178 +#2180 := [unit-resolution #2179 #2177 #2150 #2162 #2166 #2144 #2165]: #515 +#2181 := [unit-resolution #646 #2180 #2172 #2111 #2153]: false +#2182 := [lemma #2181]: #1147 +#1805 := [unit-resolution #1302 #1729]: #1122 +#2231 := (or #194 #382) +#2183 := (or #1150 #429 #1163) +#2184 := [th-lemma]: #2183 +#2185 := [unit-resolution #2184 #1333 #2182]: #1163 +#2186 := [unit-resolution #1316 #2185 #1334]: false +#2187 := [lemma #2186]: #429 +#2196 := [unit-resolution #1183 #2187]: #1153 +#2197 := [unit-resolution #1310 #2196]: #1149 +#1817 := [unit-resolution #1304 #1729]: #1123 +#2217 := [unit-resolution #1804 #1535 #2182 #1817 #2197 #1805]: #147 +#2218 := [unit-resolution #955 #2217]: #925 +#2219 := [unit-resolution #1236 #2218]: #919 +#2210 := [unit-resolution #1976 #2197 #2187]: #1161 +#2220 := (or #509 #1124 #935 #1150 #972) +#2221 := [unit-resolution #1767 #698 #672 #720 #1703 #750 #1459 #776 #1447 #802 #824 #850 #1742]: #2220 +#2222 := [unit-resolution #2221 #1620 #2069 #1805 #2182]: #509 +#2223 := (or #515 #922 #1163 #972 #1124) +#2224 := [unit-resolution #1707 #720 #1703 #750 #1459 #776 #1447 #802 #694 #824 #854 #880 #1864]: #2223 +#2225 := [unit-resolution #2224 #2219 #1805 #1620 #2210]: #515 +#2226 := [unit-resolution #646 #2225 #2111 #2222]: #631 +#2211 := (or #506 #884 #922 #1163 #1124) +#2212 := [unit-resolution #1724 #668 #694 #1703 #750 #1459 #776 #824 #854]: #2211 +#2227 := [unit-resolution #2212 #2226 #1805 #2210 #2219]: #884 +#2228 := [unit-resolution #1224 #2227]: #886 +#2229 := [unit-resolution #917 #2228]: #101 +#2230 := [th-lemma #1620 #720 #1459 #776 #1447 #802 #2033 #2229 #698 #1428 #2217]: false +#2232 := [lemma #2230]: #2231 +#2242 := [unit-resolution #2232 #1428]: #194 +#2243 := [unit-resolution #993 #2242]: #963 +#2244 := [unit-resolution #1248 #2243]: #957 +#2193 := (or #509 #1124 #1036 #935 #1150 #960 #1087) +#1814 := (or #509 #885 #1522 #1523 #1365 #1489 #998 #1124 #1371 #1036 #1357 #935 #1150 #1509 #1350 #960 #1087) +#1815 := [th-lemma]: #1814 +#2194 := [unit-resolution #1815 #698 #720 #1684 #746 #672 #772 #802 #824 #850 #1742]: #2193 +#2245 := [unit-resolution #2194 #2244 #2069 #2144 #2045 #1805 #2182]: #509 +#2205 := (or #100 #935 #1036 #382 #960 #1087) +#1834 := (or #100 #1371 #935 #1523 #1036 #1357 #998 #1509 #382 #1350 #960 #1087) +#1835 := [th-lemma]: #1834 +#2206 := [unit-resolution #1835 #698 #720 #1684 #746 #772 #802]: #2205 +#2246 := [unit-resolution #2206 #2244 #2045 #2069 #2144 #1428]: #100 +#2247 := [unit-resolution #917 #2246]: #887 +#2248 := [unit-resolution #1224 #2247]: #881 +#2215 := (or #335 #382) +#2188 := (or #335 #194) +#2189 := [unit-resolution #1777 #750 #1459 #776 #1703]: #2188 +#2190 := [unit-resolution #2189 #1422]: #194 +#2191 := [unit-resolution #993 #2190]: #963 +#2192 := [unit-resolution #1248 #2191]: #957 +#2195 := [unit-resolution #2194 #2192 #2069 #2144 #2045 #1805 #2182]: #509 +#2198 := [unit-resolution #1250 #2191]: #959 +#1840 := (or #335 #934 #1151 #961 #935 #960 #1150 #382) +#1807 := [unit-resolution #1292 #1806]: #1085 +#1808 := [hypothesis]: #933 +#1809 := (or #288 #382 #1350 #335 #1087) +#1810 := [th-lemma]: #1809 +#1811 := [unit-resolution #1810 #1422 #1807 #802 #1428]: #288 +#1812 := [unit-resolution #1069 #1811]: #1039 +#1813 := [unit-resolution #1272 #1812]: #1033 +#1816 := [unit-resolution #1815 #1813 #1808 #698 #1536 #720 #1684 #746 #672 #772 #1807 #802 #1805 #824 #1540 #850 #1742]: #509 +#1818 := (or #476 #1337 #1343 #1125 #1151 #335 #382) +#1819 := [th-lemma]: #1818 +#1820 := [unit-resolution #1819 #1422 #1817 #828 #1774 #854 #1428]: #476 +#1821 := [unit-resolution #1221 #1820]: #1191 +#1822 := [unit-resolution #1320 #1821]: #1185 +#1825 := [unit-resolution #1824 #1813 #1536 #720 #1684 #746 #698 #772 #1807 #802 #1540 #850 #1822 #876 #1808]: #516 +#1826 := [hypothesis]: #932 +#1827 := [unit-resolution #1322 #1821]: #1187 +#1830 := (or #515 #1511 #1337 #1409 #1037 #1189 #1510 #934 #1373 #1358 #1151 #961 #1364 #999 #1086) +#1831 := [th-lemma]: #1830 +#1832 := [unit-resolution #1831 #1829 #1828 #724 #1703 #750 #1459 #776 #694 #798 #1774 #854 #1827 #880 #1826]: #515 +#1833 := [unit-resolution #646 #1832 #1825 #1816]: #631 +#1836 := [unit-resolution #1835 #1813 #698 #1536 #720 #1684 #746 #1808 #772 #1807 #802 #1428]: #100 +#1837 := [unit-resolution #917 #1836]: #887 +#1838 := [unit-resolution #1224 #1837]: #881 +#1839 := [th-lemma #1838 #668 #750 #828 #854 #1703 #1817 #694 #1459 #776 #1826 #1774 #724 #798 #1828 #1829 #1833]: false +#1841 := [lemma #1839]: #1840 +#2199 := [unit-resolution #1841 #2198 #2069 #1422 #2197 #2192 #2182 #1428]: #934 +#2200 := [unit-resolution #1242 #2199]: #936 +#2201 := [unit-resolution #956 #2200]: #147 +#2202 := [unit-resolution #955 #2201]: #925 +#2203 := [unit-resolution #1236 #2202]: #919 +#2204 := [unit-resolution #2118 #2203 #1829 #720 #1468 #776 #694 #798 #2197 #854 #1864 #880 #2192]: #515 +#2207 := [unit-resolution #2206 #2192 #2045 #2069 #2144 #1428]: #100 +#2208 := [unit-resolution #917 #2207]: #887 +#2209 := [unit-resolution #1224 #2208]: #881 +#2213 := [unit-resolution #2212 #2203 #1805 #2210 #2209]: #506 +#2214 := [unit-resolution #646 #2213 #2204 #2111 #2195]: false +#2216 := [lemma #2214]: #2215 +#2249 := [unit-resolution #2216 #1428]: #335 +#2250 := [unit-resolution #1107 #2249]: #1077 +#2251 := [unit-resolution #1284 #2250]: #1071 +#2252 := (or #1084 #1074 #1357 #1523 #998 #1036 #195) +#2253 := [th-lemma]: #2252 +#2254 := [unit-resolution #2253 #2251 #746 #2144 #772 #1684 #2242]: #1084 +#2255 := [unit-resolution #1250 #2243]: #959 +#2240 := (or #934 #632 #884 #1074 #1125 #961 #1086) +#2233 := (or #515 #934 #1151 #961 #1086) +#2234 := [unit-resolution #1831 #1864 #724 #1703 #750 #1459 #776 #694 #798 #854 #880]: #2233 +#2235 := [unit-resolution #2234 #1826 #2197 #1828 #2054]: #515 +#2236 := (or #506 #884 #1125 #1036 #934 #1151 #1074) +#2237 := [unit-resolution #1788 #694 #1440 #724 #1684 #746 #668 #772 #798 #828 #854]: #2236 +#2238 := [unit-resolution #2237 #1826 #1636 #1638 #1775 #2197 #2144]: #506 +#2239 := [unit-resolution #646 #2238 #2235 #2111 #1628]: false +#2241 := [lemma #2239]: #2240 +#2256 := [unit-resolution #2241 #2245 #2248 #2251 #1817 #2255 #2254]: #934 +#2257 := [unit-resolution #1242 #2256]: #936 +#2258 := [unit-resolution #956 #2257]: #147 +#2259 := [unit-resolution #955 #2258]: #925 +#2260 := [unit-resolution #1236 #2259]: #919 +#2261 := [unit-resolution #2212 #2260 #1805 #2210 #2248]: #506 +#2262 := [unit-resolution #2118 #2260 #2254 #720 #1468 #776 #694 #798 #2197 #854 #1864 #880 #2244]: #515 +#2263 := [unit-resolution #646 #2262 #2261 #2111 #2245]: false +#2264 := [lemma #2263]: #382 +#2265 := [unit-resolution #1145 #2264]: #1115 +#2266 := [unit-resolution #1296 #2265]: #1109 +#2267 := [unit-resolution #2189 #1535]: #335 +#2268 := [unit-resolution #1107 #2267]: #1077 +#2269 := [unit-resolution #1284 #2268]: #1071 +#2270 := [unit-resolution #1966 #2269 #2142 #2266 #2182]: #1160 +#2271 := (or #1008 #998 #1036 #1357 #1074 #1358 #383) +#2272 := [th-lemma]: #2271 +#2273 := [unit-resolution #2272 #2269 #2144 #772 #1684 #798 #2264]: #1008 +#2274 := (or #509 #1010 #1113 #923 #1162) +#2275 := [unit-resolution #1608 #672 #698 #1742 #746 #1459 #776 #1447 #802 #828 #850]: #2274 +#2276 := [unit-resolution #2275 #2273 #1870 #2270 #2033]: #509 +#2277 := [unit-resolution #1960 #2264 #1870]: #1123 +#2278 := [unit-resolution #1971 #2270 #2269 #2277 #2197 #2144]: #147 +#2279 := [unit-resolution #955 #2278]: #925 +#2280 := [unit-resolution #1236 #2279]: #919 +#2281 := (or #1010 #999 #923 #100 #1371 #961 #1373) +#2282 := [th-lemma]: #2281 +#2283 := [unit-resolution #2282 #2273 #698 #1584 #724 #1703 #2033]: #100 +#2284 := [unit-resolution #917 #2283]: #887 +#2285 := [unit-resolution #1224 #2284]: #881 +#2286 := [unit-resolution #1992 #2285 #2210 #2266 #2269 #2144 #2280]: #506 +#2287 := [unit-resolution #2179 #2280 #2266 #1620 #2210 #2144 #2269]: #515 +#2288 := [unit-resolution #646 #2287 #2286 #2111 #2276]: false +#2289 := [lemma #2288]: #194 +#2305 := [unit-resolution #2253 #2302 #746 #2144 #772 #1684 #2289]: #1074 +#2306 := [unit-resolution #1284 #2305]: #1076 +#2307 := [unit-resolution #1107 #2306 #2304]: false +#2308 := [lemma #2307]: #1084 +#2300 := (or #1086 #515) +#2290 := [hypothesis]: #633 +#2291 := [unit-resolution #993 #2289]: #963 +#2292 := [unit-resolution #1250 #2291]: #959 +#2293 := [unit-resolution #2234 #2054 #2197 #2292 #2290]: #934 +#2294 := [unit-resolution #1242 #2293]: #936 +#2295 := [unit-resolution #1248 #2291]: #957 +#2296 := [unit-resolution #2118 #2054 #2290 #720 #1468 #776 #694 #798 #2197 #854 #1864 #880 #2295]: #922 +#2297 := [unit-resolution #1236 #2296]: #924 +#2298 := [unit-resolution #955 #2297]: #148 +#2299 := [unit-resolution #956 #2298 #2294]: false +#2301 := [lemma #2299]: #2300 +#1848 := [unit-resolution #2301 #2308]: #515 +#1851 := [hypothesis]: #632 +#1852 := (or #897 #1522 #509 #1523 #998 #1365 #1489 #1150 #1509 #1350 #633 #1372 #1188 #960 #1087 #1112) +#1853 := [th-lemma]: #1852 +#1846 := [unit-resolution #1853 #1851 #2295 #720 #1684 #746 #2045 #802 #2266 #824 #2182 #850 #2015 #876 #672 #1848]: #897 +#1847 := [unit-resolution #1232 #1846]: #898 +#1854 := [unit-resolution #918 #1847]: #100 +#1855 := (or #509 #1124) +#1856 := [unit-resolution #2194 #2069 #2144 #2045 #2295 #2182]: #1855 +#2309 := [unit-resolution #1856 #1851]: #1124 +#2310 := [th-lemma #1848 #876 #850 #2182 #2015 #2309 #2266 #1854]: false +#2311 := [lemma #2310]: #509 +#2312 := (or #631 #632) +#2313 := [unit-resolution #646 #2111 #1848]: #2312 +#2314 := [unit-resolution #2313 #2311]: #631 +#2315 := (or #884 #633 #1372 #1188 #1125 #1528 #506 #1364 #999 #1343 #1373 #1358 #961 #1086) +#2316 := [th-lemma]: #2315 +#2317 := [unit-resolution #2316 #668 #2292 #724 #1703 #750 #2308 #798 #2277 #828 #2015 #876 #2314 #1848]: #884 +#2318 := [unit-resolution #1224 #2317]: #886 +#2319 := (or #896 #1528 #506 #1364 #999 #1343 #1337 #1151 #1373 #1358 #634 #1511 #1189 #961 #1086 #1113) +#2320 := [th-lemma]: #2319 +#2321 := [unit-resolution #2320 #668 #2292 #724 #1703 #750 #2308 #798 #1870 #828 #2197 #854 #1864 #880 #2314 #2111]: #896 +#2322 := [unit-resolution #1230 #2321]: #898 +#2323 := [unit-resolution #918 #2322]: #100 +[unit-resolution #917 #2323 #2318]: false +unsat diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_mono_01 --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_mono_01 Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,58 @@ +(benchmark Isabelle +:extrasorts ( T2 T13 T12 T11 T10 T9 T8 T7 T6 T5 T4 T1 T3) +:extrafuns ( + (uf_37 T13) + (uf_34 T12) + (uf_31 T11) + (uf_28 T10) + (uf_25 T9) + (uf_22 T8) + (uf_19 T7) + (uf_16 T6) + (uf_13 T5) + (uf_10 T4) + (uf_7 T1) + (uf_4 T3) + (uf_36 Int T13 T13) + (uf_33 T13 T12 T12) + (uf_30 T12 T11 T11) + (uf_27 T11 T10 T10) + (uf_24 T10 T9 T9) + (uf_21 T9 T8 T8) + (uf_18 T8 T7 T7) + (uf_15 T7 T6 T6) + (uf_12 T6 T5 T5) + (uf_9 T5 T4 T4) + (uf_6 T4 T1 T1) + (uf_3 T1 T3 T3) + ) +:extrapreds ( + (up_35 Int) + (up_32 T13) + (up_29 T12) + (up_26 T11) + (up_23 T10) + (up_20 T9) + (up_17 T8) + (up_14 T7) + (up_11 T6) + (up_8 T5) + (up_5 T4) + (up_1 T1) + (up_2 T3) + ) +:assumption (forall (?x1 T1) (and (up_1 ?x1) (or (up_2 (uf_3 ?x1 uf_4)) (not (up_2 (uf_3 ?x1 uf_4)))))) +:assumption (forall (?x2 T4) (and (up_5 ?x2) (or (up_1 (uf_6 ?x2 uf_7)) (not (up_1 (uf_6 ?x2 uf_7)))))) +:assumption (forall (?x3 T5) (and (up_8 ?x3) (or (up_5 (uf_9 ?x3 uf_10)) (not (up_5 (uf_9 ?x3 uf_10)))))) +:assumption (forall (?x4 T6) (and (up_11 ?x4) (or (up_8 (uf_12 ?x4 uf_13)) (not (up_8 (uf_12 ?x4 uf_13)))))) +:assumption (forall (?x5 T7) (and (up_14 ?x5) (or (up_11 (uf_15 ?x5 uf_16)) (not (up_11 (uf_15 ?x5 uf_16)))))) +:assumption (forall (?x6 T8) (and (up_17 ?x6) (or (up_14 (uf_18 ?x6 uf_19)) (not (up_14 (uf_18 ?x6 uf_19)))))) +:assumption (forall (?x7 T9) (and (up_20 ?x7) (or (up_17 (uf_21 ?x7 uf_22)) (not (up_17 (uf_21 ?x7 uf_22)))))) +:assumption (forall (?x8 T10) (and (up_23 ?x8) (or (up_20 (uf_24 ?x8 uf_25)) (not (up_20 (uf_24 ?x8 uf_25)))))) +:assumption (forall (?x9 T11) (and (up_26 ?x9) (or (up_23 (uf_27 ?x9 uf_28)) (not (up_23 (uf_27 ?x9 uf_28)))))) +:assumption (forall (?x10 T12) (and (up_29 ?x10) (or (up_26 (uf_30 ?x10 uf_31)) (not (up_26 (uf_30 ?x10 uf_31)))))) +:assumption (forall (?x11 T13) (and (up_32 ?x11) (or (up_29 (uf_33 ?x11 uf_34)) (not (up_29 (uf_33 ?x11 uf_34)))))) +:assumption (forall (?x12 Int) (and (up_35 ?x12) (or (up_32 (uf_36 ?x12 uf_37)) (not (up_32 (uf_36 ?x12 uf_37)))))) +:assumption (not (up_35 1)) +:formula true +) diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_mono_01.proof --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_mono_01.proof Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,50 @@ +#2 := false +decl up_35 :: (-> int bool) +#112 := 1::int +#113 := (up_35 1::int) +#114 := (not #113) +#297 := [asserted]: #114 +#103 := (:var 0 int) +#104 := (up_35 #103) +#911 := (pattern #104) +#912 := (forall (vars (?x12 int)) (:pat #911) #104) +#294 := (forall (vars (?x12 int)) #104) +#915 := (iff #294 #912) +#913 := (iff #104 #104) +#914 := [refl]: #913 +#916 := [quant-intro #914]: #915 +#320 := (~ #294 #294) +#361 := (~ #104 #104) +#362 := [refl]: #361 +#321 := [nnf-pos #362]: #320 +decl up_32 :: (-> T13 bool) +decl uf_36 :: (-> int T13 T13) +decl uf_37 :: T13 +#105 := uf_37 +#106 := (uf_36 #103 uf_37) +#107 := (up_32 #106) +#108 := (not #107) +#109 := (or #107 #108) +#110 := (and #104 #109) +#111 := (forall (vars (?x12 int)) #110) +#295 := (iff #111 #294) +#292 := (iff #110 #104) +#1 := true +#287 := (and #104 true) +#290 := (iff #287 #104) +#291 := [rewrite]: #290 +#288 := (iff #110 #287) +#284 := (iff #109 true) +#286 := [rewrite]: #284 +#289 := [monotonicity #286]: #288 +#293 := [trans #289 #291]: #292 +#296 := [quant-intro #293]: #295 +#283 := [asserted]: #111 +#299 := [mp #283 #296]: #294 +#363 := [mp~ #299 #321]: #294 +#917 := [mp #363 #916]: #912 +#418 := (not #912) +#504 := (or #418 #113) +#419 := [quant-inst]: #504 +[unit-resolution #419 #917 #297]: false +unsat diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_mono_02 --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_mono_02 Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,41 @@ +(benchmark Isabelle +:extrasorts ( T5 T6 T3 T1 T2 T4 T8 T7) +:extrafuns ( + (uf_19 T1) + (uf_3 Int T3) + (uf_7 T2) + (uf_8 T4) + (uf_2 T1 T2 T2) + (uf_6 Int T4 T4) + (uf_10 T5 T1 T3) + (uf_12 T6 Int T3) + (uf_13 T2 T3) + (uf_14 T4 T3) + (uf_17 T8 T3) + (uf_15 T7 T3) + (uf_18 T1 T8) + (uf_16 Int T7) + (uf_9 T5 T2 T3) + (uf_11 T6 T4 T3) + (uf_1 T2 T3) + (uf_5 T4 T3) + (uf_4 T3 Int) + ) +:assumption (forall (?x1 T1) (?x2 T2) (= (uf_1 (uf_2 ?x1 ?x2)) (uf_3 (+ (uf_4 (uf_1 ?x2)) (uf_4 (uf_3 (+ 0 1))))))) +:assumption (forall (?x3 Int) (?x4 T4) (= (uf_5 (uf_6 ?x3 ?x4)) (uf_3 (+ (uf_4 (uf_5 ?x4)) (uf_4 (uf_3 (+ 0 1))))))) +:assumption (= (uf_1 uf_7) (uf_3 0)) +:assumption (= (uf_5 uf_8) (uf_3 0)) +:assumption (forall (?x5 T5) (?x6 T1) (?x7 T2) (= (uf_9 ?x5 (uf_2 ?x6 ?x7)) (uf_3 (+ (+ (uf_4 (uf_10 ?x5 ?x6)) (uf_4 (uf_9 ?x5 ?x7))) (uf_4 (uf_3 (+ 0 1))))))) +:assumption (forall (?x8 T6) (?x9 Int) (?x10 T4) (= (uf_11 ?x8 (uf_6 ?x9 ?x10)) (uf_3 (+ (+ (uf_4 (uf_12 ?x8 ?x9)) (uf_4 (uf_11 ?x8 ?x10))) (uf_4 (uf_3 (+ 0 1))))))) +:assumption (forall (?x11 T5) (= (uf_9 ?x11 uf_7) (uf_3 0))) +:assumption (forall (?x12 T6) (= (uf_11 ?x12 uf_8) (uf_3 0))) +:assumption (forall (?x13 T2) (= (uf_13 ?x13) (uf_1 ?x13))) +:assumption (forall (?x14 T4) (= (uf_14 ?x14) (uf_5 ?x14))) +:assumption (forall (?x15 Int) (= (uf_15 (uf_16 ?x15)) (uf_14 (uf_6 ?x15 uf_8)))) +:assumption (forall (?x16 T1) (= (uf_17 (uf_18 ?x16)) (uf_13 (uf_2 ?x16 uf_7)))) +:assumption (forall (?x17 T3) (= (uf_3 (uf_4 ?x17)) ?x17)) +:assumption (forall (?x18 Int) (implies (<= 0 ?x18) (= (uf_4 (uf_3 ?x18)) ?x18))) +:assumption (forall (?x19 Int) (implies (< ?x19 0) (= (uf_4 (uf_3 ?x19)) 0))) +:assumption (not (= (uf_15 (uf_16 3)) (uf_17 (uf_18 uf_19)))) +:formula true +) diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_mono_02.proof --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_mono_02.proof Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,492 @@ +#2 := false +decl uf_17 :: (-> T8 T3) +decl uf_18 :: (-> T1 T8) +decl uf_19 :: T1 +#104 := uf_19 +#105 := (uf_18 uf_19) +#106 := (uf_17 #105) +decl uf_15 :: (-> T7 T3) +decl uf_16 :: (-> int T7) +#101 := 3::int +#102 := (uf_16 3::int) +#103 := (uf_15 #102) +#107 := (= #103 #106) +decl uf_13 :: (-> T2 T3) +decl uf_2 :: (-> T1 T2 T2) +decl uf_7 :: T2 +#29 := uf_7 +#857 := (uf_2 uf_19 uf_7) +#859 := (uf_13 #857) +#599 := (= #859 #106) +#526 := (= #106 #859) +#79 := (:var 0 T1) +#82 := (uf_2 #79 uf_7) +#932 := (pattern #82) +#80 := (uf_18 #79) +#931 := (pattern #80) +#83 := (uf_13 #82) +#81 := (uf_17 #80) +#84 := (= #81 #83) +#933 := (forall (vars (?x16 T1)) (:pat #931 #932) #84) +#85 := (forall (vars (?x16 T1)) #84) +#936 := (iff #85 #933) +#934 := (iff #84 #84) +#935 := [refl]: #934 +#937 := [quant-intro #935]: #936 +#347 := (~ #85 #85) +#384 := (~ #84 #84) +#385 := [refl]: #384 +#348 := [nnf-pos #385]: #347 +#238 := [asserted]: #85 +#386 := [mp~ #238 #348]: #85 +#938 := [mp #386 #937]: #933 +#861 := (not #933) +#862 := (or #861 #526) +#863 := [quant-inst]: #862 +#601 := [unit-resolution #863 #938]: #526 +#588 := [symm #601]: #599 +#586 := (= #103 #859) +decl uf_1 :: (-> T2 T3) +#558 := (uf_1 #857) +#832 := (= #558 #859) +#5 := (:var 0 T2) +#66 := (uf_13 #5) +#908 := (pattern #66) +#8 := (uf_1 #5) +#907 := (pattern #8) +#222 := (= #8 #66) +#909 := (forall (vars (?x13 T2)) (:pat #907 #908) #222) +#226 := (forall (vars (?x13 T2)) #222) +#912 := (iff #226 #909) +#910 := (iff #222 #222) +#911 := [refl]: #910 +#913 := [quant-intro #911]: #912 +#341 := (~ #226 #226) +#375 := (~ #222 #222) +#376 := [refl]: #375 +#342 := [nnf-pos #376]: #341 +#67 := (= #66 #8) +#68 := (forall (vars (?x13 T2)) #67) +#227 := (iff #68 #226) +#224 := (iff #67 #222) +#225 := [rewrite]: #224 +#228 := [quant-intro #225]: #227 +#221 := [asserted]: #68 +#231 := [mp #221 #228]: #226 +#377 := [mp~ #231 #342]: #226 +#914 := [mp #377 #913]: #909 +#451 := (not #909) +#837 := (or #451 #832) +#547 := [quant-inst]: #837 +#615 := [unit-resolution #547 #914]: #832 +#585 := (= #103 #558) +decl uf_3 :: (-> int T3) +decl uf_4 :: (-> T3 int) +#30 := (uf_1 uf_7) +#806 := (uf_4 #30) +#11 := 1::int +#127 := (uf_3 1::int) +#130 := (uf_4 #127) +#649 := (+ #130 #806) +#794 := (uf_3 #649) +#597 := (= #794 #558) +#683 := (= #558 #794) +#4 := (:var 1 T1) +#6 := (uf_2 #4 #5) +#865 := (pattern #6) +#9 := (uf_4 #8) +#133 := (+ #9 #130) +#136 := (uf_3 #133) +#7 := (uf_1 #6) +#139 := (= #7 #136) +#866 := (forall (vars (?x1 T1) (?x2 T2)) (:pat #865) #139) +#142 := (forall (vars (?x1 T1) (?x2 T2)) #139) +#869 := (iff #142 #866) +#867 := (iff #139 #139) +#868 := [refl]: #867 +#870 := [quant-intro #868]: #869 +#361 := (~ #142 #142) +#359 := (~ #139 #139) +#360 := [refl]: #359 +#362 := [nnf-pos #360]: #361 +#10 := 0::int +#12 := (+ 0::int 1::int) +#13 := (uf_3 #12) +#14 := (uf_4 #13) +#15 := (+ #9 #14) +#16 := (uf_3 #15) +#17 := (= #7 #16) +#18 := (forall (vars (?x1 T1) (?x2 T2)) #17) +#143 := (iff #18 #142) +#140 := (iff #17 #139) +#137 := (= #16 #136) +#134 := (= #15 #133) +#131 := (= #14 #130) +#128 := (= #13 #127) +#125 := (= #12 1::int) +#126 := [rewrite]: #125 +#129 := [monotonicity #126]: #128 +#132 := [monotonicity #129]: #131 +#135 := [monotonicity #132]: #134 +#138 := [monotonicity #135]: #137 +#141 := [monotonicity #138]: #140 +#144 := [quant-intro #141]: #143 +#124 := [asserted]: #18 +#147 := [mp #124 #144]: #142 +#363 := [mp~ #147 #362]: #142 +#871 := [mp #363 #870]: #866 +#701 := (not #866) +#694 := (or #701 #683) +#688 := (+ #806 #130) +#689 := (uf_3 #688) +#690 := (= #558 #689) +#702 := (or #701 #690) +#704 := (iff #702 #694) +#706 := (iff #694 #694) +#799 := [rewrite]: #706 +#698 := (iff #690 #683) +#795 := (= #689 #794) +#797 := (= #688 #649) +#699 := [rewrite]: #797 +#798 := [monotonicity #699]: #795 +#700 := [monotonicity #798]: #698 +#705 := [monotonicity #700]: #704 +#796 := [trans #705 #799]: #704 +#703 := [quant-inst]: #702 +#800 := [mp #703 #796]: #694 +#614 := [unit-resolution #800 #871]: #683 +#598 := [symm #614]: #597 +#583 := (= #103 #794) +#595 := (= #127 #794) +#605 := (= #794 #127) +#618 := (= #649 1::int) +#780 := (<= #806 0::int) +#778 := (= #806 0::int) +#31 := (uf_3 0::int) +#858 := (uf_4 #31) +#855 := (= #858 0::int) +#72 := (:var 0 int) +#92 := (uf_3 #72) +#947 := (pattern #92) +#266 := (>= #72 0::int) +#267 := (not #266) +#93 := (uf_4 #92) +#248 := (= #72 #93) +#273 := (or #248 #267) +#948 := (forall (vars (?x18 int)) (:pat #947) #273) +#278 := (forall (vars (?x18 int)) #273) +#951 := (iff #278 #948) +#949 := (iff #273 #273) +#950 := [refl]: #949 +#952 := [quant-intro #950]: #951 +#351 := (~ #278 #278) +#390 := (~ #273 #273) +#391 := [refl]: #390 +#352 := [nnf-pos #391]: #351 +#94 := (= #93 #72) +#91 := (<= 0::int #72) +#95 := (implies #91 #94) +#96 := (forall (vars (?x18 int)) #95) +#281 := (iff #96 #278) +#255 := (not #91) +#256 := (or #255 #248) +#261 := (forall (vars (?x18 int)) #256) +#279 := (iff #261 #278) +#276 := (iff #256 #273) +#270 := (or #267 #248) +#274 := (iff #270 #273) +#275 := [rewrite]: #274 +#271 := (iff #256 #270) +#268 := (iff #255 #267) +#264 := (iff #91 #266) +#265 := [rewrite]: #264 +#269 := [monotonicity #265]: #268 +#272 := [monotonicity #269]: #271 +#277 := [trans #272 #275]: #276 +#280 := [quant-intro #277]: #279 +#262 := (iff #96 #261) +#259 := (iff #95 #256) +#252 := (implies #91 #248) +#257 := (iff #252 #256) +#258 := [rewrite]: #257 +#253 := (iff #95 #252) +#250 := (iff #94 #248) +#251 := [rewrite]: #250 +#254 := [monotonicity #251]: #253 +#260 := [trans #254 #258]: #259 +#263 := [quant-intro #260]: #262 +#282 := [trans #263 #280]: #281 +#247 := [asserted]: #96 +#283 := [mp #247 #282]: #278 +#392 := [mp~ #283 #352]: #278 +#953 := [mp #392 #952]: #948 +#848 := (not #948) +#850 := (or #848 #855) +#527 := (>= 0::int 0::int) +#860 := (not #527) +#864 := (= 0::int #858) +#854 := (or #864 #860) +#489 := (or #848 #854) +#851 := (iff #489 #850) +#852 := (iff #850 #850) +#838 := [rewrite]: #852 +#847 := (iff #854 #855) +#843 := (or #855 false) +#846 := (iff #843 #855) +#841 := [rewrite]: #846 +#844 := (iff #854 #843) +#505 := (iff #860 false) +#1 := true +#498 := (not true) +#503 := (iff #498 false) +#504 := [rewrite]: #503 +#840 := (iff #860 #498) +#514 := (iff #527 true) +#856 := [rewrite]: #514 +#502 := [monotonicity #856]: #840 +#842 := [trans #502 #504]: #505 +#513 := (iff #864 #855) +#518 := [rewrite]: #513 +#845 := [monotonicity #518 #842]: #844 +#484 := [trans #845 #841]: #847 +#849 := [monotonicity #484]: #851 +#839 := [trans #849 #838]: #851 +#490 := [quant-inst]: #489 +#546 := [mp #490 #839]: #850 +#644 := [unit-resolution #546 #953]: #855 +#621 := (= #806 #858) +#32 := (= #30 #31) +#159 := [asserted]: #32 +#626 := [monotonicity #159]: #621 +#616 := [trans #626 #644]: #778 +#606 := (not #778) +#608 := (or #606 #780) +#609 := [th-lemma]: #608 +#612 := [unit-resolution #609 #616]: #780 +#790 := (>= #806 0::int) +#613 := (or #606 #790) +#617 := [th-lemma]: #613 +#610 := [unit-resolution #617 #616]: #790 +#723 := (<= #130 1::int) +#746 := (= #130 1::int) +#713 := (or #848 #746) +#755 := (>= 1::int 0::int) +#756 := (not #755) +#743 := (= 1::int #130) +#744 := (or #743 #756) +#714 := (or #848 #744) +#718 := (iff #714 #713) +#720 := (iff #713 #713) +#725 := [rewrite]: #720 +#739 := (iff #744 #746) +#735 := (or #746 false) +#738 := (iff #735 #746) +#733 := [rewrite]: #738 +#736 := (iff #744 #735) +#731 := (iff #756 false) +#734 := (iff #756 #498) +#742 := (iff #755 true) +#748 := [rewrite]: #742 +#730 := [monotonicity #748]: #734 +#732 := [trans #730 #504]: #731 +#745 := (iff #743 #746) +#747 := [rewrite]: #745 +#737 := [monotonicity #747 #732]: #736 +#712 := [trans #737 #733]: #739 +#719 := [monotonicity #712]: #718 +#721 := [trans #719 #725]: #718 +#607 := [quant-inst]: #714 +#722 := [mp #607 #721]: #713 +#641 := [unit-resolution #722 #953]: #746 +#620 := (not #746) +#623 := (or #620 #723) +#627 := [th-lemma]: #623 +#629 := [unit-resolution #627 #641]: #723 +#726 := (>= #130 1::int) +#630 := (or #620 #726) +#628 := [th-lemma]: #630 +#631 := [unit-resolution #628 #641]: #726 +#611 := [th-lemma #631 #629 #610 #612]: #618 +#587 := [monotonicity #611]: #605 +#596 := [symm #587]: #595 +#581 := (= #103 #127) +decl uf_5 :: (-> T4 T3) +decl uf_8 :: T4 +#33 := uf_8 +#34 := (uf_5 uf_8) +#822 := (uf_4 #34) +#824 := (+ #130 #822) +#666 := (uf_3 #824) +#593 := (= #666 #127) +#589 := (= #127 #666) +#624 := (= 1::int #824) +#619 := (= #824 1::int) +#789 := (<= #822 0::int) +#787 := (= #822 0::int) +#632 := (= #822 #858) +#35 := (= #34 #31) +#162 := (= #31 #34) +#163 := (iff #35 #162) +#164 := [rewrite]: #163 +#160 := [asserted]: #35 +#167 := [mp #160 #164]: #162 +#662 := [symm #167]: #35 +#633 := [monotonicity #662]: #632 +#634 := [trans #633 #644]: #787 +#635 := (not #787) +#637 := (or #635 #789) +#638 := [th-lemma]: #637 +#639 := [unit-resolution #638 #634]: #789 +#781 := (>= #822 0::int) +#481 := (or #635 #781) +#640 := [th-lemma]: #481 +#636 := [unit-resolution #640 #634]: #781 +#622 := [th-lemma #631 #629 #636 #639]: #619 +#625 := [symm #622]: #624 +#590 := [monotonicity #625]: #589 +#594 := [symm #590]: #593 +#579 := (= #103 #666) +decl uf_6 :: (-> int T4 T4) +#539 := (uf_6 3::int uf_8) +#836 := (uf_5 #539) +#810 := (= #836 #666) +#813 := (= #666 #836) +#20 := (:var 0 T4) +#19 := (:var 1 int) +#21 := (uf_6 #19 #20) +#872 := (pattern #21) +#23 := (uf_5 #20) +#24 := (uf_4 #23) +#146 := (+ #24 #130) +#150 := (uf_3 #146) +#22 := (uf_5 #21) +#153 := (= #22 #150) +#873 := (forall (vars (?x3 int) (?x4 T4)) (:pat #872) #153) +#156 := (forall (vars (?x3 int) (?x4 T4)) #153) +#876 := (iff #156 #873) +#874 := (iff #153 #153) +#875 := [refl]: #874 +#877 := [quant-intro #875]: #876 +#328 := (~ #156 #156) +#364 := (~ #153 #153) +#365 := [refl]: #364 +#326 := [nnf-pos #365]: #328 +#25 := (+ #24 #14) +#26 := (uf_3 #25) +#27 := (= #22 #26) +#28 := (forall (vars (?x3 int) (?x4 T4)) #27) +#157 := (iff #28 #156) +#154 := (iff #27 #153) +#151 := (= #26 #150) +#148 := (= #25 #146) +#149 := [monotonicity #132]: #148 +#152 := [monotonicity #149]: #151 +#155 := [monotonicity #152]: #154 +#158 := [quant-intro #155]: #157 +#145 := [asserted]: #28 +#161 := [mp #145 #158]: #156 +#366 := [mp~ #161 #326]: #156 +#878 := [mp #366 #877]: #873 +#809 := (not #873) +#816 := (or #809 #813) +#817 := (+ #822 #130) +#818 := (uf_3 #817) +#823 := (= #836 #818) +#645 := (or #809 #823) +#648 := (iff #645 #816) +#802 := (iff #816 #816) +#804 := [rewrite]: #802 +#814 := (iff #823 #813) +#807 := (iff #810 #813) +#808 := [rewrite]: #807 +#811 := (iff #823 #810) +#667 := (= #818 #666) +#819 := (= #817 #824) +#825 := [rewrite]: #819 +#668 := [monotonicity #825]: #667 +#812 := [monotonicity #668]: #811 +#815 := [trans #812 #808]: #814 +#801 := [monotonicity #815]: #648 +#805 := [trans #801 #804]: #648 +#647 := [quant-inst]: #645 +#803 := [mp #647 #805]: #816 +#658 := [unit-resolution #803 #878]: #813 +#592 := [symm #658]: #810 +#600 := (= #103 #836) +decl uf_14 :: (-> T4 T3) +#540 := (uf_14 #539) +#548 := (= #540 #836) +#69 := (uf_14 #20) +#916 := (pattern #69) +#915 := (pattern #23) +#230 := (= #23 #69) +#917 := (forall (vars (?x14 T4)) (:pat #915 #916) #230) +#234 := (forall (vars (?x14 T4)) #230) +#920 := (iff #234 #917) +#918 := (iff #230 #230) +#919 := [refl]: #918 +#921 := [quant-intro #919]: #920 +#343 := (~ #234 #234) +#378 := (~ #230 #230) +#379 := [refl]: #378 +#344 := [nnf-pos #379]: #343 +#70 := (= #69 #23) +#71 := (forall (vars (?x14 T4)) #70) +#235 := (iff #71 #234) +#232 := (iff #70 #230) +#233 := [rewrite]: #232 +#236 := [quant-intro #233]: #235 +#229 := [asserted]: #71 +#239 := [mp #229 #236]: #234 +#380 := [mp~ #239 #344]: #234 +#922 := [mp #380 #921]: #917 +#541 := (not #917) +#828 := (or #541 #548) +#833 := (= #836 #540) +#829 := (or #541 #833) +#826 := (iff #829 #828) +#827 := (iff #828 #828) +#831 := [rewrite]: #827 +#549 := (iff #833 #548) +#550 := [rewrite]: #549 +#830 := [monotonicity #550]: #826 +#820 := [trans #830 #831]: #826 +#543 := [quant-inst]: #829 +#821 := [mp #543 #820]: #828 +#657 := [unit-resolution #821 #922]: #548 +#521 := (= #103 #540) +#75 := (uf_6 #72 uf_8) +#924 := (pattern #75) +#73 := (uf_16 #72) +#923 := (pattern #73) +#76 := (uf_14 #75) +#74 := (uf_15 #73) +#77 := (= #74 #76) +#925 := (forall (vars (?x15 int)) (:pat #923 #924) #77) +#78 := (forall (vars (?x15 int)) #77) +#928 := (iff #78 #925) +#926 := (iff #77 #77) +#927 := [refl]: #926 +#929 := [quant-intro #927]: #928 +#345 := (~ #78 #78) +#381 := (~ #77 #77) +#382 := [refl]: #381 +#346 := [nnf-pos #382]: #345 +#237 := [asserted]: #78 +#383 := [mp~ #237 #346]: #78 +#930 := [mp #383 #929]: #925 +#515 := (not #925) +#646 := (or #515 #521) +#853 := [quant-inst]: #646 +#603 := [unit-resolution #853 #930]: #521 +#577 := [trans #603 #657]: #600 +#580 := [trans #577 #592]: #579 +#582 := [trans #580 #594]: #581 +#584 := [trans #582 #596]: #583 +#578 := [trans #584 #598]: #585 +#571 := [trans #578 #615]: #586 +#572 := [trans #571 #588]: #107 +#108 := (not #107) +#325 := [asserted]: #108 +[unit-resolution #325 #572]: false +unsat diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_nat_arith_01 --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_nat_arith_01 Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,13 @@ +(benchmark Isabelle +:extrasorts ( T1) +:extrafuns ( + (uf_1 Int T1) + (uf_2 T1 Int) + (uf_3 T1) + ) +:assumption (forall (?x1 T1) (= (uf_1 (uf_2 ?x1)) ?x1)) +:assumption (forall (?x2 Int) (implies (<= 0 ?x2) (= (uf_2 (uf_1 ?x2)) ?x2))) +:assumption (forall (?x3 Int) (implies (< ?x3 0) (= (uf_2 (uf_1 ?x3)) 0))) +:assumption (= (uf_1 (* 2 (uf_2 uf_3))) (uf_1 1)) +:formula true +) diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_nat_arith_01.proof --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_nat_arith_01.proof Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,246 @@ +#2 := false +#9 := 0::int +decl uf_2 :: (-> T1 int) +decl uf_1 :: (-> int T1) +decl uf_3 :: T1 +#22 := uf_3 +#23 := (uf_2 uf_3) +#21 := 2::int +#24 := (* 2::int #23) +#25 := (uf_1 #24) +#293 := (uf_2 #25) +#292 := -1::int +#296 := (* -1::int #293) +#275 := (+ #24 #296) +#258 := (<= #275 0::int) +#611 := (= #275 0::int) +#204 := (>= #24 0::int) +#596 := (= #293 0::int) +#541 := (not #596) +#300 := (<= #293 0::int) +#460 := (not #300) +#26 := 1::int +#570 := (>= #293 1::int) +#569 := (= #293 1::int) +#27 := (uf_1 1::int) +#318 := (uf_2 #27) +#311 := (= #318 1::int) +#10 := (:var 0 int) +#12 := (uf_1 #10) +#627 := (pattern #12) +#70 := (>= #10 0::int) +#71 := (not #70) +#13 := (uf_2 #12) +#52 := (= #10 #13) +#77 := (or #52 #71) +#628 := (forall (vars (?x2 int)) (:pat #627) #77) +#82 := (forall (vars (?x2 int)) #77) +#631 := (iff #82 #628) +#629 := (iff #77 #77) +#630 := [refl]: #629 +#632 := [quant-intro #630]: #631 +#132 := (~ #82 #82) +#144 := (~ #77 #77) +#145 := [refl]: #144 +#130 := [nnf-pos #145]: #132 +#14 := (= #13 #10) +#11 := (<= 0::int #10) +#15 := (implies #11 #14) +#16 := (forall (vars (?x2 int)) #15) +#85 := (iff #16 #82) +#59 := (not #11) +#60 := (or #59 #52) +#65 := (forall (vars (?x2 int)) #60) +#83 := (iff #65 #82) +#80 := (iff #60 #77) +#74 := (or #71 #52) +#78 := (iff #74 #77) +#79 := [rewrite]: #78 +#75 := (iff #60 #74) +#72 := (iff #59 #71) +#68 := (iff #11 #70) +#69 := [rewrite]: #68 +#73 := [monotonicity #69]: #72 +#76 := [monotonicity #73]: #75 +#81 := [trans #76 #79]: #80 +#84 := [quant-intro #81]: #83 +#66 := (iff #16 #65) +#63 := (iff #15 #60) +#56 := (implies #11 #52) +#61 := (iff #56 #60) +#62 := [rewrite]: #61 +#57 := (iff #15 #56) +#54 := (iff #14 #52) +#55 := [rewrite]: #54 +#58 := [monotonicity #55]: #57 +#64 := [trans #58 #62]: #63 +#67 := [quant-intro #64]: #66 +#86 := [trans #67 #84]: #85 +#51 := [asserted]: #16 +#87 := [mp #51 #86]: #82 +#146 := [mp~ #87 #130]: #82 +#633 := [mp #146 #632]: #628 +#612 := (not #628) +#575 := (or #612 #311) +#316 := (>= 1::int 0::int) +#317 := (not #316) +#211 := (= 1::int #318) +#588 := (or #211 #317) +#576 := (or #612 #588) +#572 := (iff #576 #575) +#578 := (iff #575 #575) +#573 := [rewrite]: #578 +#585 := (iff #588 #311) +#583 := (or #311 false) +#584 := (iff #583 #311) +#581 := [rewrite]: #584 +#297 := (iff #588 #583) +#304 := (iff #317 false) +#1 := true +#587 := (not true) +#302 := (iff #587 false) +#303 := [rewrite]: #302 +#591 := (iff #317 #587) +#586 := (iff #316 true) +#590 := [rewrite]: #586 +#301 := [monotonicity #590]: #591 +#582 := [trans #301 #303]: #304 +#589 := (iff #211 #311) +#312 := [rewrite]: #589 +#580 := [monotonicity #312 #582]: #297 +#574 := [trans #580 #581]: #585 +#577 := [monotonicity #574]: #572 +#579 := [trans #577 #573]: #572 +#571 := [quant-inst]: #576 +#420 := [mp #571 #579]: #575 +#437 := [unit-resolution #420 #633]: #311 +#452 := (= #293 #318) +#28 := (= #25 #27) +#129 := [asserted]: #28 +#454 := [monotonicity #129]: #452 +#455 := [trans #454 #437]: #569 +#448 := (not #569) +#456 := (or #448 #570) +#457 := [th-lemma]: #456 +#458 := [unit-resolution #457 #455]: #570 +#459 := (not #570) +#553 := (or #459 #460) +#550 := [th-lemma]: #553 +#554 := [unit-resolution #550 #458]: #460 +#543 := (or #541 #300) +#535 := [th-lemma]: #543 +#532 := [unit-resolution #535 #554]: #541 +#598 := (or #204 #596) +#18 := (= #13 0::int) +#118 := (or #18 #70) +#634 := (forall (vars (?x3 int)) (:pat #627) #118) +#123 := (forall (vars (?x3 int)) #118) +#637 := (iff #123 #634) +#635 := (iff #118 #118) +#636 := [refl]: #635 +#638 := [quant-intro #636]: #637 +#133 := (~ #123 #123) +#147 := (~ #118 #118) +#148 := [refl]: #147 +#134 := [nnf-pos #148]: #133 +#17 := (< #10 0::int) +#19 := (implies #17 #18) +#20 := (forall (vars (?x3 int)) #19) +#126 := (iff #20 #123) +#89 := (= 0::int #13) +#95 := (not #17) +#96 := (or #95 #89) +#101 := (forall (vars (?x3 int)) #96) +#124 := (iff #101 #123) +#121 := (iff #96 #118) +#115 := (or #70 #18) +#119 := (iff #115 #118) +#120 := [rewrite]: #119 +#116 := (iff #96 #115) +#113 := (iff #89 #18) +#114 := [rewrite]: #113 +#111 := (iff #95 #70) +#106 := (not #71) +#109 := (iff #106 #70) +#110 := [rewrite]: #109 +#107 := (iff #95 #106) +#104 := (iff #17 #71) +#105 := [rewrite]: #104 +#108 := [monotonicity #105]: #107 +#112 := [trans #108 #110]: #111 +#117 := [monotonicity #112 #114]: #116 +#122 := [trans #117 #120]: #121 +#125 := [quant-intro #122]: #124 +#102 := (iff #20 #101) +#99 := (iff #19 #96) +#92 := (implies #17 #89) +#97 := (iff #92 #96) +#98 := [rewrite]: #97 +#93 := (iff #19 #92) +#90 := (iff #18 #89) +#91 := [rewrite]: #90 +#94 := [monotonicity #91]: #93 +#100 := [trans #94 #98]: #99 +#103 := [quant-intro #100]: #102 +#127 := [trans #103 #125]: #126 +#88 := [asserted]: #20 +#128 := [mp #88 #127]: #123 +#149 := [mp~ #128 #134]: #123 +#639 := [mp #149 #638]: #634 +#595 := (not #634) +#601 := (or #595 #204 #596) +#597 := (or #596 #204) +#238 := (or #595 #597) +#606 := (iff #238 #601) +#604 := (or #595 #598) +#605 := (iff #604 #601) +#603 := [rewrite]: #605 +#243 := (iff #238 #604) +#599 := (iff #597 #598) +#600 := [rewrite]: #599 +#244 := [monotonicity #600]: #243 +#592 := [trans #244 #603]: #606 +#602 := [quant-inst]: #238 +#593 := [mp #602 #592]: #601 +#534 := [unit-resolution #593 #639]: #598 +#544 := [unit-resolution #534 #532]: #204 +#290 := (not #204) +#281 := (or #290 #611) +#618 := (or #612 #290 #611) +#294 := (= #24 #293) +#295 := (or #294 #290) +#608 := (or #612 #295) +#594 := (iff #608 #618) +#272 := (or #612 #281) +#610 := (iff #272 #618) +#252 := [rewrite]: #610 +#609 := (iff #608 #272) +#616 := (iff #295 #281) +#400 := (or #611 #290) +#614 := (iff #400 #281) +#615 := [rewrite]: #614 +#607 := (iff #295 #400) +#613 := (iff #294 #611) +#269 := [rewrite]: #613 +#280 := [monotonicity #269]: #607 +#617 := [trans #280 #615]: #616 +#268 := [monotonicity #617]: #609 +#256 := [trans #268 #252]: #594 +#267 := [quant-inst]: #608 +#257 := [mp #267 #256]: #618 +#545 := [unit-resolution #257 #633]: #281 +#546 := [unit-resolution #545 #544]: #611 +#542 := (not #611) +#547 := (or #542 #258) +#536 := [th-lemma]: #547 +#537 := [unit-resolution #536 #546]: #258 +#259 := (>= #275 0::int) +#538 := (or #542 #259) +#539 := [th-lemma]: #538 +#533 := [unit-resolution #539 #546]: #259 +#563 := (<= #293 1::int) +#540 := (or #448 #563) +#524 := [th-lemma]: #540 +#525 := [unit-resolution #524 #455]: #563 +[th-lemma #458 #525 #533 #537]: false +unsat diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_nat_arith_02 --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_nat_arith_02 Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,14 @@ +(benchmark Isabelle +:extrasorts ( T1) +:extrafuns ( + (uf_1 Int T1) + (uf_2 T1 Int) + (uf_3 T1) + ) +:assumption (forall (?x1 T1) (= (uf_1 (uf_2 ?x1)) ?x1)) +:assumption (forall (?x2 Int) (implies (<= 0 ?x2) (= (uf_2 (uf_1 ?x2)) ?x2))) +:assumption (forall (?x3 Int) (implies (< ?x3 0) (= (uf_2 (uf_1 ?x3)) 0))) +:assumption (< (uf_2 uf_3) 3) +:assumption (not (< (uf_2 (uf_1 (* 2 (uf_2 uf_3)))) 7)) +:formula true +) diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_nat_arith_02.proof --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_nat_arith_02.proof Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,199 @@ +#2 := false +#23 := 3::int +decl uf_2 :: (-> T1 int) +decl uf_3 :: T1 +#21 := uf_3 +#22 := (uf_2 uf_3) +#137 := (>= #22 3::int) +#135 := (not #137) +#24 := (< #22 3::int) +#136 := (iff #24 #135) +#138 := [rewrite]: #136 +#132 := [asserted]: #24 +#139 := [mp #132 #138]: #135 +#9 := 0::int +decl uf_1 :: (-> int T1) +#25 := 2::int +#26 := (* 2::int #22) +#27 := (uf_1 #26) +#28 := (uf_2 #27) +#297 := -1::int +#633 := (* -1::int #28) +#635 := (+ #26 #633) +#278 := (>= #635 0::int) +#291 := (= #635 0::int) +#315 := (>= #26 0::int) +#279 := (= #28 0::int) +#627 := (not #279) +#624 := (<= #28 0::int) +#281 := (not #624) +#29 := 7::int +#143 := (>= #28 7::int) +#30 := (< #28 7::int) +#31 := (not #30) +#150 := (iff #31 #143) +#141 := (not #143) +#145 := (not #141) +#148 := (iff #145 #143) +#149 := [rewrite]: #148 +#146 := (iff #31 #145) +#142 := (iff #30 #141) +#144 := [rewrite]: #142 +#147 := [monotonicity #144]: #146 +#151 := [trans #147 #149]: #150 +#133 := [asserted]: #31 +#152 := [mp #133 #151]: #143 +#618 := (or #281 #141) +#265 := [th-lemma]: #618 +#266 := [unit-resolution #265 #152]: #281 +#625 := (or #627 #624) +#628 := [th-lemma]: #625 +#614 := [unit-resolution #628 #266]: #627 +#10 := (:var 0 int) +#12 := (uf_1 #10) +#649 := (pattern #12) +#73 := (>= #10 0::int) +#13 := (uf_2 #12) +#18 := (= #13 0::int) +#121 := (or #18 #73) +#656 := (forall (vars (?x3 int)) (:pat #649) #121) +#126 := (forall (vars (?x3 int)) #121) +#659 := (iff #126 #656) +#657 := (iff #121 #121) +#658 := [refl]: #657 +#660 := [quant-intro #658]: #659 +#154 := (~ #126 #126) +#170 := (~ #121 #121) +#171 := [refl]: #170 +#155 := [nnf-pos #171]: #154 +#17 := (< #10 0::int) +#19 := (implies #17 #18) +#20 := (forall (vars (?x3 int)) #19) +#129 := (iff #20 #126) +#92 := (= 0::int #13) +#98 := (not #17) +#99 := (or #98 #92) +#104 := (forall (vars (?x3 int)) #99) +#127 := (iff #104 #126) +#124 := (iff #99 #121) +#118 := (or #73 #18) +#122 := (iff #118 #121) +#123 := [rewrite]: #122 +#119 := (iff #99 #118) +#116 := (iff #92 #18) +#117 := [rewrite]: #116 +#114 := (iff #98 #73) +#74 := (not #73) +#109 := (not #74) +#112 := (iff #109 #73) +#113 := [rewrite]: #112 +#110 := (iff #98 #109) +#107 := (iff #17 #74) +#108 := [rewrite]: #107 +#111 := [monotonicity #108]: #110 +#115 := [trans #111 #113]: #114 +#120 := [monotonicity #115 #117]: #119 +#125 := [trans #120 #123]: #124 +#128 := [quant-intro #125]: #127 +#105 := (iff #20 #104) +#102 := (iff #19 #99) +#95 := (implies #17 #92) +#100 := (iff #95 #99) +#101 := [rewrite]: #100 +#96 := (iff #19 #95) +#93 := (iff #18 #92) +#94 := [rewrite]: #93 +#97 := [monotonicity #94]: #96 +#103 := [trans #97 #101]: #102 +#106 := [quant-intro #103]: #105 +#130 := [trans #106 #128]: #129 +#91 := [asserted]: #20 +#131 := [mp #91 #130]: #126 +#172 := [mp~ #131 #155]: #126 +#661 := [mp #172 #660]: #656 +#619 := (not #656) +#620 := (or #619 #279 #315) +#280 := (or #279 #315) +#621 := (or #619 #280) +#617 := (iff #621 #620) +#623 := [rewrite]: #617 +#622 := [quant-inst]: #621 +#260 := [mp #622 #623]: #620 +#615 := [unit-resolution #260 #661 #614]: #315 +#316 := (not #315) +#302 := (or #291 #316) +#55 := (= #10 #13) +#80 := (or #55 #74) +#650 := (forall (vars (?x2 int)) (:pat #649) #80) +#85 := (forall (vars (?x2 int)) #80) +#653 := (iff #85 #650) +#651 := (iff #80 #80) +#652 := [refl]: #651 +#654 := [quant-intro #652]: #653 +#153 := (~ #85 #85) +#167 := (~ #80 #80) +#168 := [refl]: #167 +#134 := [nnf-pos #168]: #153 +#14 := (= #13 #10) +#11 := (<= 0::int #10) +#15 := (implies #11 #14) +#16 := (forall (vars (?x2 int)) #15) +#88 := (iff #16 #85) +#62 := (not #11) +#63 := (or #62 #55) +#68 := (forall (vars (?x2 int)) #63) +#86 := (iff #68 #85) +#83 := (iff #63 #80) +#77 := (or #74 #55) +#81 := (iff #77 #80) +#82 := [rewrite]: #81 +#78 := (iff #63 #77) +#75 := (iff #62 #74) +#71 := (iff #11 #73) +#72 := [rewrite]: #71 +#76 := [monotonicity #72]: #75 +#79 := [monotonicity #76]: #78 +#84 := [trans #79 #82]: #83 +#87 := [quant-intro #84]: #86 +#69 := (iff #16 #68) +#66 := (iff #15 #63) +#59 := (implies #11 #55) +#64 := (iff #59 #63) +#65 := [rewrite]: #64 +#60 := (iff #15 #59) +#57 := (iff #14 #55) +#58 := [rewrite]: #57 +#61 := [monotonicity #58]: #60 +#67 := [trans #61 #65]: #66 +#70 := [quant-intro #67]: #69 +#89 := [trans #70 #87]: #88 +#54 := [asserted]: #16 +#90 := [mp #54 #89]: #85 +#169 := [mp~ #90 #134]: #85 +#655 := [mp #169 #654]: #650 +#637 := (not #650) +#638 := (or #637 #291 #316) +#314 := (= #26 #28) +#318 := (or #314 #316) +#639 := (or #637 #318) +#290 := (iff #639 #638) +#640 := (or #637 #302) +#294 := (iff #640 #638) +#631 := [rewrite]: #294 +#630 := (iff #639 #640) +#303 := (iff #318 #302) +#422 := (iff #314 #291) +#629 := [rewrite]: #422 +#636 := [monotonicity #629]: #303 +#289 := [monotonicity #636]: #630 +#632 := [trans #289 #631]: #290 +#634 := [quant-inst]: #639 +#274 := [mp #634 #632]: #638 +#322 := [unit-resolution #274 #655]: #302 +#337 := [unit-resolution #322 #615]: #291 +#338 := (not #291) +#339 := (or #338 #278) +#340 := [th-lemma]: #339 +#232 := [unit-resolution #340 #337]: #278 +[th-lemma #152 #232 #139]: false +unsat diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_nat_arith_03 --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_nat_arith_03 Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,13 @@ +(benchmark Isabelle +:extrasorts ( T1) +:extrafuns ( + (uf_1 Int T1) + (uf_2 T1 Int) + (uf_3 T1) + ) +:assumption (forall (?x1 T1) (= (uf_1 (uf_2 ?x1)) ?x1)) +:assumption (forall (?x2 Int) (implies (<= 0 ?x2) (= (uf_2 (uf_1 ?x2)) ?x2))) +:assumption (forall (?x3 Int) (implies (< ?x3 0) (= (uf_2 (uf_1 ?x3)) 0))) +:assumption (not (let (?x4 (uf_1 (+ 1 (uf_2 uf_3)))) (< (uf_2 (uf_1 (* 0 (uf_2 ?x4)))) (uf_2 (uf_1 (- (uf_2 ?x4) (uf_2 uf_3))))))) +:formula true +) diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_nat_arith_03.proof --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_nat_arith_03.proof Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,367 @@ +#2 := false +#9 := 0::int +decl uf_2 :: (-> T1 int) +decl uf_3 :: T1 +#22 := uf_3 +#23 := (uf_2 uf_3) +#469 := (= #23 0::int) +decl uf_1 :: (-> int T1) +#251 := (uf_1 #23) +#557 := (uf_2 #251) +#558 := (= #557 0::int) +#556 := (>= #23 0::int) +#477 := (not #556) +#144 := -1::int +#348 := (>= #23 -1::int) +#628 := (not #348) +#21 := 1::int +#24 := (+ 1::int #23) +#25 := (uf_1 #24) +#26 := (uf_2 #25) +#632 := (* -1::int #26) +#636 := (+ #23 #632) +#633 := (= #636 -1::int) +#471 := (not #633) +#613 := (<= #636 -1::int) +#527 := (not #613) +#145 := (* -1::int #23) +#146 := (+ #145 #26) +#149 := (uf_1 #146) +#152 := (uf_2 #149) +#504 := (+ #632 #152) +#505 := (+ #23 #504) +#573 := (>= #505 0::int) +#502 := (= #505 0::int) +#599 := (<= #636 0::int) +#526 := [hypothesis]: #613 +#491 := (or #527 #599) +#515 := [th-lemma]: #491 +#516 := [unit-resolution #515 #526]: #599 +#587 := (not #599) +#578 := (or #502 #587) +#10 := (:var 0 int) +#12 := (uf_1 #10) +#673 := (pattern #12) +#76 := (>= #10 0::int) +#77 := (not #76) +#13 := (uf_2 #12) +#58 := (= #10 #13) +#83 := (or #58 #77) +#674 := (forall (vars (?x2 int)) (:pat #673) #83) +#88 := (forall (vars (?x2 int)) #83) +#677 := (iff #88 #674) +#675 := (iff #83 #83) +#676 := [refl]: #675 +#678 := [quant-intro #676]: #677 +#179 := (~ #88 #88) +#191 := (~ #83 #83) +#192 := [refl]: #191 +#177 := [nnf-pos #192]: #179 +#14 := (= #13 #10) +#11 := (<= 0::int #10) +#15 := (implies #11 #14) +#16 := (forall (vars (?x2 int)) #15) +#91 := (iff #16 #88) +#65 := (not #11) +#66 := (or #65 #58) +#71 := (forall (vars (?x2 int)) #66) +#89 := (iff #71 #88) +#86 := (iff #66 #83) +#80 := (or #77 #58) +#84 := (iff #80 #83) +#85 := [rewrite]: #84 +#81 := (iff #66 #80) +#78 := (iff #65 #77) +#74 := (iff #11 #76) +#75 := [rewrite]: #74 +#79 := [monotonicity #75]: #78 +#82 := [monotonicity #79]: #81 +#87 := [trans #82 #85]: #86 +#90 := [quant-intro #87]: #89 +#72 := (iff #16 #71) +#69 := (iff #15 #66) +#62 := (implies #11 #58) +#67 := (iff #62 #66) +#68 := [rewrite]: #67 +#63 := (iff #15 #62) +#60 := (iff #14 #58) +#61 := [rewrite]: #60 +#64 := [monotonicity #61]: #63 +#70 := [trans #64 #68]: #69 +#73 := [quant-intro #70]: #72 +#92 := [trans #73 #90]: #91 +#57 := [asserted]: #16 +#93 := [mp #57 #92]: #88 +#193 := [mp~ #93 #177]: #88 +#679 := [mp #193 #678]: #674 +#644 := (not #674) +#591 := (or #644 #502 #587) +#498 := (>= #146 0::int) +#500 := (not #498) +#501 := (= #146 #152) +#494 := (or #501 #500) +#592 := (or #644 #494) +#579 := (iff #592 #591) +#593 := (or #644 #578) +#584 := (iff #593 #591) +#585 := [rewrite]: #584 +#582 := (iff #592 #593) +#580 := (iff #494 #578) +#589 := (iff #500 #587) +#596 := (iff #498 #599) +#600 := [rewrite]: #596 +#581 := [monotonicity #600]: #589 +#503 := (iff #501 #502) +#506 := [rewrite]: #503 +#590 := [monotonicity #506 #581]: #580 +#583 := [monotonicity #590]: #582 +#586 := [trans #583 #585]: #579 +#588 := [quant-inst]: #592 +#570 := [mp #588 #586]: #591 +#511 := [unit-resolution #570 #679]: #578 +#517 := [unit-resolution #511 #516]: #502 +#485 := (not #502) +#492 := (or #485 #573) +#451 := [th-lemma]: #492 +#482 := [unit-resolution #451 #517]: #573 +#554 := (<= #152 0::int) +#163 := (* -1::int #152) +#138 := (uf_1 0::int) +#141 := (uf_2 #138) +#164 := (+ #141 #163) +#162 := (>= #164 0::int) +#30 := (- #26 #23) +#31 := (uf_1 #30) +#32 := (uf_2 #31) +#27 := (* 0::int #26) +#28 := (uf_1 #27) +#29 := (uf_2 #28) +#33 := (< #29 #32) +#34 := (not #33) +#174 := (iff #34 #162) +#155 := (< #141 #152) +#158 := (not #155) +#172 := (iff #158 #162) +#161 := (not #162) +#167 := (not #161) +#170 := (iff #167 #162) +#171 := [rewrite]: #170 +#168 := (iff #158 #167) +#165 := (iff #155 #161) +#166 := [rewrite]: #165 +#169 := [monotonicity #166]: #168 +#173 := [trans #169 #171]: #172 +#159 := (iff #34 #158) +#156 := (iff #33 #155) +#153 := (= #32 #152) +#150 := (= #31 #149) +#147 := (= #30 #146) +#148 := [rewrite]: #147 +#151 := [monotonicity #148]: #150 +#154 := [monotonicity #151]: #153 +#142 := (= #29 #141) +#139 := (= #28 #138) +#136 := (= #27 0::int) +#137 := [rewrite]: #136 +#140 := [monotonicity #137]: #139 +#143 := [monotonicity #140]: #142 +#157 := [monotonicity #143 #154]: #156 +#160 := [monotonicity #157]: #159 +#175 := [trans #160 #173]: #174 +#135 := [asserted]: #34 +#176 := [mp #135 #175]: #162 +#651 := (<= #141 0::int) +#662 := (= #141 0::int) +#645 := (or #644 #662) +#316 := (>= 0::int 0::int) +#446 := (not #316) +#328 := (= 0::int #141) +#660 := (or #328 #446) +#646 := (or #644 #660) +#647 := (iff #646 #645) +#648 := (iff #645 #645) +#650 := [rewrite]: #648 +#642 := (iff #660 #662) +#640 := (or #662 false) +#305 := (iff #640 #662) +#306 := [rewrite]: #305 +#303 := (iff #660 #640) +#656 := (iff #446 false) +#1 := true +#654 := (not true) +#655 := (iff #654 false) +#315 := [rewrite]: #655 +#314 := (iff #446 #654) +#658 := (iff #316 true) +#664 := [rewrite]: #658 +#319 := [monotonicity #664]: #314 +#299 := [trans #319 #315]: #656 +#661 := (iff #328 #662) +#663 := [rewrite]: #661 +#304 := [monotonicity #663 #299]: #303 +#643 := [trans #304 #306]: #642 +#285 := [monotonicity #643]: #647 +#290 := [trans #285 #650]: #647 +#641 := [quant-inst]: #646 +#291 := [mp #641 #290]: #645 +#484 := [unit-resolution #291 #679]: #662 +#486 := (not #662) +#493 := (or #486 #651) +#495 := [th-lemma]: #493 +#496 := [unit-resolution #495 #484]: #651 +#497 := (not #651) +#507 := (or #554 #497 #161) +#487 := [th-lemma]: #507 +#508 := [unit-resolution #487 #496 #176]: #554 +#463 := [th-lemma #508 #526 #482]: false +#464 := [lemma #463]: #527 +#472 := (or #471 #613) +#473 := [th-lemma]: #472 +#474 := [unit-resolution #473 #464]: #471 +#631 := (or #628 #633) +#618 := (or #644 #628 #633) +#634 := (>= #24 0::int) +#635 := (not #634) +#357 := (= #24 #26) +#358 := (or #357 #635) +#623 := (or #644 #358) +#610 := (iff #623 #618) +#619 := (or #644 #631) +#467 := (iff #619 #618) +#468 := [rewrite]: #467 +#625 := (iff #623 #619) +#622 := (iff #358 #631) +#626 := (or #633 #628) +#620 := (iff #626 #631) +#621 := [rewrite]: #620 +#630 := (iff #358 #626) +#629 := (iff #635 #628) +#349 := (iff #634 #348) +#350 := [rewrite]: #349 +#344 := [monotonicity #350]: #629 +#637 := (iff #357 #633) +#347 := [rewrite]: #637 +#627 := [monotonicity #347 #344]: #630 +#617 := [trans #627 #621]: #622 +#466 := [monotonicity #617]: #625 +#611 := [trans #466 #468]: #610 +#624 := [quant-inst]: #623 +#612 := [mp #624 #611]: #618 +#475 := [unit-resolution #612 #679]: #631 +#476 := [unit-resolution #475 #474]: #628 +#478 := (or #477 #348) +#479 := [th-lemma]: #478 +#480 := [unit-resolution #479 #476]: #477 +#560 := (or #556 #558) +#18 := (= #13 0::int) +#124 := (or #18 #76) +#680 := (forall (vars (?x3 int)) (:pat #673) #124) +#129 := (forall (vars (?x3 int)) #124) +#683 := (iff #129 #680) +#681 := (iff #124 #124) +#682 := [refl]: #681 +#684 := [quant-intro #682]: #683 +#180 := (~ #129 #129) +#194 := (~ #124 #124) +#195 := [refl]: #194 +#181 := [nnf-pos #195]: #180 +#17 := (< #10 0::int) +#19 := (implies #17 #18) +#20 := (forall (vars (?x3 int)) #19) +#132 := (iff #20 #129) +#95 := (= 0::int #13) +#101 := (not #17) +#102 := (or #101 #95) +#107 := (forall (vars (?x3 int)) #102) +#130 := (iff #107 #129) +#127 := (iff #102 #124) +#121 := (or #76 #18) +#125 := (iff #121 #124) +#126 := [rewrite]: #125 +#122 := (iff #102 #121) +#119 := (iff #95 #18) +#120 := [rewrite]: #119 +#117 := (iff #101 #76) +#112 := (not #77) +#115 := (iff #112 #76) +#116 := [rewrite]: #115 +#113 := (iff #101 #112) +#110 := (iff #17 #77) +#111 := [rewrite]: #110 +#114 := [monotonicity #111]: #113 +#118 := [trans #114 #116]: #117 +#123 := [monotonicity #118 #120]: #122 +#128 := [trans #123 #126]: #127 +#131 := [quant-intro #128]: #130 +#108 := (iff #20 #107) +#105 := (iff #19 #102) +#98 := (implies #17 #95) +#103 := (iff #98 #102) +#104 := [rewrite]: #103 +#99 := (iff #19 #98) +#96 := (iff #18 #95) +#97 := [rewrite]: #96 +#100 := [monotonicity #97]: #99 +#106 := [trans #100 #104]: #105 +#109 := [quant-intro #106]: #108 +#133 := [trans #109 #131]: #132 +#94 := [asserted]: #20 +#134 := [mp #94 #133]: #129 +#196 := [mp~ #134 #181]: #129 +#685 := [mp #196 #684]: #680 +#604 := (not #680) +#562 := (or #604 #556 #558) +#559 := (or #558 #556) +#540 := (or #604 #559) +#542 := (iff #540 #562) +#543 := (or #604 #560) +#546 := (iff #543 #562) +#547 := [rewrite]: #546 +#544 := (iff #540 #543) +#561 := (iff #559 #560) +#551 := [rewrite]: #561 +#545 := [monotonicity #551]: #544 +#548 := [trans #545 #547]: #542 +#541 := [quant-inst]: #540 +#534 := [mp #541 #548]: #562 +#465 := [unit-resolution #534 #685]: #560 +#481 := [unit-resolution #465 #480]: #558 +#443 := (= #23 #557) +#337 := (= uf_3 #251) +#4 := (:var 0 T1) +#5 := (uf_2 #4) +#665 := (pattern #5) +#6 := (uf_1 #5) +#51 := (= #4 #6) +#666 := (forall (vars (?x1 T1)) (:pat #665) #51) +#54 := (forall (vars (?x1 T1)) #51) +#667 := (iff #54 #666) +#669 := (iff #666 #666) +#670 := [rewrite]: #669 +#668 := [rewrite]: #667 +#671 := [trans #668 #670]: #667 +#188 := (~ #54 #54) +#186 := (~ #51 #51) +#187 := [refl]: #186 +#189 := [nnf-pos #187]: #188 +#7 := (= #6 #4) +#8 := (forall (vars (?x1 T1)) #7) +#55 := (iff #8 #54) +#52 := (iff #7 #51) +#53 := [rewrite]: #52 +#56 := [quant-intro #53]: #55 +#50 := [asserted]: #8 +#59 := [mp #50 #56]: #54 +#190 := [mp~ #59 #189]: #54 +#672 := [mp #190 #671]: #666 +#252 := (not #666) +#342 := (or #252 #337) +#339 := [quant-inst]: #342 +#442 := [unit-resolution #339 #672]: #337 +#450 := [monotonicity #442]: #443 +#452 := [trans #450 #481]: #469 +#453 := (not #469) +#454 := (or #453 #556) +#456 := [th-lemma]: #454 +[unit-resolution #456 #480 #452]: false +unsat diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_nat_arith_04 --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_nat_arith_04 Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,13 @@ +(benchmark Isabelle +:extrasorts ( T1 T2) +:extrafuns ( + (uf_1 Int T1) + (uf_2 T1 Int) + (uf_3 T1) + ) +:assumption (forall (?x1 T1) (= (uf_1 (uf_2 ?x1)) ?x1)) +:assumption (forall (?x2 Int) (implies (<= 0 ?x2) (= (uf_2 (uf_1 ?x2)) ?x2))) +:assumption (forall (?x3 Int) (implies (< ?x3 0) (= (uf_2 (uf_1 ?x3)) 0))) +:assumption (not (let (?x4 (uf_1 (+ 1 (uf_2 uf_3)))) (flet ($x5 (if_then_else (< 0 (uf_2 ?x4)) true false)) (or (iff $x5 (= (uf_1 (- (uf_2 ?x4) 1)) uf_3)) $x5)))) +:formula true +) diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_nat_arith_04.proof --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_nat_arith_04.proof Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,302 @@ +#2 := false +#9 := 0::int +decl uf_2 :: (-> T1 int) +decl uf_1 :: (-> int T1) +decl uf_3 :: T1 +#22 := uf_3 +#23 := (uf_2 uf_3) +#21 := 1::int +#24 := (+ 1::int #23) +#25 := (uf_1 #24) +#26 := (uf_2 #25) +#138 := -1::int +#139 := (+ -1::int #26) +#142 := (uf_1 #139) +#289 := (uf_2 #142) +#383 := (* -1::int #289) +#542 := (+ #23 #383) +#544 := (>= #542 0::int) +#541 := (= #23 #289) +#148 := (= uf_3 #142) +#167 := (<= #26 0::int) +#168 := (not #167) +#174 := (iff #148 #168) +#189 := (not #174) +#220 := (iff #189 #148) +#210 := (not #148) +#215 := (not #210) +#218 := (iff #215 #148) +#219 := [rewrite]: #218 +#216 := (iff #189 #215) +#213 := (iff #174 #210) +#207 := (iff #148 false) +#211 := (iff #207 #210) +#212 := [rewrite]: #211 +#208 := (iff #174 #207) +#205 := (iff #168 false) +#1 := true +#200 := (not true) +#203 := (iff #200 false) +#204 := [rewrite]: #203 +#201 := (iff #168 #200) +#198 := (iff #167 true) +#179 := (or #168 #174) +#182 := (not #179) +#27 := (< 0::int #26) +#28 := (ite #27 true false) +#29 := (- #26 1::int) +#30 := (uf_1 #29) +#31 := (= #30 uf_3) +#32 := (iff #28 #31) +#33 := (or #32 #28) +#34 := (not #33) +#185 := (iff #34 #182) +#153 := (iff #27 #148) +#159 := (or #27 #153) +#164 := (not #159) +#183 := (iff #164 #182) +#180 := (iff #159 #179) +#177 := (iff #153 #174) +#171 := (iff #168 #148) +#175 := (iff #171 #174) +#176 := [rewrite]: #175 +#172 := (iff #153 #171) +#169 := (iff #27 #168) +#170 := [rewrite]: #169 +#173 := [monotonicity #170]: #172 +#178 := [trans #173 #176]: #177 +#181 := [monotonicity #170 #178]: #180 +#184 := [monotonicity #181]: #183 +#165 := (iff #34 #164) +#162 := (iff #33 #159) +#156 := (or #153 #27) +#160 := (iff #156 #159) +#161 := [rewrite]: #160 +#157 := (iff #33 #156) +#136 := (iff #28 #27) +#137 := [rewrite]: #136 +#154 := (iff #32 #153) +#151 := (iff #31 #148) +#145 := (= #142 uf_3) +#149 := (iff #145 #148) +#150 := [rewrite]: #149 +#146 := (iff #31 #145) +#143 := (= #30 #142) +#140 := (= #29 #139) +#141 := [rewrite]: #140 +#144 := [monotonicity #141]: #143 +#147 := [monotonicity #144]: #146 +#152 := [trans #147 #150]: #151 +#155 := [monotonicity #137 #152]: #154 +#158 := [monotonicity #155 #137]: #157 +#163 := [trans #158 #161]: #162 +#166 := [monotonicity #163]: #165 +#186 := [trans #166 #184]: #185 +#135 := [asserted]: #34 +#187 := [mp #135 #186]: #182 +#188 := [not-or-elim #187]: #167 +#199 := [iff-true #188]: #198 +#202 := [monotonicity #199]: #201 +#206 := [trans #202 #204]: #205 +#209 := [monotonicity #206]: #208 +#214 := [trans #209 #212]: #213 +#217 := [monotonicity #214]: #216 +#221 := [trans #217 #219]: #220 +#190 := [not-or-elim #187]: #189 +#222 := [mp #190 #221]: #148 +#624 := [monotonicity #222]: #541 +#618 := (not #541) +#625 := (or #618 #544) +#609 := [th-lemma]: #625 +#610 := [unit-resolution #609 #624]: #544 +#698 := (* -1::int #26) +#355 := (+ #23 #698) +#324 := (<= #355 -1::int) +#485 := (= #355 -1::int) +#367 := (>= #23 -1::int) +#533 := (>= #289 0::int) +#643 := (= #289 0::int) +#659 := (>= #26 1::int) +#656 := (not #659) +#612 := (or #656 #168) +#613 := [th-lemma]: #612 +#614 := [unit-resolution #613 #188]: #656 +#10 := (:var 0 int) +#12 := (uf_1 #10) +#712 := (pattern #12) +#76 := (>= #10 0::int) +#13 := (uf_2 #12) +#18 := (= #13 0::int) +#124 := (or #18 #76) +#719 := (forall (vars (?x3 int)) (:pat #712) #124) +#129 := (forall (vars (?x3 int)) #124) +#722 := (iff #129 #719) +#720 := (iff #124 #124) +#721 := [refl]: #720 +#723 := [quant-intro #721]: #722 +#229 := (~ #129 #129) +#227 := (~ #124 #124) +#228 := [refl]: #227 +#230 := [nnf-pos #228]: #229 +#17 := (< #10 0::int) +#19 := (implies #17 #18) +#20 := (forall (vars (?x3 int)) #19) +#132 := (iff #20 #129) +#95 := (= 0::int #13) +#101 := (not #17) +#102 := (or #101 #95) +#107 := (forall (vars (?x3 int)) #102) +#130 := (iff #107 #129) +#127 := (iff #102 #124) +#121 := (or #76 #18) +#125 := (iff #121 #124) +#126 := [rewrite]: #125 +#122 := (iff #102 #121) +#119 := (iff #95 #18) +#120 := [rewrite]: #119 +#117 := (iff #101 #76) +#77 := (not #76) +#112 := (not #77) +#115 := (iff #112 #76) +#116 := [rewrite]: #115 +#113 := (iff #101 #112) +#110 := (iff #17 #77) +#111 := [rewrite]: #110 +#114 := [monotonicity #111]: #113 +#118 := [trans #114 #116]: #117 +#123 := [monotonicity #118 #120]: #122 +#128 := [trans #123 #126]: #127 +#131 := [quant-intro #128]: #130 +#108 := (iff #20 #107) +#105 := (iff #19 #102) +#98 := (implies #17 #95) +#103 := (iff #98 #102) +#104 := [rewrite]: #103 +#99 := (iff #19 #98) +#96 := (iff #18 #95) +#97 := [rewrite]: #96 +#100 := [monotonicity #97]: #99 +#106 := [trans #100 #104]: #105 +#109 := [quant-intro #106]: #108 +#133 := [trans #109 #131]: #132 +#94 := [asserted]: #20 +#134 := [mp #94 #133]: #129 +#231 := [mp~ #134 #230]: #129 +#724 := [mp #231 #723]: #719 +#402 := (not #719) +#528 := (or #402 #643 #659) +#388 := (>= #139 0::int) +#644 := (or #643 #388) +#529 := (or #402 #644) +#522 := (iff #529 #528) +#642 := (or #643 #659) +#636 := (or #402 #642) +#634 := (iff #636 #528) +#637 := [rewrite]: #634 +#538 := (iff #529 #636) +#645 := (iff #644 #642) +#660 := (iff #388 #659) +#661 := [rewrite]: #660 +#527 := [monotonicity #661]: #645 +#633 := [monotonicity #527]: #538 +#537 := [trans #633 #637]: #522 +#488 := [quant-inst]: #529 +#539 := [mp #488 #537]: #528 +#615 := [unit-resolution #539 #724 #614]: #643 +#611 := (not #643) +#616 := (or #611 #533) +#602 := [th-lemma]: #616 +#603 := [unit-resolution #602 #615]: #533 +#606 := (not #544) +#605 := (not #533) +#607 := (or #367 #605 #606) +#604 := [th-lemma]: #607 +#608 := [unit-resolution #604 #603 #610]: #367 +#701 := (not #367) +#358 := (or #701 #485) +#58 := (= #10 #13) +#83 := (or #58 #77) +#713 := (forall (vars (?x2 int)) (:pat #712) #83) +#88 := (forall (vars (?x2 int)) #83) +#716 := (iff #88 #713) +#714 := (iff #83 #83) +#715 := [refl]: #714 +#717 := [quant-intro #715]: #716 +#191 := (~ #88 #88) +#195 := (~ #83 #83) +#193 := [refl]: #195 +#225 := [nnf-pos #193]: #191 +#14 := (= #13 #10) +#11 := (<= 0::int #10) +#15 := (implies #11 #14) +#16 := (forall (vars (?x2 int)) #15) +#91 := (iff #16 #88) +#65 := (not #11) +#66 := (or #65 #58) +#71 := (forall (vars (?x2 int)) #66) +#89 := (iff #71 #88) +#86 := (iff #66 #83) +#80 := (or #77 #58) +#84 := (iff #80 #83) +#85 := [rewrite]: #84 +#81 := (iff #66 #80) +#78 := (iff #65 #77) +#74 := (iff #11 #76) +#75 := [rewrite]: #74 +#79 := [monotonicity #75]: #78 +#82 := [monotonicity #79]: #81 +#87 := [trans #82 #85]: #86 +#90 := [quant-intro #87]: #89 +#72 := (iff #16 #71) +#69 := (iff #15 #66) +#62 := (implies #11 #58) +#67 := (iff #62 #66) +#68 := [rewrite]: #67 +#63 := (iff #15 #62) +#60 := (iff #14 #58) +#61 := [rewrite]: #60 +#64 := [monotonicity #61]: #63 +#70 := [trans #64 #68]: #69 +#73 := [quant-intro #70]: #72 +#92 := [trans #73 #90]: #91 +#57 := [asserted]: #16 +#93 := [mp #57 #92]: #88 +#226 := [mp~ #93 #225]: #88 +#718 := [mp #226 #717]: #713 +#679 := (not #713) +#342 := (or #679 #701 #485) +#380 := (>= #24 0::int) +#381 := (not #380) +#361 := (= #24 #26) +#696 := (or #361 #381) +#343 := (or #679 #696) +#685 := (iff #343 #342) +#345 := (or #679 #358) +#683 := (iff #345 #342) +#684 := [rewrite]: #683 +#681 := (iff #343 #345) +#695 := (iff #696 #358) +#703 := (or #485 #701) +#694 := (iff #703 #358) +#354 := [rewrite]: #694 +#693 := (iff #696 #703) +#702 := (iff #381 #701) +#699 := (iff #380 #367) +#700 := [rewrite]: #699 +#697 := [monotonicity #700]: #702 +#692 := (iff #361 #485) +#366 := [rewrite]: #692 +#353 := [monotonicity #366 #697]: #693 +#338 := [trans #353 #354]: #695 +#682 := [monotonicity #338]: #681 +#680 := [trans #682 #684]: #685 +#344 := [quant-inst]: #343 +#686 := [mp #344 #680]: #342 +#588 := [unit-resolution #686 #718]: #358 +#589 := [unit-resolution #588 #608]: #485 +#591 := (not #485) +#592 := (or #591 #324) +#593 := [th-lemma]: #592 +#594 := [unit-resolution #593 #589]: #324 +[th-lemma #603 #188 #594 #610]: false +unsat diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_nat_arith_05 --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_nat_arith_05 Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,13 @@ +(benchmark Isabelle +:extrasorts ( T1) +:extrafuns ( + (uf_1 Int T1) + (uf_2 T1 Int) + (uf_3 T1) + ) +:assumption (forall (?x1 T1) (= (uf_1 (uf_2 ?x1)) ?x1)) +:assumption (forall (?x2 Int) (implies (<= 0 ?x2) (= (uf_2 (uf_1 ?x2)) ?x2))) +:assumption (forall (?x3 Int) (implies (< ?x3 0) (= (uf_2 (uf_1 ?x3)) 0))) +:assumption (not (distinct (uf_1 (+ (uf_2 uf_3) 1)) (uf_1 (+ (uf_2 (uf_1 (* (uf_2 uf_3) 2))) 3)) (uf_1 (- (uf_2 uf_3) (uf_2 uf_3))))) +:formula true +) diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_nat_arith_05.proof --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_nat_arith_05.proof Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,525 @@ +#2 := false +decl uf_2 :: (-> T1 int) +decl uf_1 :: (-> int T1) +decl uf_3 :: T1 +#21 := uf_3 +#22 := (uf_2 uf_3) +#23 := 1::int +#138 := (+ 1::int #22) +#141 := (uf_1 #138) +#656 := (uf_2 #141) +#26 := 2::int +#144 := (* 2::int #22) +#147 := (uf_1 #144) +#150 := (uf_2 #147) +#30 := 3::int +#156 := (+ 3::int #150) +#161 := (uf_1 #156) +#494 := (uf_2 #161) +#288 := (= #494 #656) +#266 := (= #161 #141) +#249 := (= #141 #161) +#9 := 0::int +#166 := (uf_1 0::int) +#251 := (= #161 #166) +#337 := (not #251) +#567 := (= #494 0::int) +#543 := (uf_2 #166) +#547 := (= #543 0::int) +#10 := (:var 0 int) +#12 := (uf_1 #10) +#673 := (pattern #12) +#78 := (>= #10 0::int) +#79 := (not #78) +#13 := (uf_2 #12) +#60 := (= #10 #13) +#85 := (or #60 #79) +#674 := (forall (vars (?x2 int)) (:pat #673) #85) +#90 := (forall (vars (?x2 int)) #85) +#677 := (iff #90 #674) +#675 := (iff #85 #85) +#676 := [refl]: #675 +#678 := [quant-intro #676]: #677 +#178 := (~ #90 #90) +#190 := (~ #85 #85) +#191 := [refl]: #190 +#175 := [nnf-pos #191]: #178 +#14 := (= #13 #10) +#11 := (<= 0::int #10) +#15 := (implies #11 #14) +#16 := (forall (vars (?x2 int)) #15) +#93 := (iff #16 #90) +#67 := (not #11) +#68 := (or #67 #60) +#73 := (forall (vars (?x2 int)) #68) +#91 := (iff #73 #90) +#88 := (iff #68 #85) +#82 := (or #79 #60) +#86 := (iff #82 #85) +#87 := [rewrite]: #86 +#83 := (iff #68 #82) +#80 := (iff #67 #79) +#76 := (iff #11 #78) +#77 := [rewrite]: #76 +#81 := [monotonicity #77]: #80 +#84 := [monotonicity #81]: #83 +#89 := [trans #84 #87]: #88 +#92 := [quant-intro #89]: #91 +#74 := (iff #16 #73) +#71 := (iff #15 #68) +#64 := (implies #11 #60) +#69 := (iff #64 #68) +#70 := [rewrite]: #69 +#65 := (iff #15 #64) +#62 := (iff #14 #60) +#63 := [rewrite]: #62 +#66 := [monotonicity #63]: #65 +#72 := [trans #66 #70]: #71 +#75 := [quant-intro #72]: #74 +#94 := [trans #75 #92]: #93 +#59 := [asserted]: #16 +#95 := [mp #59 #94]: #90 +#192 := [mp~ #95 #175]: #90 +#679 := [mp #192 #678]: #674 +#290 := (not #674) +#519 := (or #290 #547) +#540 := (>= 0::int 0::int) +#541 := (not #540) +#544 := (= 0::int #543) +#545 := (or #544 #541) +#520 := (or #290 #545) +#521 := (iff #520 #519) +#523 := (iff #519 #519) +#526 := [rewrite]: #523 +#407 := (iff #545 #547) +#533 := (or #547 false) +#513 := (iff #533 #547) +#514 := [rewrite]: #513 +#539 := (iff #545 #533) +#537 := (iff #541 false) +#1 := true +#530 := (not true) +#535 := (iff #530 false) +#536 := [rewrite]: #535 +#531 := (iff #541 #530) +#548 := (iff #540 true) +#534 := [rewrite]: #548 +#532 := [monotonicity #534]: #531 +#538 := [trans #532 #536]: #537 +#546 := (iff #544 #547) +#542 := [rewrite]: #546 +#512 := [monotonicity #542 #538]: #539 +#518 := [trans #512 #514]: #407 +#522 := [monotonicity #518]: #521 +#527 := [trans #522 #526]: #521 +#525 := [quant-inst]: #520 +#528 := [mp #525 #527]: #519 +#316 := [unit-resolution #528 #679]: #547 +#286 := (= #494 #543) +#287 := [hypothesis]: #251 +#292 := [monotonicity #287]: #286 +#267 := [trans #292 #316]: #567 +#296 := (not #567) +#551 := (<= #494 0::int) +#300 := (not #551) +#501 := (>= #150 0::int) +#622 := (>= #144 0::int) +#302 := -1::int +#303 := (* -1::int #656) +#304 := (+ #22 #303) +#635 := (>= #304 -1::int) +#305 := (= #304 -1::int) +#644 := (>= #22 -1::int) +#511 := (>= #22 0::int) +#487 := (= #22 0::int) +#660 := (uf_1 #22) +#517 := (uf_2 #660) +#485 := (= #517 0::int) +#389 := (not #511) +#390 := [hypothesis]: #389 +#492 := (or #485 #511) +#18 := (= #13 0::int) +#126 := (or #18 #78) +#680 := (forall (vars (?x3 int)) (:pat #673) #126) +#131 := (forall (vars (?x3 int)) #126) +#683 := (iff #131 #680) +#681 := (iff #126 #126) +#682 := [refl]: #681 +#684 := [quant-intro #682]: #683 +#179 := (~ #131 #131) +#193 := (~ #126 #126) +#194 := [refl]: #193 +#180 := [nnf-pos #194]: #179 +#17 := (< #10 0::int) +#19 := (implies #17 #18) +#20 := (forall (vars (?x3 int)) #19) +#134 := (iff #20 #131) +#97 := (= 0::int #13) +#103 := (not #17) +#104 := (or #103 #97) +#109 := (forall (vars (?x3 int)) #104) +#132 := (iff #109 #131) +#129 := (iff #104 #126) +#123 := (or #78 #18) +#127 := (iff #123 #126) +#128 := [rewrite]: #127 +#124 := (iff #104 #123) +#121 := (iff #97 #18) +#122 := [rewrite]: #121 +#119 := (iff #103 #78) +#114 := (not #79) +#117 := (iff #114 #78) +#118 := [rewrite]: #117 +#115 := (iff #103 #114) +#112 := (iff #17 #79) +#113 := [rewrite]: #112 +#116 := [monotonicity #113]: #115 +#120 := [trans #116 #118]: #119 +#125 := [monotonicity #120 #122]: #124 +#130 := [trans #125 #128]: #129 +#133 := [quant-intro #130]: #132 +#110 := (iff #20 #109) +#107 := (iff #19 #104) +#100 := (implies #17 #97) +#105 := (iff #100 #104) +#106 := [rewrite]: #105 +#101 := (iff #19 #100) +#98 := (iff #18 #97) +#99 := [rewrite]: #98 +#102 := [monotonicity #99]: #101 +#108 := [trans #102 #106]: #107 +#111 := [quant-intro #108]: #110 +#135 := [trans #111 #133]: #134 +#96 := [asserted]: #20 +#136 := [mp #96 #135]: #131 +#195 := [mp~ #136 #180]: #131 +#685 := [mp #195 #684]: #680 +#637 := (not #680) +#484 := (or #637 #485 #511) +#486 := (or #637 #492) +#495 := (iff #486 #484) +#496 := [rewrite]: #495 +#493 := [quant-inst]: #486 +#497 := [mp #493 #496]: #484 +#391 := [unit-resolution #497 #685]: #492 +#392 := [unit-resolution #391 #390]: #485 +#394 := (= #22 #517) +#661 := (= uf_3 #660) +#4 := (:var 0 T1) +#5 := (uf_2 #4) +#665 := (pattern #5) +#6 := (uf_1 #5) +#53 := (= #4 #6) +#666 := (forall (vars (?x1 T1)) (:pat #665) #53) +#56 := (forall (vars (?x1 T1)) #53) +#667 := (iff #56 #666) +#669 := (iff #666 #666) +#670 := [rewrite]: #669 +#668 := [rewrite]: #667 +#671 := [trans #668 #670]: #667 +#187 := (~ #56 #56) +#185 := (~ #53 #53) +#186 := [refl]: #185 +#188 := [nnf-pos #186]: #187 +#7 := (= #6 #4) +#8 := (forall (vars (?x1 T1)) #7) +#57 := (iff #8 #56) +#54 := (iff #7 #53) +#55 := [rewrite]: #54 +#58 := [quant-intro #55]: #57 +#52 := [asserted]: #8 +#61 := [mp #52 #58]: #56 +#189 := [mp~ #61 #188]: #56 +#672 := [mp #189 #671]: #666 +#658 := (not #666) +#664 := (or #658 #661) +#654 := [quant-inst]: #664 +#393 := [unit-resolution #654 #672]: #661 +#395 := [monotonicity #393]: #394 +#396 := [trans #395 #392]: #487 +#397 := (not #487) +#398 := (or #397 #511) +#399 := [th-lemma]: #398 +#388 := [unit-resolution #399 #390 #396]: false +#400 := [lemma #388]: #511 +#366 := (or #389 #644) +#367 := [th-lemma]: #366 +#352 := [unit-resolution #367 #400]: #644 +#641 := (not #644) +#648 := (or #305 #641) +#651 := (or #290 #305 #641) +#313 := (>= #138 0::int) +#318 := (not #313) +#298 := (= #138 #656) +#640 := (or #298 #318) +#649 := (or #290 #640) +#363 := (iff #649 #651) +#638 := (or #290 #648) +#361 := (iff #638 #651) +#362 := [rewrite]: #361 +#639 := (iff #649 #638) +#650 := (iff #640 #648) +#647 := (iff #318 #641) +#645 := (iff #313 #644) +#646 := [rewrite]: #645 +#284 := [monotonicity #646]: #647 +#642 := (iff #298 #305) +#643 := [rewrite]: #642 +#289 := [monotonicity #643 #284]: #650 +#346 := [monotonicity #289]: #639 +#364 := [trans #346 #362]: #363 +#652 := [quant-inst]: #649 +#257 := [mp #652 #364]: #651 +#424 := [unit-resolution #257 #679]: #648 +#353 := [unit-resolution #424 #352]: #305 +#439 := (not #305) +#281 := (or #439 #635) +#440 := [th-lemma]: #281 +#330 := [unit-resolution #440 #353]: #635 +#620 := (<= #656 0::int) +#441 := (not #620) +#634 := (<= #304 -1::int) +#344 := (or #439 #634) +#354 := [th-lemma]: #344 +#355 := [unit-resolution #354 #353]: #634 +#345 := (not #634) +#356 := (or #441 #389 #345) +#322 := [th-lemma]: #356 +#324 := [unit-resolution #322 #355 #400]: #441 +#432 := (not #635) +#331 := (or #622 #432 #620) +#319 := [th-lemma]: #331 +#320 := [unit-resolution #319 #324 #330]: #622 +#624 := (* -1::int #150) +#619 := (+ #144 #624) +#606 := (<= #619 0::int) +#625 := (= #619 0::int) +#617 := (not #622) +#612 := (or #617 #625) +#615 := (or #290 #617 #625) +#618 := (= #144 #150) +#623 := (or #618 #617) +#609 := (or #290 #623) +#604 := (iff #609 #615) +#445 := (or #290 #612) +#601 := (iff #445 #615) +#602 := [rewrite]: #601 +#447 := (iff #609 #445) +#608 := (iff #623 #612) +#468 := (or #625 #617) +#613 := (iff #468 #612) +#607 := [rewrite]: #613 +#610 := (iff #623 #468) +#466 := (iff #618 #625) +#467 := [rewrite]: #466 +#611 := [monotonicity #467]: #610 +#614 := [trans #611 #607]: #608 +#448 := [monotonicity #614]: #447 +#605 := [trans #448 #602]: #604 +#616 := [quant-inst]: #609 +#603 := [mp #616 #605]: #615 +#480 := [unit-resolution #603 #679]: #612 +#299 := [unit-resolution #480 #320]: #625 +#406 := (not #625) +#408 := (or #406 #606) +#409 := [th-lemma]: #408 +#301 := [unit-resolution #409 #299]: #606 +#413 := (not #606) +#306 := (or #501 #413 #617) +#307 := [th-lemma]: #306 +#308 := [unit-resolution #307 #301 #320]: #501 +#506 := -3::int +#504 := (* -1::int #494) +#505 := (+ #150 #504) +#564 := (<= #505 -3::int) +#599 := (= #505 -3::int) +#587 := (>= #150 -3::int) +#417 := (or #587 #413 #617) +#410 := [th-lemma]: #417 +#309 := [unit-resolution #410 #301 #320]: #587 +#578 := (not #587) +#593 := (or #578 #599) +#579 := (or #290 #578 #599) +#449 := (>= #156 0::int) +#597 := (not #449) +#502 := (= #156 #494) +#503 := (or #502 #597) +#586 := (or #290 #503) +#572 := (iff #586 #579) +#571 := (or #290 #593) +#575 := (iff #571 #579) +#576 := [rewrite]: #575 +#573 := (iff #586 #571) +#584 := (iff #503 #593) +#591 := (or #599 #578) +#582 := (iff #591 #593) +#583 := [rewrite]: #582 +#592 := (iff #503 #591) +#580 := (iff #597 #578) +#589 := (iff #449 #587) +#581 := [rewrite]: #589 +#590 := [monotonicity #581]: #580 +#596 := (iff #502 #599) +#600 := [rewrite]: #596 +#588 := [monotonicity #600 #590]: #592 +#585 := [trans #588 #583]: #584 +#574 := [monotonicity #585]: #573 +#577 := [trans #574 #576]: #572 +#570 := [quant-inst]: #586 +#563 := [mp #570 #577]: #579 +#458 := [unit-resolution #563 #679]: #593 +#310 := [unit-resolution #458 #309]: #599 +#460 := (not #599) +#461 := (or #460 #564) +#444 := [th-lemma]: #461 +#311 := [unit-resolution #444 #310]: #564 +#434 := (not #564) +#453 := (not #501) +#312 := (or #300 #453 #434) +#293 := [th-lemma]: #312 +#295 := [unit-resolution #293 #311 #308]: #300 +#294 := (or #296 #551) +#297 := [th-lemma]: #294 +#285 := [unit-resolution #297 #295]: #296 +#271 := [unit-resolution #285 #267]: false +#272 := [lemma #271]: #337 +#282 := (or #249 #251) +#250 := (= #141 #166) +#336 := (not #250) +#357 := (= #656 0::int) +#332 := (= #656 #543) +#329 := [hypothesis]: #250 +#333 := [monotonicity #329]: #332 +#323 := [trans #333 #316]: #357 +#429 := (not #357) +#430 := (or #429 #620) +#428 := [th-lemma]: #430 +#325 := [unit-resolution #428 #324]: #429 +#334 := [unit-resolution #325 #323]: false +#317 := [lemma #334]: #336 +#279 := (or #249 #250 #251) +#335 := (not #249) +#328 := (and #335 #336 #337) +#339 := (not #328) +#169 := (distinct #141 #161 #166) +#172 := (not #169) +#33 := (- #22 #22) +#34 := (uf_1 #33) +#27 := (* #22 2::int) +#28 := (uf_1 #27) +#29 := (uf_2 #28) +#31 := (+ #29 3::int) +#32 := (uf_1 #31) +#24 := (+ #22 1::int) +#25 := (uf_1 #24) +#35 := (distinct #25 #32 #34) +#36 := (not #35) +#173 := (iff #36 #172) +#170 := (iff #35 #169) +#167 := (= #34 #166) +#164 := (= #33 0::int) +#165 := [rewrite]: #164 +#168 := [monotonicity #165]: #167 +#162 := (= #32 #161) +#159 := (= #31 #156) +#153 := (+ #150 3::int) +#157 := (= #153 #156) +#158 := [rewrite]: #157 +#154 := (= #31 #153) +#151 := (= #29 #150) +#148 := (= #28 #147) +#145 := (= #27 #144) +#146 := [rewrite]: #145 +#149 := [monotonicity #146]: #148 +#152 := [monotonicity #149]: #151 +#155 := [monotonicity #152]: #154 +#160 := [trans #155 #158]: #159 +#163 := [monotonicity #160]: #162 +#142 := (= #25 #141) +#139 := (= #24 #138) +#140 := [rewrite]: #139 +#143 := [monotonicity #140]: #142 +#171 := [monotonicity #143 #163 #168]: #170 +#174 := [monotonicity #171]: #173 +#137 := [asserted]: #36 +#177 := [mp #137 #174]: #172 +#326 := (or #169 #339) +#327 := [def-axiom]: #326 +#277 := [unit-resolution #327 #177]: #339 +#659 := (or #328 #249 #250 #251) +#315 := [def-axiom]: #659 +#280 := [unit-resolution #315 #277]: #279 +#278 := [unit-resolution #280 #317]: #282 +#283 := [unit-resolution #278 #272]: #249 +#269 := [symm #283]: #266 +#270 := [monotonicity #269]: #288 +#508 := (+ #494 #303) +#473 := (<= #508 0::int) +#433 := (not #473) +#477 := [hypothesis]: #473 +#421 := (or #622 #433) +#489 := (= #150 0::int) +#478 := [hypothesis]: #617 +#490 := (or #489 #622) +#499 := (or #637 #489 #622) +#594 := (or #637 #490) +#598 := (iff #594 #499) +#483 := [rewrite]: #598 +#595 := [quant-inst]: #594 +#498 := [mp #595 #483]: #499 +#465 := [unit-resolution #498 #685]: #490 +#481 := [unit-resolution #465 #478]: #489 +#442 := (not #489) +#443 := (or #442 #501) +#450 := [th-lemma]: #443 +#452 := [unit-resolution #450 #481]: #501 +#454 := (or #453 #587) +#456 := [th-lemma]: #454 +#457 := [unit-resolution #456 #452]: #587 +#459 := [unit-resolution #458 #457]: #599 +#462 := [unit-resolution #444 #459]: #564 +#435 := (or #432 #622 #433 #453 #434) +#437 := [th-lemma]: #435 +#438 := [unit-resolution #437 #478 #452 #462 #477]: #432 +#436 := [unit-resolution #440 #438]: #439 +#420 := (or #441 #433 #453 #434) +#423 := [th-lemma]: #420 +#427 := [unit-resolution #423 #452 #462 #477]: #441 +#431 := [unit-resolution #428 #427]: #429 +#632 := (or #357 #644) +#347 := (or #637 #357 #644) +#358 := (or #357 #313) +#348 := (or #637 #358) +#630 := (iff #348 #347) +#350 := (or #637 #632) +#343 := (iff #350 #347) +#626 := [rewrite]: #343 +#628 := (iff #348 #350) +#636 := (iff #358 #632) +#633 := [monotonicity #646]: #636 +#629 := [monotonicity #633]: #628 +#627 := [trans #629 #626]: #630 +#349 := [quant-inst]: #348 +#631 := [mp #349 #627]: #347 +#419 := [unit-resolution #631 #685]: #632 +#422 := [unit-resolution #419 #431]: #644 +#425 := [unit-resolution #424 #422 #436]: false +#426 := [lemma #425]: #421 +#479 := [unit-resolution #426 #477]: #622 +#416 := [unit-resolution #480 #479]: #625 +#412 := [unit-resolution #409 #416]: #606 +#418 := [unit-resolution #410 #412 #479]: #587 +#411 := [unit-resolution #458 #418]: #599 +#414 := [unit-resolution #444 #411]: #564 +#415 := (or #644 #617) +#401 := [th-lemma]: #415 +#403 := [unit-resolution #401 #479]: #644 +#404 := [unit-resolution #424 #403]: #305 +#402 := [unit-resolution #440 #404]: #635 +#405 := [th-lemma #418 #402 #477 #414 #412]: false +#387 := [lemma #405]: #433 +#273 := (not #288) +#274 := (or #273 #473) +#275 := [th-lemma]: #274 +[unit-resolution #275 #387 #270]: false +unsat diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_nat_arith_06 --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_nat_arith_06 Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,13 @@ +(benchmark Isabelle +:extrasorts ( T1) +:extrafuns ( + (uf_1 Int T1) + (uf_2 T1 Int) + (uf_3 Int) + ) +:assumption (forall (?x1 T1) (= (uf_1 (uf_2 ?x1)) ?x1)) +:assumption (forall (?x2 Int) (implies (<= 0 ?x2) (= (uf_2 (uf_1 ?x2)) ?x2))) +:assumption (forall (?x3 Int) (implies (< ?x3 0) (= (uf_2 (uf_1 ?x3)) 0))) +:assumption (not (= (uf_2 (uf_1 (ite (< uf_3 0) (~ uf_3) uf_3))) (ite (< uf_3 0) (~ uf_3) uf_3))) +:formula true +) diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_nat_arith_06.proof --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_nat_arith_06.proof Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,161 @@ +#2 := false +#9 := 0::int +decl uf_3 :: int +#21 := uf_3 +#130 := -1::int +#131 := (* -1::int uf_3) +#154 := (>= uf_3 0::int) +#161 := (ite #154 uf_3 #131) +#648 := (* -1::int #161) +#651 := (+ #131 #648) +#657 := (<= #651 0::int) +#341 := (= #131 #161) +#155 := (not #154) +#649 := (+ uf_3 #648) +#650 := (<= #649 0::int) +#254 := (= uf_3 #161) +#646 := [hypothesis]: #154 +#255 := (or #155 #254) +#342 := [def-axiom]: #255 +#652 := [unit-resolution #342 #646]: #254 +#290 := (not #254) +#653 := (or #290 #650) +#655 := [th-lemma]: #653 +#295 := [unit-resolution #655 #652]: #650 +#346 := (>= #161 0::int) +#274 := (not #346) +decl uf_2 :: (-> T1 int) +decl uf_1 :: (-> int T1) +#166 := (uf_1 #161) +#169 := (uf_2 #166) +#172 := (= #161 #169) +#175 := (not #172) +#23 := (- uf_3) +#22 := (< uf_3 0::int) +#24 := (ite #22 #23 uf_3) +#25 := (uf_1 #24) +#26 := (uf_2 #25) +#27 := (= #26 #24) +#28 := (not #27) +#178 := (iff #28 #175) +#134 := (ite #22 #131 uf_3) +#137 := (uf_1 #134) +#140 := (uf_2 #137) +#146 := (= #134 #140) +#151 := (not #146) +#176 := (iff #151 #175) +#173 := (iff #146 #172) +#170 := (= #140 #169) +#167 := (= #137 #166) +#164 := (= #134 #161) +#158 := (ite #155 #131 uf_3) +#162 := (= #158 #161) +#163 := [rewrite]: #162 +#159 := (= #134 #158) +#156 := (iff #22 #155) +#157 := [rewrite]: #156 +#160 := [monotonicity #157]: #159 +#165 := [trans #160 #163]: #164 +#168 := [monotonicity #165]: #167 +#171 := [monotonicity #168]: #170 +#174 := [monotonicity #165 #171]: #173 +#177 := [monotonicity #174]: #176 +#152 := (iff #28 #151) +#149 := (iff #27 #146) +#143 := (= #140 #134) +#147 := (iff #143 #146) +#148 := [rewrite]: #147 +#144 := (iff #27 #143) +#135 := (= #24 #134) +#132 := (= #23 #131) +#133 := [rewrite]: #132 +#136 := [monotonicity #133]: #135 +#141 := (= #26 #140) +#138 := (= #25 #137) +#139 := [monotonicity #136]: #138 +#142 := [monotonicity #139]: #141 +#145 := [monotonicity #142 #136]: #144 +#150 := [trans #145 #148]: #149 +#153 := [monotonicity #150]: #152 +#179 := [trans #153 #177]: #178 +#129 := [asserted]: #28 +#180 := [mp #129 #179]: #175 +#10 := (:var 0 int) +#12 := (uf_1 #10) +#678 := (pattern #12) +#70 := (>= #10 0::int) +#71 := (not #70) +#13 := (uf_2 #12) +#52 := (= #10 #13) +#77 := (or #52 #71) +#679 := (forall (vars (?x2 int)) (:pat #678) #77) +#82 := (forall (vars (?x2 int)) #77) +#682 := (iff #82 #679) +#680 := (iff #77 #77) +#681 := [refl]: #680 +#683 := [quant-intro #681]: #682 +#183 := (~ #82 #82) +#195 := (~ #77 #77) +#196 := [refl]: #195 +#181 := [nnf-pos #196]: #183 +#14 := (= #13 #10) +#11 := (<= 0::int #10) +#15 := (implies #11 #14) +#16 := (forall (vars (?x2 int)) #15) +#85 := (iff #16 #82) +#59 := (not #11) +#60 := (or #59 #52) +#65 := (forall (vars (?x2 int)) #60) +#83 := (iff #65 #82) +#80 := (iff #60 #77) +#74 := (or #71 #52) +#78 := (iff #74 #77) +#79 := [rewrite]: #78 +#75 := (iff #60 #74) +#72 := (iff #59 #71) +#68 := (iff #11 #70) +#69 := [rewrite]: #68 +#73 := [monotonicity #69]: #72 +#76 := [monotonicity #73]: #75 +#81 := [trans #76 #79]: #80 +#84 := [quant-intro #81]: #83 +#66 := (iff #16 #65) +#63 := (iff #15 #60) +#56 := (implies #11 #52) +#61 := (iff #56 #60) +#62 := [rewrite]: #61 +#57 := (iff #15 #56) +#54 := (iff #14 #52) +#55 := [rewrite]: #54 +#58 := [monotonicity #55]: #57 +#64 := [trans #58 #62]: #63 +#67 := [quant-intro #64]: #66 +#86 := [trans #67 #84]: #85 +#51 := [asserted]: #16 +#87 := [mp #51 #86]: #82 +#197 := [mp~ #87 #181]: #82 +#684 := [mp #197 #683]: #679 +#321 := (not #679) +#451 := (or #321 #172 #274) +#327 := (or #172 #274) +#658 := (or #321 #327) +#333 := (iff #658 #451) +#665 := [rewrite]: #333 +#332 := [quant-inst]: #658 +#666 := [mp #332 #665]: #451 +#296 := [unit-resolution #666 #684 #180]: #274 +#656 := [th-lemma #646 #296 #295]: false +#654 := [lemma #656]: #155 +#256 := (or #154 #341) +#343 := [def-axiom]: #256 +#644 := [unit-resolution #343 #654]: #341 +#366 := (not #341) +#367 := (or #366 #657) +#368 := [th-lemma]: #367 +#369 := [unit-resolution #368 #644]: #657 +#647 := (<= #161 0::int) +#262 := (or #647 #346) +#639 := [th-lemma]: #262 +#640 := [unit-resolution #639 #296]: #647 +[th-lemma #654 #640 #369]: false +unsat diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_nat_arith_07 --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_nat_arith_07 Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,19 @@ +(benchmark Isabelle +:extrasorts ( T1 T2) +:extrafuns ( + (uf_1 Int T1) + (uf_2 T1 Int) + (uf_5 T1) + ) +:extrapreds ( + (up_3 T1) + (up_4 T1 T1) + ) +:assumption (forall (?x1 T1) (= (uf_1 (uf_2 ?x1)) ?x1)) +:assumption (forall (?x2 Int) (implies (<= 0 ?x2) (= (uf_2 (uf_1 ?x2)) ?x2))) +:assumption (forall (?x3 Int) (implies (< ?x3 0) (= (uf_2 (uf_1 ?x3)) 0))) +:assumption (forall (?x4 T1) (iff (up_3 ?x4) (and (< 1 (uf_2 ?x4)) (forall (?x5 T1) (implies (up_4 ?x5 ?x4) (or (= ?x5 (uf_1 1)) (= ?x5 ?x4))))))) +:assumption (up_3 (uf_1 (+ (uf_2 (uf_1 (* 4 (uf_2 uf_5)))) 1))) +:assumption (not (<= 1 (uf_2 uf_5))) +:formula true +) diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_nat_arith_07.proof --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_nat_arith_07.proof Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,553 @@ +#2 := false +#9 := 0::int +decl uf_2 :: (-> T1 int) +decl uf_1 :: (-> int T1) +decl uf_5 :: T1 +#36 := uf_5 +#37 := (uf_2 uf_5) +#35 := 4::int +#38 := (* 4::int #37) +#39 := (uf_1 #38) +#40 := (uf_2 #39) +#549 := (= #40 0::int) +#963 := (not #549) +#537 := (<= #40 0::int) +#958 := (not #537) +#22 := 1::int +#186 := (+ 1::int #40) +#189 := (uf_1 #186) +#524 := (uf_2 #189) +#452 := (<= #524 1::int) +#874 := (not #452) +decl up_4 :: (-> T1 T1 bool) +#4 := (:var 0 T1) +#456 := (up_4 #4 #189) +#440 := (pattern #456) +#446 := (not #456) +#455 := (= #4 #189) +#26 := (uf_1 1::int) +#27 := (= #4 #26) +#434 := (or #27 #455 #446) +#416 := (forall (vars (?x5 T1)) (:pat #440) #434) +#417 := (not #416) +#409 := (or #417 #452) +#400 := (not #409) +decl up_3 :: (-> T1 bool) +#192 := (up_3 #189) +#429 := (not #192) +#405 := (or #429 #400) +#389 := (not #405) +decl ?x5!0 :: (-> T1 T1) +#478 := (?x5!0 #189) +#479 := (= #26 #478) +#468 := (= #189 #478) +#445 := (up_4 #478 #189) +#447 := (not #445) +#396 := (or #447 #468 #479) +#391 := (not #396) +#386 := (or #192 #391 #452) +#377 := (not #386) +#843 := (or #377 #389) +#848 := (not #843) +#5 := (uf_2 #4) +#788 := (pattern #5) +#21 := (up_3 #4) +#836 := (pattern #21) +#210 := (?x5!0 #4) +#274 := (= #4 #210) +#271 := (= #26 #210) +#232 := (up_4 #210 #4) +#233 := (not #232) +#277 := (or #233 #271 #274) +#280 := (not #277) +#163 := (<= #5 1::int) +#289 := (or #21 #163 #280) +#304 := (not #289) +#24 := (:var 1 T1) +#25 := (up_4 #4 #24) +#809 := (pattern #25) +#28 := (= #4 #24) +#147 := (not #25) +#167 := (or #147 #27 #28) +#810 := (forall (vars (?x5 T1)) (:pat #809) #167) +#815 := (not #810) +#818 := (or #163 #815) +#821 := (not #818) +#253 := (not #21) +#824 := (or #253 #821) +#827 := (not #824) +#830 := (or #827 #304) +#833 := (not #830) +#837 := (forall (vars (?x4 T1)) (:pat #836 #788) #833) +#170 := (forall (vars (?x5 T1)) #167) +#236 := (not #170) +#239 := (or #163 #236) +#240 := (not #239) +#215 := (or #253 #240) +#303 := (not #215) +#305 := (or #303 #304) +#306 := (not #305) +#311 := (forall (vars (?x4 T1)) #306) +#838 := (iff #311 #837) +#834 := (iff #306 #833) +#831 := (iff #305 #830) +#828 := (iff #303 #827) +#825 := (iff #215 #824) +#822 := (iff #240 #821) +#819 := (iff #239 #818) +#816 := (iff #236 #815) +#813 := (iff #170 #810) +#811 := (iff #167 #167) +#812 := [refl]: #811 +#814 := [quant-intro #812]: #813 +#817 := [monotonicity #814]: #816 +#820 := [monotonicity #817]: #819 +#823 := [monotonicity #820]: #822 +#826 := [monotonicity #823]: #825 +#829 := [monotonicity #826]: #828 +#832 := [monotonicity #829]: #831 +#835 := [monotonicity #832]: #834 +#839 := [quant-intro #835]: #838 +#164 := (not #163) +#173 := (and #164 #170) +#259 := (or #253 #173) +#294 := (and #259 #289) +#297 := (forall (vars (?x4 T1)) #294) +#312 := (iff #297 #311) +#309 := (iff #294 #306) +#214 := (and #215 #289) +#307 := (iff #214 #306) +#308 := [rewrite]: #307 +#301 := (iff #294 #214) +#216 := (iff #259 #215) +#268 := (iff #173 #240) +#300 := [rewrite]: #268 +#213 := [monotonicity #300]: #216 +#302 := [monotonicity #213]: #301 +#310 := [trans #302 #308]: #309 +#313 := [quant-intro #310]: #312 +#230 := (= #210 #4) +#231 := (= #210 #26) +#234 := (or #233 #231 #230) +#235 := (not #234) +#228 := (not #164) +#241 := (or #228 #235) +#258 := (or #21 #241) +#260 := (and #259 #258) +#263 := (forall (vars (?x4 T1)) #260) +#298 := (iff #263 #297) +#295 := (iff #260 #294) +#292 := (iff #258 #289) +#283 := (or #163 #280) +#286 := (or #21 #283) +#290 := (iff #286 #289) +#291 := [rewrite]: #290 +#287 := (iff #258 #286) +#284 := (iff #241 #283) +#281 := (iff #235 #280) +#278 := (iff #234 #277) +#275 := (iff #230 #274) +#276 := [rewrite]: #275 +#272 := (iff #231 #271) +#273 := [rewrite]: #272 +#279 := [monotonicity #273 #276]: #278 +#282 := [monotonicity #279]: #281 +#269 := (iff #228 #163) +#270 := [rewrite]: #269 +#285 := [monotonicity #270 #282]: #284 +#288 := [monotonicity #285]: #287 +#293 := [trans #288 #291]: #292 +#296 := [monotonicity #293]: #295 +#299 := [quant-intro #296]: #298 +#176 := (iff #21 #173) +#179 := (forall (vars (?x4 T1)) #176) +#264 := (~ #179 #263) +#261 := (~ #176 #260) +#251 := (~ #173 #173) +#249 := (~ #170 #170) +#247 := (~ #167 #167) +#248 := [refl]: #247 +#250 := [nnf-pos #248]: #249 +#245 := (~ #164 #164) +#246 := [refl]: #245 +#252 := [monotonicity #246 #250]: #251 +#242 := (not #173) +#243 := (~ #242 #241) +#237 := (~ #236 #235) +#238 := [sk]: #237 +#229 := (~ #228 #228) +#209 := [refl]: #229 +#244 := [nnf-neg #209 #238]: #243 +#256 := (~ #21 #21) +#257 := [refl]: #256 +#254 := (~ #253 #253) +#255 := [refl]: #254 +#262 := [nnf-pos #255 #257 #244 #252]: #261 +#265 := [nnf-pos #262]: #264 +#29 := (or #27 #28) +#30 := (implies #25 #29) +#31 := (forall (vars (?x5 T1)) #30) +#23 := (< 1::int #5) +#32 := (and #23 #31) +#33 := (iff #21 #32) +#34 := (forall (vars (?x4 T1)) #33) +#182 := (iff #34 #179) +#148 := (or #147 #29) +#151 := (forall (vars (?x5 T1)) #148) +#154 := (and #23 #151) +#157 := (iff #21 #154) +#160 := (forall (vars (?x4 T1)) #157) +#180 := (iff #160 #179) +#177 := (iff #157 #176) +#174 := (iff #154 #173) +#171 := (iff #151 #170) +#168 := (iff #148 #167) +#169 := [rewrite]: #168 +#172 := [quant-intro #169]: #171 +#165 := (iff #23 #164) +#166 := [rewrite]: #165 +#175 := [monotonicity #166 #172]: #174 +#178 := [monotonicity #175]: #177 +#181 := [quant-intro #178]: #180 +#161 := (iff #34 #160) +#158 := (iff #33 #157) +#155 := (iff #32 #154) +#152 := (iff #31 #151) +#149 := (iff #30 #148) +#150 := [rewrite]: #149 +#153 := [quant-intro #150]: #152 +#156 := [monotonicity #153]: #155 +#159 := [monotonicity #156]: #158 +#162 := [quant-intro #159]: #161 +#183 := [trans #162 #181]: #182 +#146 := [asserted]: #34 +#184 := [mp #146 #183]: #179 +#266 := [mp~ #184 #265]: #263 +#267 := [mp #266 #299]: #297 +#314 := [mp #267 #313]: #311 +#840 := [mp #314 #839]: #837 +#754 := (not #837) +#851 := (or #754 #848) +#448 := (or #447 #479 #468) +#439 := (not #448) +#453 := (or #192 #452 #439) +#454 := (not #453) +#457 := (or #446 #27 #455) +#442 := (forall (vars (?x5 T1)) (:pat #440) #457) +#443 := (not #442) +#422 := (or #452 #443) +#424 := (not #422) +#430 := (or #429 #424) +#431 := (not #430) +#432 := (or #431 #454) +#433 := (not #432) +#852 := (or #754 #433) +#854 := (iff #852 #851) +#856 := (iff #851 #851) +#857 := [rewrite]: #856 +#849 := (iff #433 #848) +#846 := (iff #432 #843) +#379 := (or #389 #377) +#844 := (iff #379 #843) +#845 := [rewrite]: #844 +#841 := (iff #432 #379) +#378 := (iff #454 #377) +#388 := (iff #453 #386) +#381 := (or #192 #452 #391) +#387 := (iff #381 #386) +#383 := [rewrite]: #387 +#382 := (iff #453 #381) +#399 := (iff #439 #391) +#397 := (iff #448 #396) +#398 := [rewrite]: #397 +#384 := [monotonicity #398]: #399 +#385 := [monotonicity #384]: #382 +#375 := [trans #385 #383]: #388 +#376 := [monotonicity #375]: #378 +#392 := (iff #431 #389) +#401 := (iff #430 #405) +#402 := (iff #424 #400) +#394 := (iff #422 #409) +#410 := (or #452 #417) +#415 := (iff #410 #409) +#390 := [rewrite]: #415 +#411 := (iff #422 #410) +#420 := (iff #443 #417) +#418 := (iff #442 #416) +#423 := (iff #457 #434) +#435 := [rewrite]: #423 +#419 := [quant-intro #435]: #418 +#408 := [monotonicity #419]: #420 +#414 := [monotonicity #408]: #411 +#395 := [trans #414 #390]: #394 +#403 := [monotonicity #395]: #402 +#406 := [monotonicity #403]: #401 +#393 := [monotonicity #406]: #392 +#842 := [monotonicity #393 #376]: #841 +#847 := [trans #842 #845]: #846 +#850 := [monotonicity #847]: #849 +#855 := [monotonicity #850]: #854 +#858 := [trans #855 #857]: #854 +#853 := [quant-inst]: #852 +#859 := [mp #853 #858]: #851 +#934 := [unit-resolution #859 #840]: #848 +#893 := (or #843 #405) +#894 := [def-axiom]: #893 +#935 := [unit-resolution #894 #934]: #405 +#938 := (or #389 #400) +#41 := (+ #40 1::int) +#42 := (uf_1 #41) +#43 := (up_3 #42) +#193 := (iff #43 #192) +#190 := (= #42 #189) +#187 := (= #41 #186) +#188 := [rewrite]: #187 +#191 := [monotonicity #188]: #190 +#194 := [monotonicity #191]: #193 +#185 := [asserted]: #43 +#197 := [mp #185 #194]: #192 +#889 := (or #389 #429 #400) +#890 := [def-axiom]: #889 +#939 := [unit-resolution #890 #197]: #938 +#940 := [unit-resolution #939 #935]: #400 +#881 := (or #409 #874) +#882 := [def-axiom]: #881 +#941 := [unit-resolution #882 #940]: #874 +#555 := -1::int +#525 := (* -1::int #524) +#528 := (+ #40 #525) +#494 := (>= #528 -1::int) +#510 := (= #528 -1::int) +#514 := (>= #40 -1::int) +#495 := (= #524 0::int) +#946 := (not #495) +#467 := (<= #524 0::int) +#942 := (not #467) +#943 := (or #942 #452) +#944 := [th-lemma]: #943 +#945 := [unit-resolution #944 #941]: #942 +#947 := (or #946 #467) +#948 := [th-lemma]: #947 +#949 := [unit-resolution #948 #945]: #946 +#498 := (or #495 #514) +#10 := (:var 0 int) +#12 := (uf_1 #10) +#796 := (pattern #12) +#87 := (>= #10 0::int) +#13 := (uf_2 #12) +#18 := (= #13 0::int) +#135 := (or #18 #87) +#803 := (forall (vars (?x3 int)) (:pat #796) #135) +#140 := (forall (vars (?x3 int)) #135) +#806 := (iff #140 #803) +#804 := (iff #135 #135) +#805 := [refl]: #804 +#807 := [quant-intro #805]: #806 +#207 := (~ #140 #140) +#225 := (~ #135 #135) +#226 := [refl]: #225 +#208 := [nnf-pos #226]: #207 +#17 := (< #10 0::int) +#19 := (implies #17 #18) +#20 := (forall (vars (?x3 int)) #19) +#143 := (iff #20 #140) +#106 := (= 0::int #13) +#112 := (not #17) +#113 := (or #112 #106) +#118 := (forall (vars (?x3 int)) #113) +#141 := (iff #118 #140) +#138 := (iff #113 #135) +#132 := (or #87 #18) +#136 := (iff #132 #135) +#137 := [rewrite]: #136 +#133 := (iff #113 #132) +#130 := (iff #106 #18) +#131 := [rewrite]: #130 +#128 := (iff #112 #87) +#88 := (not #87) +#123 := (not #88) +#126 := (iff #123 #87) +#127 := [rewrite]: #126 +#124 := (iff #112 #123) +#121 := (iff #17 #88) +#122 := [rewrite]: #121 +#125 := [monotonicity #122]: #124 +#129 := [trans #125 #127]: #128 +#134 := [monotonicity #129 #131]: #133 +#139 := [trans #134 #137]: #138 +#142 := [quant-intro #139]: #141 +#119 := (iff #20 #118) +#116 := (iff #19 #113) +#109 := (implies #17 #106) +#114 := (iff #109 #113) +#115 := [rewrite]: #114 +#110 := (iff #19 #109) +#107 := (iff #18 #106) +#108 := [rewrite]: #107 +#111 := [monotonicity #108]: #110 +#117 := [trans #111 #115]: #116 +#120 := [quant-intro #117]: #119 +#144 := [trans #120 #142]: #143 +#105 := [asserted]: #20 +#145 := [mp #105 #144]: #140 +#227 := [mp~ #145 #208]: #140 +#808 := [mp #227 #807]: #803 +#532 := (not #803) +#488 := (or #532 #495 #514) +#529 := (>= #186 0::int) +#496 := (or #495 #529) +#489 := (or #532 #496) +#474 := (iff #489 #488) +#482 := (or #532 #498) +#483 := (iff #482 #488) +#493 := [rewrite]: #483 +#491 := (iff #489 #482) +#497 := (iff #496 #498) +#515 := (iff #529 #514) +#516 := [rewrite]: #515 +#499 := [monotonicity #516]: #497 +#492 := [monotonicity #499]: #491 +#475 := [trans #492 #493]: #474 +#490 := [quant-inst]: #489 +#476 := [mp #490 #475]: #488 +#950 := [unit-resolution #476 #808]: #498 +#951 := [unit-resolution #950 #949]: #514 +#517 := (not #514) +#520 := (or #510 #517) +#69 := (= #10 #13) +#94 := (or #69 #88) +#797 := (forall (vars (?x2 int)) (:pat #796) #94) +#99 := (forall (vars (?x2 int)) #94) +#800 := (iff #99 #797) +#798 := (iff #94 #94) +#799 := [refl]: #798 +#801 := [quant-intro #799]: #800 +#206 := (~ #99 #99) +#222 := (~ #94 #94) +#223 := [refl]: #222 +#196 := [nnf-pos #223]: #206 +#14 := (= #13 #10) +#11 := (<= 0::int #10) +#15 := (implies #11 #14) +#16 := (forall (vars (?x2 int)) #15) +#102 := (iff #16 #99) +#76 := (not #11) +#77 := (or #76 #69) +#82 := (forall (vars (?x2 int)) #77) +#100 := (iff #82 #99) +#97 := (iff #77 #94) +#91 := (or #88 #69) +#95 := (iff #91 #94) +#96 := [rewrite]: #95 +#92 := (iff #77 #91) +#89 := (iff #76 #88) +#85 := (iff #11 #87) +#86 := [rewrite]: #85 +#90 := [monotonicity #86]: #89 +#93 := [monotonicity #90]: #92 +#98 := [trans #93 #96]: #97 +#101 := [quant-intro #98]: #100 +#83 := (iff #16 #82) +#80 := (iff #15 #77) +#73 := (implies #11 #69) +#78 := (iff #73 #77) +#79 := [rewrite]: #78 +#74 := (iff #15 #73) +#71 := (iff #14 #69) +#72 := [rewrite]: #71 +#75 := [monotonicity #72]: #74 +#81 := [trans #75 #79]: #80 +#84 := [quant-intro #81]: #83 +#103 := [trans #84 #101]: #102 +#68 := [asserted]: #16 +#104 := [mp #68 #103]: #99 +#224 := [mp~ #104 #196]: #99 +#802 := [mp #224 #801]: #797 +#559 := (not #797) +#511 := (or #559 #510 #517) +#531 := (not #529) +#526 := (= #186 #524) +#527 := (or #526 #531) +#523 := (or #559 #527) +#507 := (iff #523 #511) +#502 := (or #559 #520) +#505 := (iff #502 #511) +#506 := [rewrite]: #505 +#503 := (iff #523 #502) +#521 := (iff #527 #520) +#518 := (iff #531 #517) +#519 := [monotonicity #516]: #518 +#512 := (iff #526 #510) +#513 := [rewrite]: #512 +#522 := [monotonicity #513 #519]: #521 +#504 := [monotonicity #522]: #503 +#508 := [trans #504 #506]: #507 +#500 := [quant-inst]: #523 +#501 := [mp #500 #508]: #511 +#952 := [unit-resolution #501 #802]: #520 +#953 := [unit-resolution #952 #951]: #510 +#954 := (not #510) +#955 := (or #954 #494) +#956 := [th-lemma]: #955 +#957 := [unit-resolution #956 #953]: #494 +#959 := (not #494) +#960 := (or #958 #452 #959) +#961 := [th-lemma]: #960 +#962 := [unit-resolution #961 #957 #941]: #958 +#964 := (or #963 #537) +#965 := [th-lemma]: #964 +#966 := [unit-resolution #965 #962]: #963 +#583 := (>= #38 0::int) +#584 := (not #583) +#556 := (* -1::int #40) +#557 := (+ #38 #556) +#558 := (= #557 0::int) +#971 := (not #558) +#544 := (>= #557 0::int) +#967 := (not #544) +#201 := (>= #37 1::int) +#202 := (not #201) +#44 := (<= 1::int #37) +#45 := (not #44) +#203 := (iff #45 #202) +#199 := (iff #44 #201) +#200 := [rewrite]: #199 +#204 := [monotonicity #200]: #203 +#195 := [asserted]: #45 +#205 := [mp #195 #204]: #202 +#968 := (or #967 #201 #452 #959) +#969 := [th-lemma]: #968 +#970 := [unit-resolution #969 #205 #957 #941]: #967 +#972 := (or #971 #544) +#973 := [th-lemma]: #972 +#974 := [unit-resolution #973 #970]: #971 +#562 := (or #558 #584) +#564 := (or #559 #558 #584) +#567 := (= #38 #40) +#585 := (or #567 #584) +#543 := (or #559 #585) +#542 := (iff #543 #564) +#550 := (or #559 #562) +#551 := (iff #550 #564) +#554 := [rewrite]: #551 +#552 := (iff #543 #550) +#404 := (iff #585 #562) +#560 := (iff #567 #558) +#561 := [rewrite]: #560 +#563 := [monotonicity #561]: #404 +#553 := [monotonicity #563]: #552 +#545 := [trans #553 #554]: #542 +#546 := [quant-inst]: #543 +#547 := [mp #546 #545]: #564 +#975 := [unit-resolution #547 #802]: #562 +#976 := [unit-resolution #975 #974]: #584 +#539 := (or #549 #583) +#535 := (or #532 #549 #583) +#536 := (or #532 #539) +#533 := (iff #536 #535) +#541 := [rewrite]: #533 +#540 := [quant-inst]: #536 +#534 := [mp #540 #541]: #535 +#977 := [unit-resolution #534 #808]: #539 +[unit-resolution #977 #976 #966]: false +unsat diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_nlarith_01 --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_nlarith_01 Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,10 @@ +(benchmark Isabelle +:extrafuns ( + (uf_1 Int) + (uf_2 Int) + ) +:assumption (< 0 uf_1) +:assumption (< 0 (* uf_1 uf_2)) +:assumption (not (< 0 uf_2)) +:formula true +) diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_nlarith_01.proof --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_nlarith_01.proof Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,1 @@ +unsat diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_nlarith_02 --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_nlarith_02 Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,9 @@ +(benchmark Isabelle +:extrafuns ( + (uf_1 Int) + (uf_2 Int) + (uf_3 Int) + ) +:assumption (not (= (* uf_1 (+ (+ uf_2 1) uf_3)) (+ (* uf_1 uf_2) (* uf_1 (+ uf_3 1))))) +:formula true +) diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_nlarith_02.proof --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_nlarith_02.proof Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,75 @@ +#2 := false +#6 := 1::int +decl uf_3 :: int +#8 := uf_3 +#12 := (+ uf_3 1::int) +decl uf_1 :: int +#4 := uf_1 +#13 := (* uf_1 #12) +decl uf_2 :: int +#5 := uf_2 +#11 := (* uf_1 uf_2) +#14 := (+ #11 #13) +#7 := (+ uf_2 1::int) +#9 := (+ #7 uf_3) +#10 := (* uf_1 #9) +#15 := (= #10 #14) +#16 := (not #15) +#85 := (iff #16 false) +#1 := true +#80 := (not true) +#83 := (iff #80 false) +#84 := [rewrite]: #83 +#81 := (iff #16 #80) +#78 := (iff #15 true) +#48 := (* uf_1 uf_3) +#49 := (+ #11 #48) +#50 := (+ uf_1 #49) +#73 := (= #50 #50) +#76 := (iff #73 true) +#77 := [rewrite]: #76 +#74 := (iff #15 #73) +#71 := (= #14 #50) +#61 := (+ uf_1 #48) +#66 := (+ #11 #61) +#69 := (= #66 #50) +#70 := [rewrite]: #69 +#67 := (= #14 #66) +#64 := (= #13 #61) +#55 := (+ 1::int uf_3) +#58 := (* uf_1 #55) +#62 := (= #58 #61) +#63 := [rewrite]: #62 +#59 := (= #13 #58) +#56 := (= #12 #55) +#57 := [rewrite]: #56 +#60 := [monotonicity #57]: #59 +#65 := [trans #60 #63]: #64 +#68 := [monotonicity #65]: #67 +#72 := [trans #68 #70]: #71 +#53 := (= #10 #50) +#39 := (+ uf_2 uf_3) +#40 := (+ 1::int #39) +#45 := (* uf_1 #40) +#51 := (= #45 #50) +#52 := [rewrite]: #51 +#46 := (= #10 #45) +#43 := (= #9 #40) +#33 := (+ 1::int uf_2) +#36 := (+ #33 uf_3) +#41 := (= #36 #40) +#42 := [rewrite]: #41 +#37 := (= #9 #36) +#34 := (= #7 #33) +#35 := [rewrite]: #34 +#38 := [monotonicity #35]: #37 +#44 := [trans #38 #42]: #43 +#47 := [monotonicity #44]: #46 +#54 := [trans #47 #52]: #53 +#75 := [monotonicity #54 #72]: #74 +#79 := [trans #75 #77]: #78 +#82 := [monotonicity #79]: #81 +#86 := [trans #82 #84]: #85 +#32 := [asserted]: #16 +[mp #32 #86]: false +unsat diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_nlarith_03 --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_nlarith_03 Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,8 @@ +(benchmark Isabelle +:extrafuns ( + (uf_1 Real) + (uf_2 Real) + ) +:assumption (not (= (- (* uf_1 (+ 1.0 uf_2)) (* uf_1 (- 1.0 uf_2))) (* (* 2.0 uf_1) uf_2))) +:formula true +) diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_nlarith_03.proof --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_nlarith_03.proof Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,62 @@ +#2 := false +decl uf_2 :: real +#6 := uf_2 +decl uf_1 :: real +#4 := uf_1 +#12 := 2::real +#13 := (* 2::real uf_1) +#14 := (* #13 uf_2) +#5 := 1::real +#9 := (- 1::real uf_2) +#10 := (* uf_1 #9) +#7 := (+ 1::real uf_2) +#8 := (* uf_1 #7) +#11 := (- #8 #10) +#15 := (= #11 #14) +#16 := (not #15) +#73 := (iff #16 false) +#1 := true +#68 := (not true) +#71 := (iff #68 false) +#72 := [rewrite]: #71 +#69 := (iff #16 #68) +#66 := (iff #15 true) +#33 := (* uf_1 uf_2) +#55 := (* 2::real #33) +#61 := (= #55 #55) +#64 := (iff #61 true) +#65 := [rewrite]: #64 +#62 := (iff #15 #61) +#59 := (= #14 #55) +#60 := [rewrite]: #59 +#57 := (= #11 #55) +#37 := -1::real +#45 := (* -1::real #33) +#46 := (+ uf_1 #45) +#34 := (+ uf_1 #33) +#51 := (- #34 #46) +#54 := (= #51 #55) +#56 := [rewrite]: #54 +#52 := (= #11 #51) +#49 := (= #10 #46) +#38 := (* -1::real uf_2) +#39 := (+ 1::real #38) +#42 := (* uf_1 #39) +#47 := (= #42 #46) +#48 := [rewrite]: #47 +#43 := (= #10 #42) +#40 := (= #9 #39) +#41 := [rewrite]: #40 +#44 := [monotonicity #41]: #43 +#50 := [trans #44 #48]: #49 +#35 := (= #8 #34) +#36 := [rewrite]: #35 +#53 := [monotonicity #36 #50]: #52 +#58 := [trans #53 #56]: #57 +#63 := [monotonicity #58 #60]: #62 +#67 := [trans #63 #65]: #66 +#70 := [monotonicity #67]: #69 +#74 := [trans #70 #72]: #73 +#32 := [asserted]: #16 +[mp #32 #74]: false +unsat diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_nlarith_04 --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_nlarith_04 Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,11 @@ +(benchmark Isabelle +:extrafuns ( + (uf_1 Int) + (uf_3 Int) + (uf_5 Int) + (uf_4 Int) + (uf_2 Int) + ) +:assumption (not (= (+ (+ uf_1 (* (+ 1 uf_2) (+ uf_3 uf_4))) (* uf_2 uf_5)) (- (+ uf_1 (+ (+ (* (* 2 (+ 1 uf_2)) (+ uf_3 uf_4)) (* (+ 1 uf_2) uf_5)) (* uf_5 uf_2))) (* (+ 1 uf_2) (+ (+ uf_3 uf_5) uf_4))))) +:formula true +) diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_nlarith_04.proof --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_nlarith_04.proof Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,141 @@ +#2 := false +decl uf_4 :: int +#9 := uf_4 +decl uf_5 :: int +#13 := uf_5 +decl uf_3 :: int +#8 := uf_3 +#24 := (+ uf_3 uf_5) +#25 := (+ #24 uf_4) +decl uf_2 :: int +#6 := uf_2 +#5 := 1::int +#7 := (+ 1::int uf_2) +#26 := (* #7 #25) +#21 := (* uf_5 uf_2) +#19 := (* #7 uf_5) +#10 := (+ uf_3 uf_4) +#16 := 2::int +#17 := (* 2::int #7) +#18 := (* #17 #10) +#20 := (+ #18 #19) +#22 := (+ #20 #21) +decl uf_1 :: int +#4 := uf_1 +#23 := (+ uf_1 #22) +#27 := (- #23 #26) +#14 := (* uf_2 uf_5) +#11 := (* #7 #10) +#12 := (+ uf_1 #11) +#15 := (+ #12 #14) +#28 := (= #15 #27) +#29 := (not #28) +#149 := (iff #29 false) +#1 := true +#144 := (not true) +#147 := (iff #144 false) +#148 := [rewrite]: #147 +#145 := (iff #29 #144) +#142 := (iff #28 true) +#47 := (* uf_2 uf_4) +#46 := (* uf_2 uf_3) +#48 := (+ #46 #47) +#59 := (+ #14 #48) +#60 := (+ uf_4 #59) +#61 := (+ uf_3 #60) +#62 := (+ uf_1 #61) +#136 := (= #62 #62) +#140 := (iff #136 true) +#141 := [rewrite]: #140 +#135 := (iff #28 #136) +#138 := (= #27 #62) +#123 := (+ uf_5 #59) +#124 := (+ uf_4 #123) +#125 := (+ uf_3 #124) +#77 := (* 2::int #47) +#75 := (* 2::int #46) +#78 := (+ #75 #77) +#104 := (* 2::int #14) +#105 := (+ #104 #78) +#106 := (+ uf_5 #105) +#76 := (* 2::int uf_4) +#107 := (+ #76 #106) +#74 := (* 2::int uf_3) +#108 := (+ #74 #107) +#113 := (+ uf_1 #108) +#130 := (- #113 #125) +#133 := (= #130 #62) +#139 := [rewrite]: #133 +#131 := (= #27 #130) +#128 := (= #26 #125) +#116 := (+ uf_4 uf_5) +#117 := (+ uf_3 #116) +#120 := (* #7 #117) +#126 := (= #120 #125) +#127 := [rewrite]: #126 +#121 := (= #26 #120) +#118 := (= #25 #117) +#119 := [rewrite]: #118 +#122 := [monotonicity #119]: #121 +#129 := [trans #122 #127]: #128 +#114 := (= #23 #113) +#111 := (= #22 #108) +#91 := (+ #14 #78) +#92 := (+ uf_5 #91) +#93 := (+ #76 #92) +#94 := (+ #74 #93) +#101 := (+ #94 #14) +#109 := (= #101 #108) +#110 := [rewrite]: #109 +#102 := (= #22 #101) +#99 := (= #21 #14) +#100 := [rewrite]: #99 +#97 := (= #20 #94) +#85 := (+ uf_5 #14) +#79 := (+ #76 #78) +#80 := (+ #74 #79) +#88 := (+ #80 #85) +#95 := (= #88 #94) +#96 := [rewrite]: #95 +#89 := (= #20 #88) +#86 := (= #19 #85) +#87 := [rewrite]: #86 +#83 := (= #18 #80) +#67 := (* 2::int uf_2) +#68 := (+ 2::int #67) +#71 := (* #68 #10) +#81 := (= #71 #80) +#82 := [rewrite]: #81 +#72 := (= #18 #71) +#69 := (= #17 #68) +#70 := [rewrite]: #69 +#73 := [monotonicity #70]: #72 +#84 := [trans #73 #82]: #83 +#90 := [monotonicity #84 #87]: #89 +#98 := [trans #90 #96]: #97 +#103 := [monotonicity #98 #100]: #102 +#112 := [trans #103 #110]: #111 +#115 := [monotonicity #112]: #114 +#132 := [monotonicity #115 #129]: #131 +#137 := [trans #132 #139]: #138 +#65 := (= #15 #62) +#49 := (+ uf_4 #48) +#50 := (+ uf_3 #49) +#53 := (+ uf_1 #50) +#56 := (+ #53 #14) +#63 := (= #56 #62) +#64 := [rewrite]: #63 +#57 := (= #15 #56) +#54 := (= #12 #53) +#51 := (= #11 #50) +#52 := [rewrite]: #51 +#55 := [monotonicity #52]: #54 +#58 := [monotonicity #55]: #57 +#66 := [trans #58 #64]: #65 +#134 := [monotonicity #66 #137]: #135 +#143 := [trans #134 #141]: #142 +#146 := [monotonicity #143]: #145 +#150 := [trans #146 #148]: #149 +#45 := [asserted]: #29 +[mp #45 #150]: false +unsat diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_pair_01 --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_pair_01 Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,17 @@ +(benchmark Isabelle +:extrasorts ( T2 T3 T1) +:extrafuns ( + (uf_2 T1 T2) + (uf_3 T1 T3) + (uf_1 T2 T3 T1) + (uf_6 T2) + (uf_4 T2) + (uf_5 T3) + ) +:assumption (forall (?x1 T1) (= (uf_1 (uf_2 ?x1) (uf_3 ?x1)) ?x1)) +:assumption (forall (?x2 T2) (?x3 T3) (= (uf_3 (uf_1 ?x2 ?x3)) ?x3)) +:assumption (forall (?x4 T2) (?x5 T3) (= (uf_2 (uf_1 ?x4 ?x5)) ?x4)) +:assumption (= (uf_2 (uf_1 uf_4 uf_5)) uf_6) +:assumption (not (= uf_4 uf_6)) +:formula true +) diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_pair_01.proof --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_pair_01.proof Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,50 @@ +#2 := false +decl uf_6 :: T2 +#23 := uf_6 +decl uf_4 :: T2 +#19 := uf_4 +#25 := (= uf_4 uf_6) +decl uf_2 :: (-> T1 T2) +decl uf_1 :: (-> T2 T3 T1) +decl uf_5 :: T3 +#20 := uf_5 +#21 := (uf_1 uf_4 uf_5) +#22 := (uf_2 #21) +#24 := (= #22 uf_6) +#65 := [asserted]: #24 +#143 := (= uf_4 #22) +#11 := (:var 0 T3) +#10 := (:var 1 T2) +#12 := (uf_1 #10 #11) +#567 := (pattern #12) +#16 := (uf_2 #12) +#58 := (= #10 #16) +#574 := (forall (vars (?x4 T2) (?x5 T3)) (:pat #567) #58) +#62 := (forall (vars (?x4 T2) (?x5 T3)) #58) +#577 := (iff #62 #574) +#575 := (iff #58 #58) +#576 := [refl]: #575 +#578 := [quant-intro #576]: #577 +#71 := (~ #62 #62) +#87 := (~ #58 #58) +#88 := [refl]: #87 +#72 := [nnf-pos #88]: #71 +#17 := (= #16 #10) +#18 := (forall (vars (?x4 T2) (?x5 T3)) #17) +#63 := (iff #18 #62) +#60 := (iff #17 #58) +#61 := [rewrite]: #60 +#64 := [quant-intro #61]: #63 +#57 := [asserted]: #18 +#67 := [mp #57 #64]: #62 +#89 := [mp~ #67 #72]: #62 +#579 := [mp #89 #578]: #574 +#214 := (not #574) +#551 := (or #214 #143) +#553 := [quant-inst]: #551 +#233 := [unit-resolution #553 #579]: #143 +#235 := [trans #233 #65]: #25 +#26 := (not #25) +#66 := [asserted]: #26 +[unit-resolution #66 #235]: false +unsat diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_pair_02 --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_pair_02 Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,24 @@ +(benchmark Isabelle +:extrasorts ( T2 T3 T1 T4) +:extrafuns ( + (uf_2 T1 T2) + (uf_5 T4 T3) + (uf_3 T1 T3) + (uf_6 T4 T2) + (uf_1 T2 T3 T1) + (uf_4 T3 T2 T4) + (uf_8 T2) + (uf_9 T3) + (uf_7 T1) + (uf_10 T4) + ) +:assumption (forall (?x1 T1) (= (uf_1 (uf_2 ?x1) (uf_3 ?x1)) ?x1)) +:assumption (forall (?x2 T4) (= (uf_4 (uf_5 ?x2) (uf_6 ?x2)) ?x2)) +:assumption (forall (?x3 T2) (?x4 T3) (= (uf_3 (uf_1 ?x3 ?x4)) ?x4)) +:assumption (forall (?x5 T3) (?x6 T2) (= (uf_6 (uf_4 ?x5 ?x6)) ?x6)) +:assumption (forall (?x7 T2) (?x8 T3) (= (uf_2 (uf_1 ?x7 ?x8)) ?x7)) +:assumption (forall (?x9 T3) (?x10 T2) (= (uf_5 (uf_4 ?x9 ?x10)) ?x9)) +:assumption (and (= uf_7 (uf_1 uf_8 uf_9)) (= uf_10 (uf_4 uf_9 uf_8))) +:assumption (not (= (uf_2 uf_7) (uf_6 uf_10))) +:formula true +) diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_pair_02.proof --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_pair_02.proof Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,105 @@ +#2 := false +decl uf_6 :: (-> T4 T2) +decl uf_10 :: T4 +#39 := uf_10 +#44 := (uf_6 uf_10) +decl uf_2 :: (-> T1 T2) +decl uf_7 :: T1 +#34 := uf_7 +#43 := (uf_2 uf_7) +#45 := (= #43 #44) +decl uf_4 :: (-> T3 T2 T4) +decl uf_8 :: T2 +#35 := uf_8 +decl uf_9 :: T3 +#36 := uf_9 +#40 := (uf_4 uf_9 uf_8) +#204 := (uf_6 #40) +#598 := (= #204 #44) +#595 := (= #44 #204) +#41 := (= uf_10 #40) +decl uf_1 :: (-> T2 T3 T1) +#37 := (uf_1 uf_8 uf_9) +#38 := (= uf_7 #37) +#42 := (and #38 #41) +#109 := [asserted]: #42 +#114 := [and-elim #109]: #41 +#256 := [monotonicity #114]: #595 +#599 := [symm #256]: #598 +#596 := (= #43 #204) +#269 := (= uf_8 #204) +#23 := (:var 0 T2) +#22 := (:var 1 T3) +#24 := (uf_4 #22 #23) +#643 := (pattern #24) +#25 := (uf_6 #24) +#86 := (= #23 #25) +#644 := (forall (vars (?x5 T3) (?x6 T2)) (:pat #643) #86) +#90 := (forall (vars (?x5 T3) (?x6 T2)) #86) +#647 := (iff #90 #644) +#645 := (iff #86 #86) +#646 := [refl]: #645 +#648 := [quant-intro #646]: #647 +#119 := (~ #90 #90) +#144 := (~ #86 #86) +#145 := [refl]: #144 +#120 := [nnf-pos #145]: #119 +#26 := (= #25 #23) +#27 := (forall (vars (?x5 T3) (?x6 T2)) #26) +#91 := (iff #27 #90) +#88 := (iff #26 #86) +#89 := [rewrite]: #88 +#92 := [quant-intro #89]: #91 +#85 := [asserted]: #27 +#95 := [mp #85 #92]: #90 +#146 := [mp~ #95 #120]: #90 +#649 := [mp #146 #648]: #644 +#613 := (not #644) +#619 := (or #613 #269) +#609 := [quant-inst]: #619 +#267 := [unit-resolution #609 #649]: #269 +#600 := (= #43 uf_8) +#289 := (uf_2 #37) +#259 := (= #289 uf_8) +#296 := (= uf_8 #289) +#17 := (:var 0 T3) +#16 := (:var 1 T2) +#18 := (uf_1 #16 #17) +#636 := (pattern #18) +#28 := (uf_2 #18) +#94 := (= #16 #28) +#650 := (forall (vars (?x7 T2) (?x8 T3)) (:pat #636) #94) +#98 := (forall (vars (?x7 T2) (?x8 T3)) #94) +#653 := (iff #98 #650) +#651 := (iff #94 #94) +#652 := [refl]: #651 +#654 := [quant-intro #652]: #653 +#121 := (~ #98 #98) +#147 := (~ #94 #94) +#148 := [refl]: #147 +#122 := [nnf-pos #148]: #121 +#29 := (= #28 #16) +#30 := (forall (vars (?x7 T2) (?x8 T3)) #29) +#99 := (iff #30 #98) +#96 := (iff #29 #94) +#97 := [rewrite]: #96 +#100 := [quant-intro #97]: #99 +#93 := [asserted]: #30 +#103 := [mp #93 #100]: #98 +#149 := [mp~ #103 #122]: #98 +#655 := [mp #149 #654]: #650 +#615 := (not #650) +#616 := (or #615 #296) +#617 := [quant-inst]: #616 +#618 := [unit-resolution #617 #655]: #296 +#597 := [symm #618]: #259 +#611 := (= #43 #289) +#113 := [and-elim #109]: #38 +#252 := [monotonicity #113]: #611 +#601 := [trans #252 #597]: #600 +#602 := [trans #601 #267]: #596 +#238 := [trans #602 #599]: #45 +#46 := (not #45) +#110 := [asserted]: #46 +[unit-resolution #110 #238]: false +unsat diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_prop_01 --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_prop_01 Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,4 @@ +(benchmark Isabelle +:assumption (not true) +:formula true +) diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_prop_01.proof --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_prop_01.proof Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,8 @@ +#2 := false +#1 := true +#4 := (not true) +#21 := (iff #4 false) +#22 := [rewrite]: #21 +#20 := [asserted]: #4 +[mp #20 #22]: false +unsat diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_prop_02 --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_prop_02 Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,8 @@ +(benchmark Isabelle +:extrasorts ( T1) +:extrapreds ( + (up_1) + ) +:assumption (not (or up_1 (not up_1))) +:formula true +) diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_prop_02.proof --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_prop_02.proof Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,19 @@ +#2 := false +decl up_1 :: bool +#4 := up_1 +#5 := (not up_1) +#6 := (or up_1 #5) +#7 := (not #6) +#31 := (iff #7 false) +#1 := true +#26 := (not true) +#29 := (iff #26 false) +#30 := [rewrite]: #29 +#27 := (iff #7 #26) +#24 := (iff #6 true) +#25 := [rewrite]: #24 +#28 := [monotonicity #25]: #27 +#32 := [trans #28 #30]: #31 +#23 := [asserted]: #7 +[mp #23 #32]: false +unsat diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_prop_03 --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_prop_03 Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,8 @@ +(benchmark Isabelle +:extrasorts ( T1) +:extrapreds ( + (up_1) + ) +:assumption (not (iff (and up_1 true) up_1)) +:formula true +) diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_prop_03.proof --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_prop_03.proof Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,25 @@ +#2 := false +decl up_1 :: bool +#4 := up_1 +#1 := true +#5 := (and up_1 true) +#6 := (iff #5 up_1) +#7 := (not #6) +#37 := (iff #7 false) +#32 := (not true) +#35 := (iff #32 false) +#36 := [rewrite]: #35 +#33 := (iff #7 #32) +#30 := (iff #6 true) +#25 := (iff up_1 up_1) +#28 := (iff #25 true) +#29 := [rewrite]: #28 +#26 := (iff #6 #25) +#24 := [rewrite]: #6 +#27 := [monotonicity #24]: #26 +#31 := [trans #27 #29]: #30 +#34 := [monotonicity #31]: #33 +#38 := [trans #34 #36]: #37 +#23 := [asserted]: #7 +[mp #23 #38]: false +unsat diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_prop_04 --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_prop_04 Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,10 @@ +(benchmark Isabelle +:extrasorts ( T1) +:extrapreds ( + (up_1) + (up_2) + ) +:assumption (and (or up_1 up_2) (not up_1)) +:assumption (not up_2) +:formula true +) diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_prop_04.proof --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_prop_04.proof Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,38 @@ +#2 := false +decl up_2 :: bool +#5 := up_2 +decl up_1 :: bool +#4 := up_1 +#6 := (or up_1 up_2) +#51 := (iff #6 false) +#46 := (or false false) +#49 := (iff #46 false) +#50 := [rewrite]: #49 +#47 := (iff #6 #46) +#40 := (iff up_2 false) +#9 := (not up_2) +#43 := (iff #9 #40) +#41 := (iff #40 #9) +#42 := [rewrite]: #41 +#44 := [symm #42]: #43 +#32 := [asserted]: #9 +#45 := [mp #32 #44]: #40 +#35 := (iff up_1 false) +#7 := (not up_1) +#37 := (iff #7 #35) +#33 := (iff #35 #7) +#36 := [rewrite]: #33 +#38 := [symm #36]: #37 +#26 := (and #7 #6) +#8 := (and #6 #7) +#27 := (iff #8 #26) +#28 := [rewrite]: #27 +#25 := [asserted]: #8 +#31 := [mp #25 #28]: #26 +#29 := [and-elim #31]: #7 +#39 := [mp #29 #38]: #35 +#48 := [monotonicity #39 #45]: #47 +#52 := [trans #48 #50]: #51 +#30 := [and-elim #31]: #6 +[mp #30 #52]: false +unsat diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_prop_05 --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_prop_05 Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,12 @@ +(benchmark Isabelle +:extrasorts ( T1) +:extrapreds ( + (up_1) + (up_2) + (up_3) + (up_4) + ) +:assumption (or (and up_1 up_2) (and up_3 up_4)) +:assumption (not (or (and up_1 up_2) (and up_3 up_4))) +:formula true +) diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_prop_05.proof --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_prop_05.proof Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,1 @@ +unsat diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_prop_06 --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_prop_06 Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,10 @@ +(benchmark Isabelle +:extrasorts ( T1) +:extrapreds ( + (up_1) + (up_2) + (up_3) + ) +:assumption (not (implies (or (and up_1 up_2) up_3) (or (implies up_1 (or (and up_3 up_2) (and up_1 up_3))) up_1))) +:formula true +) diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_prop_06.proof --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_prop_06.proof Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,71 @@ +#2 := false +decl up_1 :: bool +#4 := up_1 +decl up_3 :: bool +#7 := up_3 +#10 := (and up_1 up_3) +decl up_2 :: bool +#5 := up_2 +#9 := (and up_3 up_2) +#11 := (or #9 #10) +#12 := (implies up_1 #11) +#13 := (or #12 up_1) +#6 := (and up_1 up_2) +#8 := (or #6 up_3) +#14 := (implies #8 #13) +#15 := (not #14) +#81 := (iff #15 false) +#32 := (and up_2 up_3) +#38 := (or #10 #32) +#46 := (not up_1) +#47 := (or #46 #38) +#55 := (or up_1 #47) +#63 := (not #8) +#64 := (or #63 #55) +#69 := (not #64) +#79 := (iff #69 false) +#1 := true +#74 := (not true) +#77 := (iff #74 false) +#78 := [rewrite]: #77 +#75 := (iff #69 #74) +#72 := (iff #64 true) +#73 := [rewrite]: #72 +#76 := [monotonicity #73]: #75 +#80 := [trans #76 #78]: #79 +#70 := (iff #15 #69) +#67 := (iff #14 #64) +#60 := (implies #8 #55) +#65 := (iff #60 #64) +#66 := [rewrite]: #65 +#61 := (iff #14 #60) +#58 := (iff #13 #55) +#52 := (or #47 up_1) +#56 := (iff #52 #55) +#57 := [rewrite]: #56 +#53 := (iff #13 #52) +#50 := (iff #12 #47) +#43 := (implies up_1 #38) +#48 := (iff #43 #47) +#49 := [rewrite]: #48 +#44 := (iff #12 #43) +#41 := (iff #11 #38) +#35 := (or #32 #10) +#39 := (iff #35 #38) +#40 := [rewrite]: #39 +#36 := (iff #11 #35) +#33 := (iff #9 #32) +#34 := [rewrite]: #33 +#37 := [monotonicity #34]: #36 +#42 := [trans #37 #40]: #41 +#45 := [monotonicity #42]: #44 +#51 := [trans #45 #49]: #50 +#54 := [monotonicity #51]: #53 +#59 := [trans #54 #57]: #58 +#62 := [monotonicity #59]: #61 +#68 := [trans #62 #66]: #67 +#71 := [monotonicity #68]: #70 +#82 := [trans #71 #80]: #81 +#31 := [asserted]: #15 +[mp #31 #82]: false +unsat diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_prop_07 --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_prop_07 Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,8 @@ +(benchmark Isabelle +:extrasorts ( T1) +:extrapreds ( + (up_1) + ) +:assumption (not (iff (iff (iff (iff (iff (iff (iff (iff (iff up_1 up_1) up_1) up_1) up_1) up_1) up_1) up_1) up_1) up_1)) +:formula true +) diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_prop_07.proof --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_prop_07.proof Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,57 @@ +#2 := false +decl up_1 :: bool +#4 := up_1 +#5 := (iff up_1 up_1) +#6 := (iff #5 up_1) +#7 := (iff #6 up_1) +#8 := (iff #7 up_1) +#9 := (iff #8 up_1) +#10 := (iff #9 up_1) +#11 := (iff #10 up_1) +#12 := (iff #11 up_1) +#13 := (iff #12 up_1) +#14 := (not #13) +#69 := (iff #14 false) +#1 := true +#64 := (not true) +#67 := (iff #64 false) +#68 := [rewrite]: #67 +#65 := (iff #14 #64) +#62 := (iff #13 true) +#31 := (iff #5 true) +#32 := [rewrite]: #31 +#60 := (iff #13 #5) +#33 := (iff true up_1) +#36 := (iff #33 up_1) +#37 := [rewrite]: #36 +#57 := (iff #12 #33) +#55 := (iff #11 true) +#53 := (iff #11 #5) +#50 := (iff #10 #33) +#48 := (iff #9 true) +#46 := (iff #9 #5) +#43 := (iff #8 #33) +#41 := (iff #7 true) +#39 := (iff #7 #5) +#34 := (iff #6 #33) +#35 := [monotonicity #32]: #34 +#38 := [trans #35 #37]: #7 +#40 := [monotonicity #38]: #39 +#42 := [trans #40 #32]: #41 +#44 := [monotonicity #42]: #43 +#45 := [trans #44 #37]: #9 +#47 := [monotonicity #45]: #46 +#49 := [trans #47 #32]: #48 +#51 := [monotonicity #49]: #50 +#52 := [trans #51 #37]: #11 +#54 := [monotonicity #52]: #53 +#56 := [trans #54 #32]: #55 +#58 := [monotonicity #56]: #57 +#59 := [trans #58 #37]: #13 +#61 := [monotonicity #59]: #60 +#63 := [trans #61 #32]: #62 +#66 := [monotonicity #63]: #65 +#70 := [trans #66 #68]: #69 +#30 := [asserted]: #14 +[mp #30 #70]: false +unsat diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_prop_08 --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_prop_08 Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,22 @@ +(benchmark Isabelle +:extrasorts ( T1) +:extrapreds ( + (up_1) + (up_2) + (up_3) + (up_4) + (up_5) + (up_6) + (up_8) + (up_9) + (up_7) + ) +:assumption (or up_1 (or up_2 (or up_3 up_4))) +:assumption (or up_5 (or up_6 (and up_1 up_4))) +:assumption (or (not (or up_1 (and up_3 (not up_3)))) up_2) +:assumption (or (not (and up_2 (or up_7 (not up_7)))) up_3) +:assumption (or (not (or up_4 false)) up_3) +:assumption (not (or up_3 (and (not up_8) (or up_8 (and up_9 (not up_9)))))) +:assumption (not false) +:formula true +) diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_prop_08.proof --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_prop_08.proof Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,194 @@ +#2 := false +decl up_1 :: bool +#4 := up_1 +#75 := (not up_1) +#246 := (iff #75 false) +#1 := true +#214 := (not true) +#217 := (iff #214 false) +#218 := [rewrite]: #217 +#244 := (iff #75 #214) +#238 := (iff up_1 true) +#241 := (iff up_1 #238) +#239 := (iff #238 up_1) +#240 := [rewrite]: #239 +#242 := [symm #240]: #241 +decl up_4 :: bool +#7 := up_4 +decl up_2 :: bool +#5 := up_2 +#161 := (or up_1 up_2 up_4) +#200 := (iff #161 up_1) +#195 := (or up_1 false false) +#198 := (iff #195 up_1) +#199 := [rewrite]: #198 +#196 := (iff #161 #195) +#189 := (iff up_4 false) +#102 := (not up_4) +#192 := (iff #102 #189) +#190 := (iff #189 #102) +#191 := [rewrite]: #190 +#193 := [symm #191]: #192 +decl up_3 :: bool +#6 := up_3 +#108 := (or up_3 #102) +#180 := (iff #108 #102) +#175 := (or false #102) +#178 := (iff #175 #102) +#179 := [rewrite]: #178 +#176 := (iff #108 #175) +#152 := (iff up_3 false) +#16 := (not up_3) +#155 := (iff #16 #152) +#153 := (iff #152 #16) +#154 := [rewrite]: #153 +#156 := [symm #154]: #155 +decl up_9 :: bool +#32 := up_9 +#33 := (not up_9) +#34 := (and up_9 #33) +decl up_8 :: bool +#30 := up_8 +#35 := (or up_8 #34) +#31 := (not up_8) +#36 := (and #31 #35) +#37 := (or up_3 #36) +#38 := (not #37) +#138 := (iff #38 #16) +#136 := (iff #37 up_3) +#131 := (or up_3 false) +#134 := (iff #131 up_3) +#135 := [rewrite]: #134 +#132 := (iff #37 #131) +#129 := (iff #36 false) +#124 := (and #31 up_8) +#127 := (iff #124 false) +#128 := [rewrite]: #127 +#125 := (iff #36 #124) +#122 := (iff #35 up_8) +#117 := (or up_8 false) +#120 := (iff #117 up_8) +#121 := [rewrite]: #120 +#118 := (iff #35 #117) +#114 := (iff #34 false) +#116 := [rewrite]: #114 +#119 := [monotonicity #116]: #118 +#123 := [trans #119 #121]: #122 +#126 := [monotonicity #123]: #125 +#130 := [trans #126 #128]: #129 +#133 := [monotonicity #130]: #132 +#137 := [trans #133 #135]: #136 +#139 := [monotonicity #137]: #138 +#113 := [asserted]: #38 +#142 := [mp #113 #139]: #16 +#157 := [mp #142 #156]: #152 +#177 := [monotonicity #157]: #176 +#181 := [trans #177 #179]: #180 +#27 := (or up_4 false) +#28 := (not #27) +#29 := (or #28 up_3) +#111 := (iff #29 #108) +#105 := (or #102 up_3) +#109 := (iff #105 #108) +#110 := [rewrite]: #109 +#106 := (iff #29 #105) +#103 := (iff #28 #102) +#99 := (iff #27 up_4) +#101 := [rewrite]: #99 +#104 := [monotonicity #101]: #103 +#107 := [monotonicity #104]: #106 +#112 := [trans #107 #110]: #111 +#98 := [asserted]: #29 +#115 := [mp #98 #112]: #108 +#182 := [mp #115 #181]: #102 +#194 := [mp #182 #193]: #189 +#183 := (iff up_2 false) +#92 := (not up_2) +#186 := (iff #92 #183) +#184 := (iff #183 #92) +#185 := [rewrite]: #184 +#187 := [symm #185]: #186 +#95 := (or #92 up_3) +#172 := (iff #95 #92) +#167 := (or #92 false) +#170 := (iff #167 #92) +#171 := [rewrite]: #170 +#168 := (iff #95 #167) +#169 := [monotonicity #157]: #168 +#173 := [trans #169 #171]: #172 +decl up_7 :: bool +#21 := up_7 +#22 := (not up_7) +#23 := (or up_7 #22) +#24 := (and up_2 #23) +#25 := (not #24) +#26 := (or #25 up_3) +#96 := (iff #26 #95) +#93 := (iff #25 #92) +#90 := (iff #24 up_2) +#85 := (and up_2 true) +#88 := (iff #85 up_2) +#89 := [rewrite]: #88 +#86 := (iff #24 #85) +#82 := (iff #23 true) +#84 := [rewrite]: #82 +#87 := [monotonicity #84]: #86 +#91 := [trans #87 #89]: #90 +#94 := [monotonicity #91]: #93 +#97 := [monotonicity #94]: #96 +#81 := [asserted]: #26 +#100 := [mp #81 #97]: #95 +#174 := [mp #100 #173]: #92 +#188 := [mp #174 #187]: #183 +#197 := [monotonicity #188 #194]: #196 +#201 := [trans #197 #199]: #200 +#58 := (or up_1 up_2 up_3 up_4) +#164 := (iff #58 #161) +#158 := (or up_1 up_2 false up_4) +#162 := (iff #158 #161) +#163 := [rewrite]: #162 +#159 := (iff #58 #158) +#160 := [monotonicity #157]: #159 +#165 := [trans #160 #163]: #164 +#8 := (or up_3 up_4) +#9 := (or up_2 #8) +#10 := (or up_1 #9) +#59 := (iff #10 #58) +#60 := [rewrite]: #59 +#55 := [asserted]: #10 +#61 := [mp #55 #60]: #58 +#166 := [mp #61 #165]: #161 +#202 := [mp #166 #201]: up_1 +#243 := [mp #202 #242]: #238 +#245 := [monotonicity #243]: #244 +#247 := [trans #245 #218]: #246 +#78 := (or #75 up_2) +#235 := (iff #78 #75) +#230 := (or #75 false) +#233 := (iff #230 #75) +#234 := [rewrite]: #233 +#231 := (iff #78 #230) +#232 := [monotonicity #188]: #231 +#236 := [trans #232 #234]: #235 +#17 := (and up_3 #16) +#18 := (or up_1 #17) +#19 := (not #18) +#20 := (or #19 up_2) +#79 := (iff #20 #78) +#76 := (iff #19 #75) +#73 := (iff #18 up_1) +#68 := (or up_1 false) +#71 := (iff #68 up_1) +#72 := [rewrite]: #71 +#69 := (iff #18 #68) +#62 := (iff #17 false) +#67 := [rewrite]: #62 +#70 := [monotonicity #67]: #69 +#74 := [trans #70 #72]: #73 +#77 := [monotonicity #74]: #76 +#80 := [monotonicity #77]: #79 +#57 := [asserted]: #20 +#83 := [mp #57 #80]: #78 +#237 := [mp #83 #236]: #75 +[mp #237 #247]: false +unsat diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_prop_09 --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_prop_09 Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,11 @@ +(benchmark Isabelle +:extrasorts ( T1) +:extrafuns ( + (uf_1 T1 T1 T1) + (uf_2 T1) + (uf_3 T1) + ) +:assumption (forall (?x1 T1) (?x2 T1) (= (uf_1 ?x1 ?x2) (uf_1 ?x2 ?x1))) +:assumption (not (and (= uf_2 uf_2) (= (uf_1 uf_2 uf_3) (uf_1 uf_3 uf_2)))) +:formula true +) diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_prop_09.proof --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_prop_09.proof Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,52 @@ +#2 := false +decl uf_1 :: (-> T1 T1 T1) +decl uf_2 :: T1 +#10 := uf_2 +decl uf_3 :: T1 +#12 := uf_3 +#14 := (uf_1 uf_3 uf_2) +#13 := (uf_1 uf_2 uf_3) +#15 := (= #13 #14) +#44 := (not #15) +#11 := (= uf_2 uf_2) +#16 := (and #11 #15) +#17 := (not #16) +#45 := (iff #17 #44) +#42 := (iff #16 #15) +#1 := true +#37 := (and true #15) +#40 := (iff #37 #15) +#41 := [rewrite]: #40 +#38 := (iff #16 #37) +#35 := (iff #11 true) +#36 := [rewrite]: #35 +#39 := [monotonicity #36]: #38 +#43 := [trans #39 #41]: #42 +#46 := [monotonicity #43]: #45 +#34 := [asserted]: #17 +#49 := [mp #34 #46]: #44 +#4 := (:var 1 T1) +#5 := (:var 0 T1) +#7 := (uf_1 #5 #4) +#530 := (pattern #7) +#6 := (uf_1 #4 #5) +#529 := (pattern #6) +#8 := (= #6 #7) +#531 := (forall (vars (?x1 T1) (?x2 T1)) (:pat #529 #530) #8) +#9 := (forall (vars (?x1 T1) (?x2 T1)) #8) +#534 := (iff #9 #531) +#532 := (iff #8 #8) +#533 := [refl]: #532 +#535 := [quant-intro #533]: #534 +#55 := (~ #9 #9) +#53 := (~ #8 #8) +#54 := [refl]: #53 +#56 := [nnf-pos #54]: #55 +#33 := [asserted]: #9 +#57 := [mp~ #33 #56]: #9 +#536 := [mp #57 #535]: #531 +#112 := (not #531) +#199 := (or #112 #15) +#113 := [quant-inst]: #199 +[unit-resolution #113 #536 #49]: false +unsat diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_prop_10 --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_prop_10 Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,251 @@ +(benchmark Isabelle +:extrasorts ( T1) +:extrapreds ( + (up_1) + (up_5) + (up_7) + (up_9) + (up_11) + (up_14) + (up_16) + (up_18) + (up_20) + (up_22) + (up_25) + (up_27) + (up_29) + (up_31) + (up_33) + (up_36) + (up_38) + (up_40) + (up_42) + (up_44) + (up_47) + (up_49) + (up_51) + (up_53) + (up_55) + (up_57) + (up_58) + (up_59) + (up_60) + (up_3) + (up_2) + (up_6) + (up_8) + (up_10) + (up_12) + (up_13) + (up_15) + (up_17) + (up_19) + (up_21) + (up_23) + (up_24) + (up_26) + (up_28) + (up_30) + (up_32) + (up_34) + (up_35) + (up_37) + (up_39) + (up_41) + (up_43) + (up_45) + (up_46) + (up_48) + (up_50) + (up_52) + (up_54) + (up_56) + (up_4) + ) +:assumption (not up_1) +:assumption (not up_2) +:assumption (not up_3) +:assumption (not up_4) +:assumption (or up_5 (or up_6 up_1)) +:assumption (or up_7 (or up_8 up_5)) +:assumption (or up_9 (or up_10 up_7)) +:assumption (or up_11 (or up_12 up_9)) +:assumption (or up_13 up_11) +:assumption (or up_14 (or up_15 up_2)) +:assumption (or up_16 (or up_17 (or up_14 up_6))) +:assumption (or up_18 (or up_19 (or up_16 up_8))) +:assumption (or up_20 (or up_21 (or up_18 up_10))) +:assumption (or up_22 (or up_23 (or up_20 up_12))) +:assumption (or up_24 (or up_22 up_13)) +:assumption (or up_25 (or up_26 up_15)) +:assumption (or up_27 (or up_28 (or up_25 up_17))) +:assumption (or up_29 (or up_30 (or up_27 up_19))) +:assumption (or up_31 (or up_32 (or up_29 up_21))) +:assumption (or up_33 (or up_34 (or up_31 up_23))) +:assumption (or up_35 (or up_33 up_24)) +:assumption (or up_36 (or up_37 up_26)) +:assumption (or up_38 (or up_39 (or up_36 up_28))) +:assumption (or up_40 (or up_41 (or up_38 up_30))) +:assumption (or up_42 (or up_43 (or up_40 up_32))) +:assumption (or up_44 (or up_45 (or up_42 up_34))) +:assumption (or up_46 (or up_44 up_35)) +:assumption (or up_47 (or up_48 up_37)) +:assumption (or up_49 (or up_50 (or up_47 up_39))) +:assumption (or up_51 (or up_52 (or up_49 up_41))) +:assumption (or up_53 (or up_54 (or up_51 up_43))) +:assumption (or up_55 (or up_56 (or up_53 up_45))) +:assumption (or up_4 (or up_55 up_46)) +:assumption (or up_57 up_48) +:assumption (or up_58 (or up_57 up_50)) +:assumption (or up_59 (or up_58 up_52)) +:assumption (or up_60 (or up_59 up_54)) +:assumption (or up_3 (or up_60 up_56)) +:assumption (or (not up_5) (not up_6)) +:assumption (or (not up_5) (not up_1)) +:assumption (or (not up_6) (not up_1)) +:assumption (or (not up_7) (not up_8)) +:assumption (or (not up_7) (not up_5)) +:assumption (or (not up_8) (not up_5)) +:assumption (or (not up_9) (not up_10)) +:assumption (or (not up_9) (not up_7)) +:assumption (or (not up_10) (not up_7)) +:assumption (or (not up_11) (not up_12)) +:assumption (or (not up_11) (not up_9)) +:assumption (or (not up_12) (not up_9)) +:assumption (or (not up_13) (not up_11)) +:assumption (or (not up_14) (not up_15)) +:assumption (or (not up_14) (not up_2)) +:assumption (or (not up_15) (not up_2)) +:assumption (or (not up_16) (not up_17)) +:assumption (or (not up_16) (not up_14)) +:assumption (or (not up_16) (not up_6)) +:assumption (or (not up_17) (not up_14)) +:assumption (or (not up_17) (not up_6)) +:assumption (or (not up_14) (not up_6)) +:assumption (or (not up_18) (not up_19)) +:assumption (or (not up_18) (not up_16)) +:assumption (or (not up_18) (not up_8)) +:assumption (or (not up_19) (not up_16)) +:assumption (or (not up_19) (not up_8)) +:assumption (or (not up_16) (not up_8)) +:assumption (or (not up_20) (not up_21)) +:assumption (or (not up_20) (not up_18)) +:assumption (or (not up_20) (not up_10)) +:assumption (or (not up_21) (not up_18)) +:assumption (or (not up_21) (not up_10)) +:assumption (or (not up_18) (not up_10)) +:assumption (or (not up_22) (not up_23)) +:assumption (or (not up_22) (not up_20)) +:assumption (or (not up_22) (not up_12)) +:assumption (or (not up_23) (not up_20)) +:assumption (or (not up_23) (not up_12)) +:assumption (or (not up_20) (not up_12)) +:assumption (or (not up_24) (not up_22)) +:assumption (or (not up_24) (not up_13)) +:assumption (or (not up_22) (not up_13)) +:assumption (or (not up_25) (not up_26)) +:assumption (or (not up_25) (not up_15)) +:assumption (or (not up_26) (not up_15)) +:assumption (or (not up_27) (not up_28)) +:assumption (or (not up_27) (not up_25)) +:assumption (or (not up_27) (not up_17)) +:assumption (or (not up_28) (not up_25)) +:assumption (or (not up_28) (not up_17)) +:assumption (or (not up_25) (not up_17)) +:assumption (or (not up_29) (not up_30)) +:assumption (or (not up_29) (not up_27)) +:assumption (or (not up_29) (not up_19)) +:assumption (or (not up_30) (not up_27)) +:assumption (or (not up_30) (not up_19)) +:assumption (or (not up_27) (not up_19)) +:assumption (or (not up_31) (not up_32)) +:assumption (or (not up_31) (not up_29)) +:assumption (or (not up_31) (not up_21)) +:assumption (or (not up_32) (not up_29)) +:assumption (or (not up_32) (not up_21)) +:assumption (or (not up_29) (not up_21)) +:assumption (or (not up_33) (not up_34)) +:assumption (or (not up_33) (not up_31)) +:assumption (or (not up_33) (not up_23)) +:assumption (or (not up_34) (not up_31)) +:assumption (or (not up_34) (not up_23)) +:assumption (or (not up_31) (not up_23)) +:assumption (or (not up_35) (not up_33)) +:assumption (or (not up_35) (not up_24)) +:assumption (or (not up_33) (not up_24)) +:assumption (or (not up_36) (not up_37)) +:assumption (or (not up_36) (not up_26)) +:assumption (or (not up_37) (not up_26)) +:assumption (or (not up_38) (not up_39)) +:assumption (or (not up_38) (not up_36)) +:assumption (or (not up_38) (not up_28)) +:assumption (or (not up_39) (not up_36)) +:assumption (or (not up_39) (not up_28)) +:assumption (or (not up_36) (not up_28)) +:assumption (or (not up_40) (not up_41)) +:assumption (or (not up_40) (not up_38)) +:assumption (or (not up_40) (not up_30)) +:assumption (or (not up_41) (not up_38)) +:assumption (or (not up_41) (not up_30)) +:assumption (or (not up_38) (not up_30)) +:assumption (or (not up_42) (not up_43)) +:assumption (or (not up_42) (not up_40)) +:assumption (or (not up_42) (not up_32)) +:assumption (or (not up_43) (not up_40)) +:assumption (or (not up_43) (not up_32)) +:assumption (or (not up_40) (not up_32)) +:assumption (or (not up_44) (not up_45)) +:assumption (or (not up_44) (not up_42)) +:assumption (or (not up_44) (not up_34)) +:assumption (or (not up_45) (not up_42)) +:assumption (or (not up_45) (not up_34)) +:assumption (or (not up_42) (not up_34)) +:assumption (or (not up_46) (not up_44)) +:assumption (or (not up_46) (not up_35)) +:assumption (or (not up_44) (not up_35)) +:assumption (or (not up_47) (not up_48)) +:assumption (or (not up_47) (not up_37)) +:assumption (or (not up_48) (not up_37)) +:assumption (or (not up_49) (not up_50)) +:assumption (or (not up_49) (not up_47)) +:assumption (or (not up_49) (not up_39)) +:assumption (or (not up_50) (not up_47)) +:assumption (or (not up_50) (not up_39)) +:assumption (or (not up_47) (not up_39)) +:assumption (or (not up_51) (not up_52)) +:assumption (or (not up_51) (not up_49)) +:assumption (or (not up_51) (not up_41)) +:assumption (or (not up_52) (not up_49)) +:assumption (or (not up_52) (not up_41)) +:assumption (or (not up_49) (not up_41)) +:assumption (or (not up_53) (not up_54)) +:assumption (or (not up_53) (not up_51)) +:assumption (or (not up_53) (not up_43)) +:assumption (or (not up_54) (not up_51)) +:assumption (or (not up_54) (not up_43)) +:assumption (or (not up_51) (not up_43)) +:assumption (or (not up_55) (not up_56)) +:assumption (or (not up_55) (not up_53)) +:assumption (or (not up_55) (not up_45)) +:assumption (or (not up_56) (not up_53)) +:assumption (or (not up_56) (not up_45)) +:assumption (or (not up_53) (not up_45)) +:assumption (or (not up_4) (not up_55)) +:assumption (or (not up_4) (not up_46)) +:assumption (or (not up_55) (not up_46)) +:assumption (or (not up_57) (not up_48)) +:assumption (or (not up_58) (not up_57)) +:assumption (or (not up_58) (not up_50)) +:assumption (or (not up_57) (not up_50)) +:assumption (or (not up_59) (not up_58)) +:assumption (or (not up_59) (not up_52)) +:assumption (or (not up_58) (not up_52)) +:assumption (or (not up_60) (not up_59)) +:assumption (or (not up_60) (not up_54)) +:assumption (or (not up_59) (not up_54)) +:assumption (or (not up_3) (not up_60)) +:assumption (or (not up_3) (not up_56)) +:assumption (or (not up_60) (not up_56)) +:assumption (not false) +:formula true +) diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Examples/cert/z3_prop_10.proof --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/cert/z3_prop_10.proof Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,1667 @@ +#2 := false +decl up_54 :: bool +#126 := up_54 +#317 := (not up_54) +decl up_60 :: bool +#145 := up_60 +decl up_56 :: bool +#131 := up_56 +#325 := (not up_56) +decl up_55 :: bool +#130 := up_55 +decl up_46 :: bool +#108 := up_46 +#291 := (not up_46) +decl up_35 :: bool +#81 := up_35 +decl up_29 :: bool +#66 := up_29 +decl up_32 :: bool +#72 := up_32 +#235 := (not up_32) +decl up_34 :: bool +#77 := up_34 +#243 := (not up_34) +decl up_33 :: bool +#76 := up_33 +#250 := (not up_35) +#1611 := [hypothesis]: #250 +decl up_24 :: bool +#54 := up_24 +#209 := (not up_24) +decl up_13 :: bool +#28 := up_13 +decl up_11 :: bool +#24 := up_11 +#165 := (not up_11) +decl up_12 :: bool +#25 := up_12 +#2327 := (or up_12 up_35) +#345 := (not up_60) +decl up_59 :: bool +#142 := up_59 +decl up_19 :: bool +#40 := up_19 +decl up_8 :: bool +#17 := up_8 +#156 := (not up_8) +decl up_7 :: bool +#16 := up_7 +#166 := (not up_12) +#1457 := [hypothesis]: #166 +#2183 := (or up_7 up_12 up_35) +#155 := (not up_7) +#1612 := [hypothesis]: #155 +decl up_10 :: bool +#21 := up_10 +#161 := (not up_10) +decl up_20 :: bool +#44 := up_20 +decl up_23 :: bool +#50 := up_23 +#202 := (not up_23) +#2170 := (or up_34 up_7 up_35 up_12) +#1605 := [hypothesis]: #243 +#2164 := (or up_29 up_34 up_7 up_35 up_12) +decl up_42 :: bool +#98 := up_42 +#275 := (not up_42) +#226 := (not up_29) +#907 := [hypothesis]: #226 +#2136 := (or up_29 up_12 up_7 up_35 up_32) +decl up_22 :: bool +#49 := up_22 +#895 := [hypothesis]: #235 +#1624 := (or up_29 up_22 up_12 up_32 up_35 up_7) +decl up_21 :: bool +#45 := up_21 +decl up_31 :: bool +#71 := up_31 +#234 := (not up_31) +decl up_9 :: bool +#20 := up_9 +#201 := (not up_22) +#1456 := [hypothesis]: #201 +#847 := (or #161 up_32 up_29 up_22 up_12) +#193 := (not up_20) +#1400 := [hypothesis]: up_10 +#964 := (or #161 #193) +#197 := (or #193 #161) +#966 := (iff #197 #964) +#967 := [rewrite]: #966 +#963 := [asserted]: #197 +#970 := [mp #963 #967]: #964 +#1399 := [unit-resolution #970 #1400]: #193 +#500 := (or up_12 up_20 up_22 up_23) +#51 := (or up_20 up_12) +#52 := (or up_23 #51) +#53 := (or up_22 #52) +#503 := (iff #53 #500) +#491 := (or up_12 up_20) +#494 := (or up_23 #491) +#497 := (or up_22 #494) +#501 := (iff #497 #500) +#502 := [rewrite]: #501 +#498 := (iff #53 #497) +#495 := (iff #52 #494) +#492 := (iff #51 #491) +#493 := [rewrite]: #492 +#496 := [monotonicity #493]: #495 +#499 := [monotonicity #496]: #498 +#504 := [trans #499 #502]: #503 +#490 := [asserted]: #53 +#505 := [mp #490 #504]: #500 +#900 := [unit-resolution #505 #1399 #1456 #1457]: up_23 +#194 := (not up_21) +#974 := (or #161 #194) +#199 := (or #194 #161) +#976 := (iff #199 #974) +#977 := [rewrite]: #976 +#973 := [asserted]: #199 +#980 := [mp #973 #977]: #974 +#902 := [unit-resolution #980 #1400]: #194 +#574 := (or up_21 up_29 up_31 up_32) +#73 := (or up_29 up_21) +#74 := (or up_32 #73) +#75 := (or up_31 #74) +#577 := (iff #75 #574) +#565 := (or up_21 up_29) +#568 := (or up_32 #565) +#571 := (or up_31 #568) +#575 := (iff #571 #574) +#576 := [rewrite]: #575 +#572 := (iff #75 #571) +#569 := (iff #74 #568) +#566 := (iff #73 #565) +#567 := [rewrite]: #566 +#570 := [monotonicity #567]: #569 +#573 := [monotonicity #570]: #572 +#578 := [trans #573 #576]: #577 +#564 := [asserted]: #75 +#579 := [mp #564 #578]: #574 +#851 := [unit-resolution #579 #902 #895 #907]: up_31 +#1135 := (or #202 #234) +#249 := (or #234 #202) +#1137 := (iff #249 #1135) +#1138 := [rewrite]: #1137 +#1134 := [asserted]: #249 +#1141 := [mp #1134 #1138]: #1135 +#858 := [unit-resolution #1141 #851 #900]: false +#853 := [lemma #858]: #847 +#1613 := [unit-resolution #853 #907 #1456 #895 #1457]: #161 +#405 := (or up_7 up_9 up_10) +#22 := (or up_10 up_7) +#23 := (or up_9 #22) +#408 := (iff #23 #405) +#399 := (or up_7 up_10) +#402 := (or up_9 #399) +#406 := (iff #402 #405) +#407 := [rewrite]: #406 +#403 := (iff #23 #402) +#400 := (iff #22 #399) +#401 := [rewrite]: #400 +#404 := [monotonicity #401]: #403 +#409 := [trans #404 #407]: #408 +#398 := [asserted]: #23 +#410 := [mp #398 #409]: #405 +#1614 := [unit-resolution #410 #1613 #1612]: up_9 +#160 := (not up_9) +#881 := (or #160 #165) +#168 := (or #165 #160) +#882 := (iff #168 #881) +#883 := [rewrite]: #882 +#879 := [asserted]: #168 +#886 := [mp #879 #883]: #881 +#1615 := [unit-resolution #886 #1614]: #165 +#425 := (or up_11 up_13) +#29 := (or up_13 up_11) +#426 := (iff #29 #425) +#427 := [rewrite]: #426 +#424 := [asserted]: #29 +#430 := [mp #424 #427]: #425 +#1616 := [unit-resolution #430 #1615]: up_13 +#170 := (not up_13) +#1015 := (or #170 #209) +#211 := (or #209 #170) +#1017 := (iff #211 #1015) +#1018 := [rewrite]: #1017 +#1014 := [asserted]: #211 +#1021 := [mp #1014 #1018]: #1015 +#1617 := [unit-resolution #1021 #1616]: #209 +#603 := (or up_24 up_33 up_35) +#82 := (or up_33 up_24) +#83 := (or up_35 #82) +#606 := (iff #83 #603) +#597 := (or up_24 up_33) +#600 := (or up_35 #597) +#604 := (iff #600 #603) +#605 := [rewrite]: #604 +#601 := (iff #83 #600) +#598 := (iff #82 #597) +#599 := [rewrite]: #598 +#602 := [monotonicity #599]: #601 +#607 := [trans #602 #605]: #606 +#596 := [asserted]: #83 +#608 := [mp #596 #607]: #603 +#1618 := [unit-resolution #608 #1617 #1611]: up_33 +#242 := (not up_33) +#1116 := (or #234 #242) +#245 := (or #242 #234) +#1117 := (iff #245 #1116) +#1118 := [rewrite]: #1117 +#1114 := [asserted]: #245 +#1121 := [mp #1114 #1118]: #1116 +#1619 := [unit-resolution #1121 #1618]: #234 +#1620 := [unit-resolution #579 #1619 #895 #907]: up_21 +#1120 := (or #202 #242) +#246 := (or #242 #202) +#1122 := (iff #246 #1120) +#1123 := [rewrite]: #1122 +#1119 := [asserted]: #246 +#1126 := [mp #1119 #1123]: #1120 +#1621 := [unit-resolution #1126 #1618]: #202 +#1622 := [unit-resolution #505 #1621 #1456 #1457]: up_20 +#195 := (or #193 #194) +#957 := [asserted]: #195 +#1623 := [unit-resolution #957 #1622 #1620]: false +#1625 := [lemma #1623]: #1624 +#2132 := [unit-resolution #1625 #907 #1611 #1457 #895 #1612]: up_22 +#1978 := (or up_32 up_35 up_29 up_21 up_12 up_7) +#1972 := [unit-resolution #1625 #895 #907 #1457 #1611 #1612]: up_22 +#1010 := (or #201 #209) +#210 := (or #209 #201) +#1012 := (iff #210 #1010) +#1013 := [rewrite]: #1012 +#1009 := [asserted]: #210 +#1016 := [mp #1009 #1013]: #1010 +#1973 := [unit-resolution #1016 #1972]: #209 +#1974 := [hypothesis]: #194 +#1975 := [unit-resolution #579 #895 #907 #1974]: up_31 +#1976 := [unit-resolution #1121 #1975]: #242 +#1977 := [unit-resolution #608 #1976 #1973 #1611]: false +#1979 := [lemma #1977]: #1978 +#2133 := [unit-resolution #1979 #907 #1611 #1457 #895 #1612]: up_21 +#1682 := (or #194 up_7 up_12 up_23) +#1673 := [hypothesis]: #202 +#1674 := [hypothesis]: up_21 +#1675 := [unit-resolution #957 #1674]: #193 +#1676 := [unit-resolution #505 #1675 #1457 #1673]: up_22 +#1020 := (or #170 #201) +#212 := (or #201 #170) +#1022 := (iff #212 #1020) +#1023 := [rewrite]: #1022 +#1019 := [asserted]: #212 +#1026 := [mp #1019 #1023]: #1020 +#1677 := [unit-resolution #1026 #1676]: #170 +#1678 := [unit-resolution #980 #1674]: #161 +#1679 := [unit-resolution #410 #1678 #1612]: up_9 +#1680 := [unit-resolution #886 #1679]: #165 +#1681 := [unit-resolution #430 #1680 #1677]: false +#1683 := [lemma #1681]: #1682 +#2134 := [unit-resolution #1683 #2133 #1457 #1612]: up_23 +#203 := (or #201 #202) +#983 := [asserted]: #203 +#2135 := [unit-resolution #983 #2134 #2132]: false +#2137 := [lemma #2135]: #2136 +#2156 := [unit-resolution #2137 #907 #1612 #1611 #1457]: up_32 +#1224 := (or #235 #275) +#279 := (or #275 #235) +#1226 := (iff #279 #1224) +#1227 := [rewrite]: #1226 +#1223 := [asserted]: #279 +#1230 := [mp #1223 #1227]: #1224 +#2157 := [unit-resolution #1230 #2156]: #275 +#2158 := (or up_12 up_29 up_7 up_54) +decl up_26 :: bool +#58 := up_26 +#214 := (not up_26) +decl up_15 :: bool +#31 := up_15 +decl up_14 :: bool +#30 := up_14 +#172 := (not up_14) +decl up_6 :: bool +#13 := up_6 +decl up_5 :: bool +#12 := up_5 +#150 := (not up_5) +decl up_25 :: bool +#57 := up_25 +#2099 := [hypothesis]: up_5 +#859 := (or #150 #155) +#158 := (or #155 #150) +#860 := (iff #158 #859) +#861 := [rewrite]: #860 +#857 := [asserted]: #158 +#864 := [mp #857 #861]: #859 +#2100 := [unit-resolution #864 #2099]: #155 +#863 := (or #150 #156) +#159 := (or #156 #150) +#865 := (iff #159 #863) +#866 := [rewrite]: #865 +#862 := [asserted]: #159 +#869 := [mp #862 #866]: #863 +#2101 := [unit-resolution #869 #2099]: #156 +#2097 := (or up_12 up_7 up_8) +#1626 := [hypothesis]: #156 +#2054 := (or up_54 up_7 up_8) +decl up_16 :: bool +#34 := up_16 +#1597 := [hypothesis]: #317 +#1888 := (or up_16 up_8 up_7 up_54) +decl up_45 :: bool +#104 := up_45 +#284 := (not up_45) +decl up_52 :: bool +#121 := up_52 +#309 := (not up_52) +decl up_51 :: bool +#120 := up_51 +#177 := (not up_16) +#1627 := [hypothesis]: #177 +#1733 := (or up_51 up_7 up_54 up_8 up_16) +decl up_53 :: bool +#125 := up_53 +#308 := (not up_51) +#1598 := [hypothesis]: #308 +decl up_43 :: bool +#99 := up_43 +#276 := (not up_43) +#1710 := (or up_32 up_16 up_8 up_7 up_51 up_54) +#1671 := (or up_35 up_16 up_8 up_32 up_7) +#1655 := (or #166 up_32 up_16 up_8 up_35 up_7) +#1642 := [hypothesis]: up_12 +#885 := (or #160 #166) +#169 := (or #166 #160) +#887 := (iff #169 #885) +#888 := [rewrite]: #887 +#884 := [asserted]: #169 +#891 := [mp #884 #888]: #885 +#1643 := [unit-resolution #891 #1642]: #160 +#1644 := [unit-resolution #410 #1643 #1612]: up_10 +#1645 := [unit-resolution #980 #1644]: #194 +#167 := (or #165 #166) +#878 := [asserted]: #167 +#1646 := [unit-resolution #878 #1642]: #165 +#1647 := [unit-resolution #430 #1646]: up_13 +#1648 := [unit-resolution #1021 #1647]: #209 +#1649 := [unit-resolution #608 #1648 #1611]: up_33 +#1650 := [unit-resolution #1121 #1649]: #234 +decl up_18 :: bool +#39 := up_18 +#185 := (not up_18) +#979 := (or #161 #185) +#200 := (or #185 #161) +#981 := (iff #200 #979) +#982 := [rewrite]: #981 +#978 := [asserted]: #200 +#985 := [mp #978 #982]: #979 +#1651 := [unit-resolution #985 #1644]: #185 +#468 := (or up_8 up_16 up_18 up_19) +#41 := (or up_16 up_8) +#42 := (or up_19 #41) +#43 := (or up_18 #42) +#471 := (iff #43 #468) +#459 := (or up_8 up_16) +#462 := (or up_19 #459) +#465 := (or up_18 #462) +#469 := (iff #465 #468) +#470 := [rewrite]: #469 +#466 := (iff #43 #465) +#463 := (iff #42 #462) +#460 := (iff #41 #459) +#461 := [rewrite]: #460 +#464 := [monotonicity #461]: #463 +#467 := [monotonicity #464]: #466 +#472 := [trans #467 #470]: #471 +#458 := [asserted]: #43 +#473 := [mp #458 #472]: #468 +#1652 := [unit-resolution #473 #1651 #1627 #1626]: up_19 +#186 := (not up_19) +#1068 := (or #186 #226) +#230 := (or #226 #186) +#1070 := (iff #230 #1068) +#1071 := [rewrite]: #1070 +#1067 := [asserted]: #230 +#1074 := [mp #1067 #1071]: #1068 +#1653 := [unit-resolution #1074 #1652]: #226 +#1654 := [unit-resolution #579 #1653 #1650 #895 #1645]: false +#1656 := [lemma #1654]: #1655 +#1657 := [unit-resolution #1656 #1611 #1627 #1626 #895 #1612]: #166 +#1640 := (or up_12 up_35 up_7 up_22 up_16 up_8 up_32) +#1628 := [unit-resolution #1625 #1457 #1456 #895 #1611 #1612]: up_29 +#1629 := [unit-resolution #1074 #1628]: #186 +#1630 := [unit-resolution #473 #1629 #1627 #1626]: up_18 +#960 := (or #185 #193) +#196 := (or #193 #185) +#961 := (iff #196 #960) +#962 := [rewrite]: #961 +#958 := [asserted]: #196 +#965 := [mp #958 #962]: #960 +#1631 := [unit-resolution #965 #1630]: #193 +#1632 := [unit-resolution #505 #1631 #1456 #1457]: up_23 +#1633 := [unit-resolution #1126 #1632]: #242 +#1634 := [unit-resolution #608 #1633 #1611]: up_24 +#1635 := [unit-resolution #985 #1630]: #161 +#1636 := [unit-resolution #410 #1635 #1612]: up_9 +#1637 := [unit-resolution #886 #1636]: #165 +#1638 := [unit-resolution #430 #1637]: up_13 +#1639 := [unit-resolution #1021 #1638 #1634]: false +#1641 := [lemma #1639]: #1640 +#1658 := [unit-resolution #1641 #1657 #1612 #1611 #1627 #1626 #895]: up_22 +#1659 := [unit-resolution #1016 #1658]: #209 +#1660 := [unit-resolution #608 #1659 #1611]: up_33 +#1661 := [unit-resolution #1121 #1660]: #234 +#1662 := [unit-resolution #1026 #1658]: #170 +#1663 := [unit-resolution #430 #1662]: up_11 +#1664 := [unit-resolution #886 #1663]: #160 +#1665 := [unit-resolution #410 #1664 #1612]: up_10 +#1666 := [unit-resolution #980 #1665]: #194 +#1667 := [unit-resolution #579 #1666 #895 #1661]: up_29 +#1668 := [unit-resolution #985 #1665]: #185 +#1669 := [unit-resolution #473 #1668 #1627 #1626]: up_19 +#1670 := [unit-resolution #1074 #1669 #1667]: false +#1672 := [lemma #1670]: #1671 +#1698 := [unit-resolution #1672 #895 #1626 #1627 #1612]: up_35 +#1609 := (or #250 up_34 up_51 up_54) +#316 := (not up_53) +#1599 := [hypothesis]: up_35 +#1275 := (or #250 #291) +#293 := (or #291 #250) +#1277 := (iff #293 #1275) +#1278 := [rewrite]: #1277 +#1274 := [asserted]: #293 +#1281 := [mp #1274 #1278]: #1275 +#1600 := [unit-resolution #1281 #1599]: #291 +#777 := (or up_46 up_55) +decl up_4 :: bool +#10 := up_4 +#783 := (or up_4 up_46 up_55) +#1514 := (iff #783 #777) +#1509 := (or false up_46 up_55) +#1512 := (iff #1509 #777) +#1513 := [rewrite]: #1512 +#1510 := (iff #783 #1509) +#1485 := (iff up_4 false) +#11 := (not up_4) +#1488 := (iff #11 #1485) +#1486 := (iff #1485 #11) +#1487 := [rewrite]: #1486 +#1489 := [symm #1487]: #1488 +#371 := [asserted]: #11 +#1490 := [mp #371 #1489]: #1485 +#1511 := [monotonicity #1490]: #1510 +#1515 := [trans #1511 #1513]: #1514 +#135 := (or up_55 up_46) +#136 := (or up_4 #135) +#786 := (iff #136 #783) +#780 := (or up_4 #777) +#784 := (iff #780 #783) +#785 := [rewrite]: #784 +#781 := (iff #136 #780) +#778 := (iff #135 #777) +#779 := [rewrite]: #778 +#782 := [monotonicity #779]: #781 +#787 := [trans #782 #785]: #786 +#776 := [asserted]: #136 +#788 := [mp #776 #787]: #783 +#1516 := [mp #788 #1515]: #777 +#1601 := [unit-resolution #1516 #1600]: up_55 +#324 := (not up_55) +#1376 := (or #316 #324) +#327 := (or #324 #316) +#1377 := (iff #327 #1376) +#1378 := [rewrite]: #1377 +#1374 := [asserted]: #327 +#1381 := [mp #1374 #1378]: #1376 +#1602 := [unit-resolution #1381 #1601]: #316 +#754 := (or up_43 up_51 up_53 up_54) +#127 := (or up_51 up_43) +#128 := (or up_54 #127) +#129 := (or up_53 #128) +#757 := (iff #129 #754) +#745 := (or up_43 up_51) +#748 := (or up_54 #745) +#751 := (or up_53 #748) +#755 := (iff #751 #754) +#756 := [rewrite]: #755 +#752 := (iff #129 #751) +#749 := (iff #128 #748) +#746 := (iff #127 #745) +#747 := [rewrite]: #746 +#750 := [monotonicity #747]: #749 +#753 := [monotonicity #750]: #752 +#758 := [trans #753 #756]: #757 +#744 := [asserted]: #129 +#759 := [mp #744 #758]: #754 +#1603 := [unit-resolution #759 #1602 #1598 #1597]: up_43 +decl up_44 :: bool +#103 := up_44 +#283 := (not up_44) +#1280 := (or #250 #283) +#294 := (or #283 #250) +#1282 := (iff #294 #1280) +#1283 := [rewrite]: #1282 +#1279 := [asserted]: #294 +#1286 := [mp #1279 #1283]: #1280 +#1604 := [unit-resolution #1286 #1599]: #283 +#1380 := (or #284 #324) +#328 := (or #324 #284) +#1382 := (iff #328 #1380) +#1383 := [rewrite]: #1382 +#1379 := [asserted]: #328 +#1386 := [mp #1379 #1383]: #1380 +#1606 := [unit-resolution #1386 #1601]: #284 +#680 := (or up_34 up_42 up_44 up_45) +#105 := (or up_42 up_34) +#106 := (or up_45 #105) +#107 := (or up_44 #106) +#683 := (iff #107 #680) +#671 := (or up_34 up_42) +#674 := (or up_45 #671) +#677 := (or up_44 #674) +#681 := (iff #677 #680) +#682 := [rewrite]: #681 +#678 := (iff #107 #677) +#675 := (iff #106 #674) +#672 := (iff #105 #671) +#673 := [rewrite]: #672 +#676 := [monotonicity #673]: #675 +#679 := [monotonicity #676]: #678 +#684 := [trans #679 #682]: #683 +#670 := [asserted]: #107 +#685 := [mp #670 #684]: #680 +#1607 := [unit-resolution #685 #1606 #1605 #1604]: up_42 +#277 := (or #275 #276) +#1217 := [asserted]: #277 +#1608 := [unit-resolution #1217 #1607 #1603]: false +#1610 := [lemma #1608]: #1609 +#1699 := [unit-resolution #1610 #1698 #1598 #1597]: up_34 +#1125 := (or #234 #243) +#247 := (or #243 #234) +#1127 := (iff #247 #1125) +#1128 := [rewrite]: #1127 +#1124 := [asserted]: #247 +#1131 := [mp #1124 #1128]: #1125 +#1700 := [unit-resolution #1131 #1699]: #234 +#1130 := (or #202 #243) +#248 := (or #243 #202) +#1132 := (iff #248 #1130) +#1133 := [rewrite]: #1132 +#1129 := [asserted]: #248 +#1136 := [mp #1129 #1133]: #1130 +#1701 := [unit-resolution #1136 #1699]: #202 +#1696 := (or up_12 up_7 up_23 up_16 up_8 up_32 up_31) +#1684 := [hypothesis]: #234 +#1685 := [unit-resolution #1683 #1457 #1612 #1673]: #194 +#1686 := [unit-resolution #579 #1685 #895 #1684]: up_29 +#1687 := [unit-resolution #1074 #1686]: #186 +#1688 := [unit-resolution #473 #1687 #1627 #1626]: up_18 +#1689 := [unit-resolution #965 #1688]: #193 +#1690 := [unit-resolution #505 #1689 #1457 #1673]: up_22 +#1691 := [unit-resolution #1026 #1690]: #170 +#1692 := [unit-resolution #985 #1688]: #161 +#1693 := [unit-resolution #410 #1692 #1612]: up_9 +#1694 := [unit-resolution #886 #1693]: #165 +#1695 := [unit-resolution #430 #1694 #1691]: false +#1697 := [lemma #1695]: #1696 +#1702 := [unit-resolution #1697 #1701 #1612 #1627 #1626 #895 #1700]: up_12 +#1703 := [unit-resolution #891 #1702]: #160 +#1704 := [unit-resolution #410 #1703 #1612]: up_10 +#1705 := [unit-resolution #980 #1704]: #194 +#1706 := [unit-resolution #579 #1705 #895 #1700]: up_29 +#1707 := [unit-resolution #985 #1704]: #185 +#1708 := [unit-resolution #473 #1707 #1627 #1626]: up_19 +#1709 := [unit-resolution #1074 #1708 #1706]: false +#1711 := [lemma #1709]: #1710 +#1712 := [unit-resolution #1711 #1598 #1626 #1612 #1627 #1597]: up_32 +#1234 := (or #235 #276) +#281 := (or #276 #235) +#1236 := (iff #281 #1234) +#1237 := [rewrite]: #1236 +#1233 := [asserted]: #281 +#1240 := [mp #1233 #1237]: #1234 +#1713 := [unit-resolution #1240 #1712]: #276 +#1714 := [unit-resolution #759 #1713 #1598 #1597]: up_53 +#1395 := (or #284 #316) +#331 := (or #316 #284) +#1397 := (iff #331 #1395) +#1398 := [rewrite]: #1397 +#1394 := [asserted]: #331 +#1401 := [mp #1394 #1398]: #1395 +#1715 := [unit-resolution #1401 #1714]: #284 +#1716 := [unit-resolution #1230 #1712]: #275 +#1717 := [unit-resolution #1381 #1714]: #324 +#1718 := [unit-resolution #1516 #1717]: up_46 +#1270 := (or #283 #291) +#292 := (or #291 #283) +#1272 := (iff #292 #1270) +#1273 := [rewrite]: #1272 +#1269 := [asserted]: #292 +#1276 := [mp #1269 #1273]: #1270 +#1719 := [unit-resolution #1276 #1718]: #283 +#1720 := [unit-resolution #685 #1719 #1716 #1715]: up_34 +#1721 := [unit-resolution #1136 #1720]: #202 +#1722 := [unit-resolution #1281 #1718]: #250 +#244 := (or #242 #243) +#1113 := [asserted]: #244 +#1723 := [unit-resolution #1113 #1720]: #242 +#1724 := [unit-resolution #608 #1723 #1722]: up_24 +#1725 := [unit-resolution #1016 #1724]: #201 +#1726 := [unit-resolution #1021 #1724]: #170 +#1727 := [unit-resolution #430 #1726]: up_11 +#1728 := [unit-resolution #878 #1727]: #166 +#1729 := [unit-resolution #505 #1728 #1725 #1721]: up_20 +#1730 := [unit-resolution #886 #1727]: #160 +#1731 := [unit-resolution #410 #1730 #1612]: up_10 +#1732 := [unit-resolution #970 #1731 #1729]: false +#1734 := [lemma #1732]: #1733 +#1858 := [unit-resolution #1734 #1627 #1597 #1626 #1612]: up_51 +#310 := (or #308 #309) +#1321 := [asserted]: #310 +#1859 := [unit-resolution #1321 #1858]: #309 +decl up_58 :: bool +#139 := up_58 +#337 := (not up_58) +decl up_49 :: bool +#115 := up_49 +#300 := (not up_49) +#1324 := (or #300 #308) +#311 := (or #308 #300) +#1325 := (iff #311 #1324) +#1326 := [rewrite]: #1325 +#1322 := [asserted]: #311 +#1329 := [mp #1322 #1326]: #1324 +#1860 := [unit-resolution #1329 #1858]: #300 +decl up_39 :: bool +#89 := up_39 +#260 := (not up_39) +decl up_38 :: bool +#88 := up_38 +decl up_40 :: bool +#93 := up_40 +#267 := (not up_40) +decl up_41 :: bool +#94 := up_41 +#268 := (not up_41) +#1328 := (or #268 #308) +#312 := (or #308 #268) +#1330 := (iff #312 #1328) +#1331 := [rewrite]: #1330 +#1327 := [asserted]: #312 +#1334 := [mp #1327 #1331]: #1328 +#1861 := [unit-resolution #1334 #1858]: #268 +#1771 := (or up_32 up_16 up_8 up_41 up_49 up_52 up_7) +#1735 := [unit-resolution #1281 #1698]: #291 +#1736 := [unit-resolution #1516 #1735]: up_55 +#1737 := [unit-resolution #1386 #1736]: #284 +#1738 := [unit-resolution #1286 #1698]: #283 +#259 := (not up_38) +decl up_50 :: bool +#116 := up_50 +#301 := (not up_50) +#1739 := [hypothesis]: #309 +#341 := (not up_59) +#326 := (or #324 #325) +#1373 := [asserted]: #326 +#1740 := [unit-resolution #1373 #1736]: #325 +#834 := (or up_56 up_60) +decl up_3 :: bool +#8 := up_3 +#840 := (or up_3 up_56 up_60) +#1522 := (iff #840 #834) +#1517 := (or false up_56 up_60) +#1520 := (iff #1517 #834) +#1521 := [rewrite]: #1520 +#1518 := (iff #840 #1517) +#1479 := (iff up_3 false) +#9 := (not up_3) +#1482 := (iff #9 #1479) +#1480 := (iff #1479 #9) +#1481 := [rewrite]: #1480 +#1483 := [symm #1481]: #1482 +#370 := [asserted]: #9 +#1484 := [mp #370 #1483]: #1479 +#1519 := [monotonicity #1484]: #1518 +#1523 := [trans #1519 #1521]: #1522 +#148 := (or up_60 up_56) +#149 := (or up_3 #148) +#843 := (iff #149 #840) +#837 := (or up_3 #834) +#841 := (iff #837 #840) +#842 := [rewrite]: #841 +#838 := (iff #149 #837) +#835 := (iff #148 #834) +#836 := [rewrite]: #835 +#839 := [monotonicity #836]: #838 +#844 := [trans #839 #842]: #843 +#833 := [asserted]: #149 +#845 := [mp #833 #844]: #840 +#1524 := [mp #845 #1523]: #834 +#1741 := [unit-resolution #1524 #1740]: up_60 +#1442 := (or #341 #345) +#346 := (or #345 #341) +#1444 := (iff #346 #1442) +#1445 := [rewrite]: #1444 +#1441 := [asserted]: #346 +#1448 := [mp #1441 #1445]: #1442 +#1742 := [unit-resolution #1448 #1741]: #341 +#814 := (or up_52 up_58 up_59) +#143 := (or up_58 up_52) +#144 := (or up_59 #143) +#817 := (iff #144 #814) +#808 := (or up_52 up_58) +#811 := (or up_59 #808) +#815 := (iff #811 #814) +#816 := [rewrite]: #815 +#812 := (iff #144 #811) +#809 := (iff #143 #808) +#810 := [rewrite]: #809 +#813 := [monotonicity #810]: #812 +#818 := [trans #813 #816]: #817 +#807 := [asserted]: #144 +#819 := [mp #807 #818]: #814 +#1743 := [unit-resolution #819 #1742 #1739]: up_58 +#1417 := (or #301 #337) +#339 := (or #337 #301) +#1419 := (iff #339 #1417) +#1420 := [rewrite]: #1419 +#1416 := [asserted]: #339 +#1423 := [mp #1416 #1420]: #1417 +#1744 := [unit-resolution #1423 #1743]: #301 +#1745 := [hypothesis]: #300 +decl up_47 :: bool +#111 := up_47 +#295 := (not up_47) +decl up_48 :: bool +#112 := up_48 +decl up_57 :: bool +#137 := up_57 +#335 := (not up_57) +#1412 := (or #335 #337) +#338 := (or #337 #335) +#1414 := (iff #338 #1412) +#1415 := [rewrite]: #1414 +#1411 := [asserted]: #338 +#1418 := [mp #1411 #1415]: #1412 +#1746 := [unit-resolution #1418 #1743]: #335 +#790 := (or up_48 up_57) +#138 := (or up_57 up_48) +#791 := (iff #138 #790) +#792 := [rewrite]: #791 +#789 := [asserted]: #138 +#795 := [mp #789 #792]: #790 +#1747 := [unit-resolution #795 #1746]: up_48 +#296 := (not up_48) +#297 := (or #295 #296) +#1284 := [asserted]: #297 +#1748 := [unit-resolution #1284 #1747]: #295 +#722 := (or up_39 up_47 up_49 up_50) +#117 := (or up_47 up_39) +#118 := (or up_50 #117) +#119 := (or up_49 #118) +#725 := (iff #119 #722) +#713 := (or up_39 up_47) +#716 := (or up_50 #713) +#719 := (or up_49 #716) +#723 := (iff #719 #722) +#724 := [rewrite]: #723 +#720 := (iff #119 #719) +#717 := (iff #118 #716) +#714 := (iff #117 #713) +#715 := [rewrite]: #714 +#718 := [monotonicity #715]: #717 +#721 := [monotonicity #718]: #720 +#726 := [trans #721 #724]: #725 +#712 := [asserted]: #119 +#727 := [mp #712 #726]: #722 +#1749 := [unit-resolution #727 #1748 #1745 #1744]: up_39 +#261 := (or #259 #260) +#1165 := [asserted]: #261 +#1750 := [unit-resolution #1165 #1749]: #259 +#1751 := [hypothesis]: #268 +decl up_30 :: bool +#67 := up_30 +#227 := (not up_30) +decl up_27 :: bool +#61 := up_27 +#213 := (not up_25) +decl up_37 :: bool +#85 := up_37 +#255 := (not up_37) +#1291 := (or #255 #296) +#299 := (or #296 #255) +#1293 := (iff #299 #1291) +#1294 := [rewrite]: #1293 +#1290 := [asserted]: #299 +#1297 := [mp #1290 #1294]: #1291 +#1752 := [unit-resolution #1297 #1747]: #255 +decl up_36 :: bool +#84 := up_36 +#254 := (not up_36) +#1177 := (or #254 #260) +#264 := (or #260 #254) +#1179 := (iff #264 #1177) +#1180 := [rewrite]: #1179 +#1176 := [asserted]: #264 +#1183 := [mp #1176 #1180]: #1177 +#1753 := [unit-resolution #1183 #1749]: #254 +#616 := (or up_26 up_36 up_37) +#86 := (or up_37 up_26) +#87 := (or up_36 #86) +#619 := (iff #87 #616) +#610 := (or up_26 up_37) +#613 := (or up_36 #610) +#617 := (iff #613 #616) +#618 := [rewrite]: #617 +#614 := (iff #87 #613) +#611 := (iff #86 #610) +#612 := [rewrite]: #611 +#615 := [monotonicity #612]: #614 +#620 := [trans #615 #618]: #619 +#609 := [asserted]: #87 +#621 := [mp #609 #620]: #616 +#1754 := [unit-resolution #621 #1753 #1752]: up_26 +#215 := (or #213 #214) +#1024 := [asserted]: #215 +#1755 := [unit-resolution #1024 #1754]: #213 +decl up_28 :: bool +#62 := up_28 +#219 := (not up_28) +#1182 := (or #219 #260) +#265 := (or #260 #219) +#1184 := (iff #265 #1182) +#1185 := [rewrite]: #1184 +#1181 := [asserted]: #265 +#1188 := [mp #1181 #1185]: #1182 +#1756 := [unit-resolution #1188 #1749]: #219 +decl up_17 :: bool +#35 := up_17 +#178 := (not up_17) +#173 := (not up_15) +#1031 := (or #173 #214) +#217 := (or #214 #173) +#1033 := (iff #217 #1031) +#1034 := [rewrite]: #1033 +#1030 := [asserted]: #217 +#1037 := [mp #1030 #1034]: #1031 +#1757 := [unit-resolution #1037 #1754]: #173 +#1503 := (or up_14 up_15) +decl up_2 :: bool +#6 := up_2 +#436 := (or up_2 up_14 up_15) +#1506 := (iff #436 #1503) +#1500 := (or false up_14 up_15) +#1504 := (iff #1500 #1503) +#1505 := [rewrite]: #1504 +#1501 := (iff #436 #1500) +#1473 := (iff up_2 false) +#7 := (not up_2) +#1476 := (iff #7 #1473) +#1474 := (iff #1473 #7) +#1475 := [rewrite]: #1474 +#1477 := [symm #1475]: #1476 +#369 := [asserted]: #7 +#1478 := [mp #369 #1477]: #1473 +#1502 := [monotonicity #1478]: #1501 +#1507 := [trans #1502 #1505]: #1506 +#32 := (or up_15 up_2) +#33 := (or up_14 #32) +#439 := (iff #33 #436) +#429 := (or up_2 up_15) +#433 := (or up_14 #429) +#437 := (iff #433 #436) +#438 := [rewrite]: #437 +#434 := (iff #33 #433) +#431 := (iff #32 #429) +#432 := [rewrite]: #431 +#435 := [monotonicity #432]: #434 +#440 := [trans #435 #438]: #439 +#428 := [asserted]: #33 +#441 := [mp #428 #440]: #436 +#1508 := [mp #441 #1507]: #1503 +#1758 := [unit-resolution #1508 #1757]: up_14 +#917 := (or #172 #178) +#182 := (or #178 #172) +#919 := (iff #182 #917) +#920 := [rewrite]: #919 +#916 := [asserted]: #182 +#923 := [mp #916 #920]: #917 +#1759 := [unit-resolution #923 #1758]: #178 +#542 := (or up_17 up_25 up_27 up_28) +#63 := (or up_25 up_17) +#64 := (or up_28 #63) +#65 := (or up_27 #64) +#545 := (iff #65 #542) +#533 := (or up_17 up_25) +#536 := (or up_28 #533) +#539 := (or up_27 #536) +#543 := (iff #539 #542) +#544 := [rewrite]: #543 +#540 := (iff #65 #539) +#537 := (iff #64 #536) +#534 := (iff #63 #533) +#535 := [rewrite]: #534 +#538 := [monotonicity #535]: #537 +#541 := [monotonicity #538]: #540 +#546 := [trans #541 #544]: #545 +#532 := [asserted]: #65 +#547 := [mp #532 #546]: #542 +#1760 := [unit-resolution #547 #1759 #1756 #1755]: up_27 +#218 := (not up_27) +#1073 := (or #218 #227) +#231 := (or #227 #218) +#1075 := (iff #231 #1073) +#1076 := [rewrite]: #1075 +#1072 := [asserted]: #231 +#1079 := [mp #1072 #1076]: #1073 +#1761 := [unit-resolution #1079 #1760]: #227 +#648 := (or up_30 up_38 up_40 up_41) +#95 := (or up_38 up_30) +#96 := (or up_41 #95) +#97 := (or up_40 #96) +#651 := (iff #97 #648) +#639 := (or up_30 up_38) +#642 := (or up_41 #639) +#645 := (or up_40 #642) +#649 := (iff #645 #648) +#650 := [rewrite]: #649 +#646 := (iff #97 #645) +#643 := (iff #96 #642) +#640 := (iff #95 #639) +#641 := [rewrite]: #640 +#644 := [monotonicity #641]: #643 +#647 := [monotonicity #644]: #646 +#652 := [trans #647 #650]: #651 +#638 := [asserted]: #97 +#653 := [mp #638 #652]: #648 +#1762 := [unit-resolution #653 #1761 #1751 #1750]: up_40 +#1220 := (or #267 #275) +#278 := (or #275 #267) +#1221 := (iff #278 #1220) +#1222 := [rewrite]: #1221 +#1218 := [asserted]: #278 +#1225 := [mp #1218 #1222]: #1220 +#1763 := [unit-resolution #1225 #1762]: #275 +#1764 := [unit-resolution #685 #1763 #1738 #1737]: up_34 +#1064 := (or #218 #226) +#229 := (or #226 #218) +#1065 := (iff #229 #1064) +#1066 := [rewrite]: #1065 +#1062 := [asserted]: #229 +#1069 := [mp #1062 #1066]: #1064 +#1765 := [unit-resolution #1069 #1760]: #226 +#1083 := (or #186 #218) +#233 := (or #218 #186) +#1085 := (iff #233 #1083) +#1086 := [rewrite]: #1085 +#1082 := [asserted]: #233 +#1089 := [mp #1082 #1086]: #1083 +#1766 := [unit-resolution #1089 #1760]: #186 +#1767 := [unit-resolution #473 #1766 #1627 #1626]: up_18 +#969 := (or #185 #194) +#198 := (or #194 #185) +#971 := (iff #198 #969) +#972 := [rewrite]: #971 +#968 := [asserted]: #198 +#975 := [mp #968 #972]: #969 +#1768 := [unit-resolution #975 #1767]: #194 +#1769 := [unit-resolution #579 #1768 #895 #1765]: up_31 +#1770 := [unit-resolution #1131 #1769 #1764]: false +#1772 := [lemma #1770]: #1771 +#1862 := [unit-resolution #1772 #1627 #1626 #1861 #1860 #1859 #1612]: up_32 +#1239 := (or #235 #267) +#282 := (or #267 #235) +#1241 := (iff #282 #1239) +#1242 := [rewrite]: #1241 +#1238 := [asserted]: #282 +#1245 := [mp #1238 #1242]: #1239 +#1863 := [unit-resolution #1245 #1862]: #267 +#1856 := (or up_12 up_52 up_49 up_41 up_16 up_8 up_7) +#1828 := [unit-resolution #1772 #1627 #1626 #1751 #1745 #1739 #1612]: up_32 +#1829 := [unit-resolution #1245 #1828]: #267 +#1830 := [unit-resolution #1230 #1828]: #275 +#1826 := (or #170 up_41 up_40 up_16 up_8 up_49 up_12 up_52 up_42) +#1804 := [hypothesis]: up_13 +#1805 := [unit-resolution #1026 #1804]: #201 +#1806 := [unit-resolution #1021 #1804]: #209 +#1798 := [hypothesis]: #275 +#1782 := [hypothesis]: #267 +#1802 := (or #242 up_42 up_52 up_49 up_41 up_40 up_16 up_8 up_12 up_22) +#1783 := [hypothesis]: up_33 +#1784 := [unit-resolution #1126 #1783]: #202 +#1785 := [unit-resolution #505 #1784 #1457 #1456]: up_20 +#1786 := [unit-resolution #965 #1785]: #185 +#1787 := [unit-resolution #473 #1786 #1627 #1626]: up_19 +#1078 := (or #186 #227) +#232 := (or #227 #186) +#1080 := (iff #232 #1078) +#1081 := [rewrite]: #1080 +#1077 := [asserted]: #232 +#1084 := [mp #1077 #1081]: #1078 +#1788 := [unit-resolution #1084 #1787]: #227 +#1789 := [unit-resolution #653 #1788 #1751 #1782]: up_38 +#1790 := [unit-resolution #1165 #1789]: #260 +#1780 := (or #337 up_49 up_39) +#1773 := [hypothesis]: up_58 +#1774 := [unit-resolution #1418 #1773]: #335 +#1775 := [unit-resolution #795 #1774]: up_48 +#1776 := [hypothesis]: #260 +#1777 := [unit-resolution #1423 #1773]: #301 +#1778 := [unit-resolution #727 #1777 #1745 #1776]: up_47 +#1779 := [unit-resolution #1284 #1778 #1775]: false +#1781 := [lemma #1779]: #1780 +#1791 := [unit-resolution #1781 #1790 #1745]: #337 +#1792 := [unit-resolution #819 #1791 #1739]: up_59 +#1793 := [unit-resolution #1448 #1792]: #345 +#1794 := [unit-resolution #1524 #1793]: up_56 +#1795 := [unit-resolution #1373 #1794]: #324 +#1796 := [unit-resolution #1516 #1795]: up_46 +#1797 := [unit-resolution #1113 #1783]: #243 +#1390 := (or #284 #325) +#330 := (or #325 #284) +#1392 := (iff #330 #1390) +#1393 := [rewrite]: #1392 +#1389 := [asserted]: #330 +#1396 := [mp #1389 #1393]: #1390 +#1799 := [unit-resolution #1396 #1794]: #284 +#1800 := [unit-resolution #685 #1799 #1798 #1797]: up_44 +#1801 := [unit-resolution #1276 #1800 #1796]: false +#1803 := [lemma #1801]: #1802 +#1807 := [unit-resolution #1803 #1805 #1739 #1745 #1751 #1782 #1627 #1626 #1457 #1798]: #242 +#1808 := [unit-resolution #608 #1807 #1806]: up_35 +#1809 := [unit-resolution #1286 #1808]: #283 +#1810 := [unit-resolution #1281 #1808]: #291 +#1811 := [unit-resolution #1516 #1810]: up_55 +#1812 := [unit-resolution #1386 #1811]: #284 +#1813 := [unit-resolution #685 #1812 #1798 #1809]: up_34 +#1814 := [unit-resolution #1136 #1813]: #202 +#1815 := [unit-resolution #505 #1814 #1457 #1805]: up_20 +#1816 := [unit-resolution #965 #1815]: #185 +#1817 := [unit-resolution #473 #1816 #1627 #1626]: up_19 +#1818 := [unit-resolution #1373 #1811]: #325 +#1819 := [unit-resolution #1524 #1818]: up_60 +#1820 := [unit-resolution #1448 #1819]: #341 +#1821 := [unit-resolution #819 #1820 #1739]: up_58 +#1822 := [unit-resolution #1781 #1821 #1745]: up_39 +#1823 := [unit-resolution #1165 #1822]: #259 +#1824 := [unit-resolution #653 #1823 #1751 #1782]: up_30 +#1825 := [unit-resolution #1084 #1824 #1817]: false +#1827 := [lemma #1825]: #1826 +#1831 := [unit-resolution #1827 #1457 #1829 #1627 #1626 #1745 #1751 #1739 #1830]: #170 +#1832 := [unit-resolution #430 #1831]: up_11 +#1833 := [unit-resolution #886 #1832]: #160 +#1834 := [unit-resolution #410 #1833 #1612]: up_10 +#1835 := [unit-resolution #985 #1834]: #185 +#1836 := [unit-resolution #473 #1835 #1627 #1626]: up_19 +#1837 := [unit-resolution #1084 #1836]: #227 +#1838 := [unit-resolution #653 #1837 #1751 #1829]: up_38 +#1839 := [unit-resolution #1165 #1838]: #260 +#1840 := [unit-resolution #1781 #1839 #1745]: #337 +#1841 := [unit-resolution #819 #1840 #1739]: up_59 +#1842 := [unit-resolution #1448 #1841]: #345 +#1843 := [unit-resolution #1524 #1842]: up_56 +#1844 := [unit-resolution #1373 #1843]: #324 +#1845 := [unit-resolution #1516 #1844]: up_46 +#1846 := [unit-resolution #1281 #1845]: #250 +#1847 := [unit-resolution #1396 #1843]: #284 +#1848 := [unit-resolution #1276 #1845]: #283 +#1849 := [unit-resolution #685 #1848 #1830 #1847]: up_34 +#1850 := [unit-resolution #1113 #1849]: #242 +#1851 := [unit-resolution #608 #1850 #1846]: up_24 +#1852 := [unit-resolution #970 #1834]: #193 +#1853 := [unit-resolution #1136 #1849]: #202 +#1854 := [unit-resolution #505 #1853 #1457 #1852]: up_22 +#1855 := [unit-resolution #1016 #1854 #1851]: false +#1857 := [lemma #1855]: #1856 +#1864 := [unit-resolution #1857 #1859 #1860 #1861 #1627 #1626 #1612]: up_12 +#1865 := [unit-resolution #891 #1864]: #160 +#1866 := [unit-resolution #410 #1865 #1612]: up_10 +#1867 := [unit-resolution #985 #1866]: #185 +#1868 := [unit-resolution #473 #1867 #1627 #1626]: up_19 +#1869 := [unit-resolution #1084 #1868]: #227 +#1870 := [unit-resolution #653 #1869 #1861 #1863]: up_38 +#1871 := [unit-resolution #1165 #1870]: #260 +#1872 := [unit-resolution #1781 #1871 #1860]: #337 +#1873 := [unit-resolution #819 #1872 #1859]: up_59 +#1874 := [unit-resolution #1448 #1873]: #345 +#1875 := [unit-resolution #1524 #1874]: up_56 +#1876 := [unit-resolution #1396 #1875]: #284 +#1877 := [unit-resolution #1230 #1862]: #275 +#1878 := [unit-resolution #1373 #1875]: #324 +#1879 := [unit-resolution #1516 #1878]: up_46 +#1880 := [unit-resolution #1276 #1879]: #283 +#1881 := [unit-resolution #685 #1880 #1877 #1876]: up_34 +#1882 := [unit-resolution #878 #1864]: #165 +#1883 := [unit-resolution #430 #1882]: up_13 +#1884 := [unit-resolution #1021 #1883]: #209 +#1885 := [unit-resolution #1281 #1879]: #250 +#1886 := [unit-resolution #608 #1885 #1884]: up_33 +#1887 := [unit-resolution #1113 #1886 #1881]: false +#1889 := [lemma #1887]: #1888 +#2026 := [unit-resolution #1889 #1597 #1612 #1626]: up_16 +#908 := (or #172 #177) +#180 := (or #177 #172) +#909 := (iff #180 #908) +#910 := [rewrite]: #909 +#906 := [asserted]: #180 +#913 := [mp #906 #910]: #908 +#2027 := [unit-resolution #913 #2026]: #172 +#2028 := [unit-resolution #1508 #2027]: up_15 +#2029 := [unit-resolution #1037 #2028]: #214 +#1027 := (or #173 #213) +#216 := (or #213 #173) +#1028 := (iff #216 #1027) +#1029 := [rewrite]: #1028 +#1025 := [asserted]: #216 +#1032 := [mp #1025 #1029]: #1027 +#2030 := [unit-resolution #1032 #2028]: #213 +#179 := (or #177 #178) +#905 := [asserted]: #179 +#2031 := [unit-resolution #905 #2026]: #178 +#1917 := (or #226 up_54 up_26 up_17 up_25) +#1890 := [hypothesis]: #214 +#1891 := [hypothesis]: #213 +#1892 := [hypothesis]: #178 +#1893 := [hypothesis]: up_29 +#1894 := [unit-resolution #1069 #1893]: #218 +#1895 := [unit-resolution #547 #1894 #1892 #1891]: up_28 +#1187 := (or #219 #254) +#266 := (or #254 #219) +#1189 := (iff #266 #1187) +#1190 := [rewrite]: #1189 +#1186 := [asserted]: #266 +#1193 := [mp #1186 #1190]: #1187 +#1896 := [unit-resolution #1193 #1895]: #254 +#1897 := [unit-resolution #621 #1896 #1890]: up_37 +#1898 := [unit-resolution #1297 #1897]: #296 +#1899 := [unit-resolution #795 #1898]: up_57 +#1900 := [unit-resolution #1418 #1899]: #337 +#1901 := [unit-resolution #1188 #1895]: #260 +#1287 := (or #255 #295) +#298 := (or #295 #255) +#1288 := (iff #298 #1287) +#1289 := [rewrite]: #1288 +#1285 := [asserted]: #298 +#1292 := [mp #1285 #1289]: #1287 +#1902 := [unit-resolution #1292 #1897]: #295 +#1422 := (or #301 #335) +#340 := (or #335 #301) +#1424 := (iff #340 #1422) +#1425 := [rewrite]: #1424 +#1421 := [asserted]: #340 +#1428 := [mp #1421 #1425]: #1422 +#1903 := [unit-resolution #1428 #1899]: #301 +#1904 := [unit-resolution #727 #1903 #1902 #1901]: up_49 +#1333 := (or #300 #309) +#313 := (or #309 #300) +#1335 := (iff #313 #1333) +#1336 := [rewrite]: #1335 +#1332 := [asserted]: #313 +#1339 := [mp #1332 #1336]: #1333 +#1905 := [unit-resolution #1339 #1904]: #309 +#1906 := [unit-resolution #819 #1905 #1900]: up_59 +#1907 := [unit-resolution #1448 #1906]: #345 +#1908 := [unit-resolution #1524 #1907]: up_56 +#1909 := [unit-resolution #1329 #1904]: #308 +#1172 := (or #219 #259) +#263 := (or #259 #219) +#1174 := (iff #263 #1172) +#1175 := [rewrite]: #1174 +#1171 := [asserted]: #263 +#1178 := [mp #1171 #1175]: #1172 +#1910 := [unit-resolution #1178 #1895]: #259 +#228 := (or #226 #227) +#1061 := [asserted]: #228 +#1911 := [unit-resolution #1061 #1893]: #227 +#1343 := (or #268 #300) +#315 := (or #300 #268) +#1345 := (iff #315 #1343) +#1346 := [rewrite]: #1345 +#1342 := [asserted]: #315 +#1349 := [mp #1342 #1346]: #1343 +#1912 := [unit-resolution #1349 #1904]: #268 +#1913 := [unit-resolution #653 #1912 #1911 #1910]: up_40 +#1229 := (or #267 #276) +#280 := (or #276 #267) +#1231 := (iff #280 #1229) +#1232 := [rewrite]: #1231 +#1228 := [asserted]: #280 +#1235 := [mp #1228 #1232]: #1229 +#1914 := [unit-resolution #1235 #1913]: #276 +#1915 := [unit-resolution #759 #1914 #1909 #1597]: up_53 +#1385 := (or #316 #325) +#329 := (or #325 #316) +#1387 := (iff #329 #1385) +#1388 := [rewrite]: #1387 +#1384 := [asserted]: #329 +#1391 := [mp #1384 #1388]: #1385 +#1916 := [unit-resolution #1391 #1915 #1908]: false +#1918 := [lemma #1916]: #1917 +#2032 := [unit-resolution #1918 #1597 #2029 #2031 #2030]: #226 +#2010 := (or up_12 up_29 up_7 up_54 up_26) +#1993 := (or up_35 up_12 up_54 up_26 up_29 up_7) +#1955 := (or #170 up_54 up_26 up_29 up_12 up_35 up_7) +#1940 := [unit-resolution #1625 #1805 #907 #1457 #1611 #1612]: up_32 +#1941 := [unit-resolution #1240 #1940]: #276 +#1942 := [unit-resolution #1230 #1940]: #275 +#1943 := [unit-resolution #608 #1806 #1611]: up_33 +#1944 := [unit-resolution #1113 #1943]: #243 +#1925 := (or #325 up_34 up_42) +#1919 := [hypothesis]: up_56 +#1920 := [unit-resolution #1373 #1919]: #324 +#1921 := [unit-resolution #1516 #1920]: up_46 +#1922 := [unit-resolution #1396 #1919]: #284 +#1923 := [unit-resolution #685 #1922 #1605 #1798]: up_44 +#1924 := [unit-resolution #1276 #1923 #1921]: false +#1926 := [lemma #1924]: #1925 +#1945 := [unit-resolution #1926 #1944 #1942]: #325 +#1946 := [unit-resolution #1524 #1945]: up_60 +#1947 := [unit-resolution #1448 #1946]: #341 +#1938 := (or #308 up_26 up_59) +#1927 := [hypothesis]: up_51 +#1928 := [unit-resolution #1329 #1927]: #300 +#1929 := [hypothesis]: #341 +#1930 := [unit-resolution #1321 #1927]: #309 +#1931 := [unit-resolution #819 #1930 #1929]: up_58 +#1932 := [unit-resolution #1781 #1931 #1928]: up_39 +#1933 := [unit-resolution #1183 #1932]: #254 +#1934 := [unit-resolution #1418 #1931]: #335 +#1935 := [unit-resolution #795 #1934]: up_48 +#1936 := [unit-resolution #1297 #1935]: #255 +#1937 := [unit-resolution #621 #1936 #1933 #1890]: false +#1939 := [lemma #1937]: #1938 +#1948 := [unit-resolution #1939 #1947 #1890]: #308 +#1949 := [unit-resolution #759 #1948 #1941 #1597]: up_53 +#1950 := [unit-resolution #1381 #1949]: #324 +#1951 := [unit-resolution #1516 #1950]: up_46 +#1952 := [unit-resolution #1401 #1949]: #284 +#1953 := [unit-resolution #685 #1952 #1944 #1942]: up_44 +#1954 := [unit-resolution #1276 #1953 #1951]: false +#1956 := [lemma #1954]: #1955 +#1980 := [unit-resolution #1956 #1611 #1890 #907 #1457 #1597 #1612]: #170 +#1981 := [unit-resolution #430 #1980]: up_11 +#1982 := [unit-resolution #886 #1981]: #160 +#1983 := [unit-resolution #410 #1982 #1612]: up_10 +#1984 := [unit-resolution #980 #1983]: #194 +#1985 := [unit-resolution #1979 #1611 #907 #1984 #1457 #1612]: up_32 +#1970 := (or #235 up_34 up_54 up_26) +#1957 := [hypothesis]: up_32 +#1958 := [unit-resolution #1240 #1957]: #276 +#1959 := [unit-resolution #1230 #1957]: #275 +#1960 := [unit-resolution #1926 #1959 #1605]: #325 +#1961 := [unit-resolution #1524 #1960]: up_60 +#1962 := [unit-resolution #1448 #1961]: #341 +#1963 := [unit-resolution #1939 #1962 #1890]: #308 +#1964 := [unit-resolution #759 #1963 #1958 #1597]: up_53 +#1965 := [unit-resolution #1381 #1964]: #324 +#1966 := [unit-resolution #1516 #1965]: up_46 +#1967 := [unit-resolution #1401 #1964]: #284 +#1968 := [unit-resolution #685 #1967 #1605 #1959]: up_44 +#1969 := [unit-resolution #1276 #1968 #1966]: false +#1971 := [lemma #1969]: #1970 +#1986 := [unit-resolution #1971 #1985 #1597 #1890]: up_34 +#1987 := [unit-resolution #1113 #1986]: #242 +#1988 := [unit-resolution #608 #1987 #1611]: up_24 +#1989 := [unit-resolution #970 #1983]: #193 +#1990 := [unit-resolution #1136 #1986]: #202 +#1991 := [unit-resolution #505 #1990 #1457 #1989]: up_22 +#1992 := [unit-resolution #1016 #1991 #1988]: false +#1994 := [lemma #1992]: #1993 +#1995 := [unit-resolution #1994 #1457 #1597 #1890 #907 #1612]: up_35 +#1996 := [unit-resolution #1281 #1995]: #291 +#1997 := [unit-resolution #1516 #1996]: up_55 +#1998 := [unit-resolution #1373 #1997]: #325 +#1999 := [unit-resolution #1524 #1998]: up_60 +#2000 := [unit-resolution #1448 #1999]: #341 +#2001 := [unit-resolution #1939 #2000 #1890]: #308 +#2002 := [unit-resolution #1610 #2001 #1995 #1597]: up_34 +#2003 := [unit-resolution #1131 #2002]: #234 +#2004 := [unit-resolution #1381 #1997]: #316 +#2005 := [unit-resolution #759 #2001 #2004 #1597]: up_43 +#2006 := [unit-resolution #1240 #2005]: #235 +#2007 := [unit-resolution #1136 #2002]: #202 +#2008 := [unit-resolution #1683 #2007 #1612 #1457]: #194 +#2009 := [unit-resolution #579 #2008 #2006 #907 #2003]: false +#2011 := [lemma #2009]: #2010 +#2033 := [unit-resolution #2011 #2032 #1612 #1597 #2029]: up_12 +#2034 := [unit-resolution #891 #2033]: #160 +#2035 := [unit-resolution #410 #2034 #1612]: up_10 +#2036 := [unit-resolution #980 #2035]: #194 +#2037 := [unit-resolution #878 #2033]: #165 +#2038 := [unit-resolution #430 #2037]: up_13 +#2039 := [unit-resolution #1021 #2038]: #209 +#2024 := (or #234 up_26 up_54 up_24) +#2012 := [hypothesis]: #209 +#2013 := [hypothesis]: up_31 +#2014 := [unit-resolution #1121 #2013]: #242 +#2015 := [unit-resolution #608 #2014 #2012]: up_35 +#2016 := [unit-resolution #1131 #2013]: #243 +#2017 := [unit-resolution #1610 #2016 #2015 #1597]: up_51 +#2018 := [unit-resolution #1939 #2017 #1890]: up_59 +#2019 := [unit-resolution #1448 #2018]: #345 +#2020 := [unit-resolution #1281 #2015]: #291 +#2021 := [unit-resolution #1516 #2020]: up_55 +#2022 := [unit-resolution #1373 #2021]: #325 +#2023 := [unit-resolution #1524 #2022 #2019]: false +#2025 := [lemma #2023]: #2024 +#2040 := [unit-resolution #2025 #2029 #1597 #2039]: #234 +#2041 := [unit-resolution #579 #2040 #2032 #2036]: up_32 +#2042 := [unit-resolution #1240 #2041]: #276 +#2043 := [unit-resolution #1971 #2041 #1597 #2029]: up_34 +#2044 := [unit-resolution #1113 #2043]: #242 +#2045 := [unit-resolution #608 #2044 #2039]: up_35 +#2046 := [unit-resolution #1281 #2045]: #291 +#2047 := [unit-resolution #1516 #2046]: up_55 +#2048 := [unit-resolution #1381 #2047]: #316 +#2049 := [unit-resolution #759 #2048 #2042 #1597]: up_51 +#2050 := [unit-resolution #1373 #2047]: #325 +#2051 := [unit-resolution #1524 #2050]: up_60 +#2052 := [unit-resolution #1448 #2051]: #341 +#2053 := [unit-resolution #1939 #2052 #2049 #2029]: false +#2055 := [lemma #2053]: #2054 +#2065 := [unit-resolution #2055 #1612 #1626]: up_54 +#1447 := (or #317 #345) +#347 := (or #345 #317) +#1449 := (iff #347 #1447) +#1450 := [rewrite]: #1449 +#1446 := [asserted]: #347 +#1453 := [mp #1446 #1450]: #1447 +#2066 := [unit-resolution #1453 #2065]: #345 +#2067 := [unit-resolution #1524 #2066]: up_56 +#2083 := (or #275 up_7 up_12 up_8) +#2063 := [hypothesis]: up_42 +#2064 := [unit-resolution #1230 #2063]: #235 +#2068 := [unit-resolution #1373 #2067]: #324 +#2069 := [unit-resolution #1516 #2068]: up_46 +#2070 := [unit-resolution #1281 #2069]: #250 +#2071 := [unit-resolution #1672 #2064 #1626 #2070 #1612]: up_16 +#2072 := [unit-resolution #913 #2071]: #172 +#2073 := [unit-resolution #1508 #2072]: up_15 +#2074 := [unit-resolution #1032 #2073]: #213 +#2075 := [unit-resolution #905 #2071]: #178 +#1452 := (or #317 #341) +#348 := (or #341 #317) +#1454 := (iff #348 #1452) +#1455 := [rewrite]: #1454 +#1451 := [asserted]: #348 +#1458 := [mp #1451 #1455]: #1452 +#2076 := [unit-resolution #1458 #2065]: #341 +#2077 := [unit-resolution #1225 #2063]: #267 +#2061 := (or #226 up_59 up_40 up_17 up_25) +#2056 := [unit-resolution #653 #1910 #1782 #1911]: up_41 +#2057 := [unit-resolution #1349 #2056]: #300 +#1338 := (or #268 #309) +#314 := (or #309 #268) +#1340 := (iff #314 #1338) +#1341 := [rewrite]: #1340 +#1337 := [asserted]: #314 +#1344 := [mp #1337 #1341]: #1338 +#2058 := [unit-resolution #1344 #2056]: #309 +#2059 := [unit-resolution #819 #2058 #1929]: up_58 +#2060 := [unit-resolution #1781 #2059 #2057 #1901]: false +#2062 := [lemma #2060]: #2061 +#2078 := [unit-resolution #2062 #2077 #2076 #2075 #2074]: #226 +#2079 := [unit-resolution #1625 #2078 #2070 #1457 #2064 #1612]: up_22 +#2080 := [unit-resolution #1979 #2078 #2070 #1457 #2064 #1612]: up_21 +#2081 := [unit-resolution #1683 #2080 #1612 #1457]: up_23 +#2082 := [unit-resolution #983 #2081 #2079]: false +#2084 := [lemma #2082]: #2083 +#2085 := [unit-resolution #2084 #1457 #1612 #1626]: #275 +#2086 := [unit-resolution #1926 #2085 #2067]: up_34 +#2087 := [unit-resolution #1136 #2086]: #202 +#2088 := [unit-resolution #1113 #2086]: #242 +#2089 := [unit-resolution #608 #2088 #2070]: up_24 +#2090 := [unit-resolution #1016 #2089]: #201 +#2091 := [unit-resolution #505 #2090 #1457 #2087]: up_20 +#2092 := [unit-resolution #970 #2091]: #161 +#2093 := [unit-resolution #1021 #2089]: #170 +#2094 := [unit-resolution #430 #2093]: up_11 +#2095 := [unit-resolution #886 #2094]: #160 +#2096 := [unit-resolution #410 #2095 #2092 #1612]: false +#2098 := [lemma #2096]: #2097 +#2102 := [unit-resolution #2098 #2100 #2101]: up_12 +#2103 := [unit-resolution #891 #2102]: #160 +#2104 := [unit-resolution #410 #2103 #2100]: up_10 +#2105 := [unit-resolution #980 #2104]: #194 +#2106 := [unit-resolution #2055 #2100 #2101]: up_54 +#2107 := [unit-resolution #1453 #2106]: #345 +#2108 := [unit-resolution #1524 #2107]: up_56 +#2109 := [unit-resolution #1373 #2108]: #324 +#2110 := [unit-resolution #1516 #2109]: up_46 +#2111 := [unit-resolution #1281 #2110]: #250 +#2112 := [unit-resolution #878 #2102]: #165 +#2113 := [unit-resolution #430 #2112]: up_13 +#2114 := [unit-resolution #1021 #2113]: #209 +#2115 := [unit-resolution #608 #2114 #2111]: up_33 +#2116 := [unit-resolution #1121 #2115]: #234 +#2117 := [unit-resolution #1276 #2110]: #283 +#2118 := [unit-resolution #1396 #2108]: #284 +#2119 := [unit-resolution #1113 #2115]: #243 +#2120 := [unit-resolution #685 #2119 #2118 #2117]: up_42 +#2121 := [unit-resolution #1230 #2120]: #235 +#2122 := [unit-resolution #579 #2121 #2116 #2105]: up_29 +#2123 := [unit-resolution #1225 #2120]: #267 +#2124 := [unit-resolution #1458 #2106]: #341 +#2125 := [unit-resolution #1672 #2121 #2101 #2111 #2100]: up_16 +#2126 := [unit-resolution #905 #2125]: #178 +#2127 := [unit-resolution #2062 #2126 #2124 #2123 #2122]: up_25 +#2128 := [unit-resolution #913 #2125]: #172 +#2129 := [unit-resolution #1508 #2128]: up_15 +#2130 := [unit-resolution #1032 #2129 #2127]: false +#2131 := [lemma #2130]: #150 +#1494 := (or up_5 up_6) +decl up_1 :: bool +#4 := up_1 +#379 := (or up_1 up_5 up_6) +#1497 := (iff #379 #1494) +#1491 := (or false up_5 up_6) +#1495 := (iff #1491 #1494) +#1496 := [rewrite]: #1495 +#1492 := (iff #379 #1491) +#1467 := (iff up_1 false) +#5 := (not up_1) +#1470 := (iff #5 #1467) +#1463 := (iff #1467 #5) +#1468 := [rewrite]: #1463 +#1471 := [symm #1468]: #1470 +#368 := [asserted]: #5 +#1472 := [mp #368 #1471]: #1467 +#1493 := [monotonicity #1472]: #1492 +#1498 := [trans #1493 #1496]: #1497 +#14 := (or up_6 up_1) +#15 := (or up_5 #14) +#382 := (iff #15 #379) +#373 := (or up_1 up_6) +#376 := (or up_5 #373) +#380 := (iff #376 #379) +#381 := [rewrite]: #380 +#377 := (iff #15 #376) +#374 := (iff #14 #373) +#375 := [rewrite]: #374 +#378 := [monotonicity #375]: #377 +#383 := [trans #378 #381]: #382 +#372 := [asserted]: #15 +#384 := [mp #372 #383]: #379 +#1499 := [mp #384 #1498]: #1494 +#2138 := [unit-resolution #1499 #2131]: up_6 +#151 := (not up_6) +#927 := (or #151 #172) +#184 := (or #172 #151) +#929 := (iff #184 #927) +#930 := [rewrite]: #929 +#926 := [asserted]: #184 +#933 := [mp #926 #930]: #927 +#2139 := [unit-resolution #933 #2138]: #172 +#2140 := [unit-resolution #1508 #2139]: up_15 +#2147 := [unit-resolution #1037 #2140]: #214 +#2159 := [unit-resolution #2011 #2147]: #2158 +#2160 := [unit-resolution #2159 #907 #1612 #1457]: up_54 +#2161 := [unit-resolution #1453 #2160]: #345 +#2162 := [unit-resolution #1524 #2161]: up_56 +#2163 := [unit-resolution #1926 #2162 #2157 #1605]: false +#2165 := [lemma #2163]: #2164 +#2166 := [unit-resolution #2165 #1605 #1612 #1611 #1457]: up_29 +#2148 := (or #226 up_54) +#2141 := [unit-resolution #1032 #2140]: #213 +#922 := (or #151 #178) +#183 := (or #178 #151) +#924 := (iff #183 #922) +#925 := [rewrite]: #924 +#921 := [asserted]: #183 +#928 := [mp #921 #925]: #922 +#2142 := [unit-resolution #928 #2138]: #178 +#2149 := [unit-resolution #1918 #2147 #2142 #2141]: #2148 +#2167 := [unit-resolution #2149 #2166]: up_54 +#2154 := (or #226 up_34 up_59) +#2143 := (or #226 up_59 up_40) +#2144 := [unit-resolution #2062 #2142 #2141]: #2143 +#2145 := [unit-resolution #2144 #1893 #1929]: up_40 +#2146 := [unit-resolution #1225 #2145]: #275 +#2150 := [unit-resolution #2149 #1893]: up_54 +#2151 := [unit-resolution #1453 #2150]: #345 +#2152 := [unit-resolution #1524 #2151]: up_56 +#2153 := [unit-resolution #1926 #2152 #2146 #1605]: false +#2155 := [lemma #2153]: #2154 +#2168 := [unit-resolution #2155 #2166 #1605]: up_59 +#2169 := [unit-resolution #1458 #2168 #2167]: false +#2171 := [lemma #2169]: #2170 +#2172 := [unit-resolution #2171 #1612 #1611 #1457]: up_34 +#2173 := [unit-resolution #1136 #2172]: #202 +#2174 := [unit-resolution #1113 #2172]: #242 +#2175 := [unit-resolution #608 #2174 #1611]: up_24 +#2176 := [unit-resolution #1016 #2175]: #201 +#2177 := [unit-resolution #505 #2176 #1457 #2173]: up_20 +#2178 := [unit-resolution #970 #2177]: #161 +#2179 := [unit-resolution #1021 #2175]: #170 +#2180 := [unit-resolution #430 #2179]: up_11 +#2181 := [unit-resolution #886 #2180]: #160 +#2182 := [unit-resolution #410 #2181 #2178 #1612]: false +#2184 := [lemma #2182]: #2183 +#2235 := [unit-resolution #2184 #1457 #1611]: up_7 +#157 := (or #155 #156) +#856 := [asserted]: #157 +#2236 := [unit-resolution #856 #2235]: #156 +#2299 := (or up_34 up_35 up_12) +#2283 := (or #186 up_34) +#2185 := [hypothesis]: up_19 +#2191 := [unit-resolution #1084 #2185]: #227 +#2186 := [unit-resolution #1089 #2185]: #218 +#2187 := (or up_27 up_28) +#2188 := [unit-resolution #547 #2142 #2141]: #2187 +#2189 := [unit-resolution #2188 #2186]: up_28 +#2192 := [unit-resolution #1178 #2189]: #259 +#2265 := [unit-resolution #1193 #2189]: #254 +#2266 := (or up_36 up_37) +#2267 := [unit-resolution #621 #2147]: #2266 +#2268 := [unit-resolution #2267 #2265]: up_37 +#2269 := [unit-resolution #1292 #2268]: #295 +#2190 := [unit-resolution #1188 #2189]: #260 +#2270 := [unit-resolution #1297 #2268]: #296 +#2271 := [unit-resolution #795 #2270]: up_57 +#2272 := [unit-resolution #1428 #2271]: #301 +#2273 := [unit-resolution #727 #2272 #2190 #2269]: up_49 +#2274 := [unit-resolution #1349 #2273]: #268 +#2275 := [unit-resolution #653 #2274 #2192 #2191]: up_40 +#2276 := [unit-resolution #1225 #2275]: #275 +#2277 := [unit-resolution #1418 #2271]: #337 +#2278 := [unit-resolution #1339 #2273]: #309 +#2279 := [unit-resolution #819 #2278 #2277]: up_59 +#2280 := [unit-resolution #1448 #2279]: #345 +#2281 := [unit-resolution #1524 #2280]: up_56 +#2282 := [unit-resolution #1926 #2281 #2276 #1605]: false +#2284 := [lemma #2282]: #2283 +#2292 := [unit-resolution #2284 #1605]: #186 +#2223 := (or up_8 up_18 up_19) +#912 := (or #151 #177) +#181 := (or #177 #151) +#914 := (iff #181 #912) +#915 := [rewrite]: #914 +#911 := [asserted]: #181 +#918 := [mp #911 #915]: #912 +#2222 := [unit-resolution #918 #2138]: #177 +#2224 := [unit-resolution #473 #2222]: #2223 +#2293 := [unit-resolution #2224 #2292 #2236]: up_18 +#2257 := (or #235 up_34) +#2252 := (or #235 up_34 up_54) +#2253 := [unit-resolution #1971 #2147]: #2252 +#2254 := [unit-resolution #2253 #1957 #1605]: up_54 +#2255 := [unit-resolution #1453 #2254]: #345 +#2256 := [unit-resolution #1524 #2255 #1960]: false +#2258 := [lemma #2256]: #2257 +#2294 := [unit-resolution #2258 #1605]: #235 +#2290 := (or up_29 up_35 up_32 up_12 #185) +#2200 := [hypothesis]: up_18 +#2206 := (or #185 up_29 up_32 up_12 up_22) +#2201 := [unit-resolution #965 #2200]: #193 +#2202 := [unit-resolution #505 #2201 #1457 #1456]: up_23 +#2203 := [unit-resolution #975 #2200]: #194 +#2204 := [unit-resolution #579 #2203 #907 #895]: up_31 +#2205 := [unit-resolution #1141 #2204 #2202]: false +#2207 := [lemma #2205]: #2206 +#2285 := [unit-resolution #2207 #907 #895 #1457 #2200]: up_22 +#2286 := [unit-resolution #1016 #2285]: #209 +#2287 := [unit-resolution #579 #907 #895 #2203]: up_31 +#2288 := [unit-resolution #1121 #2287]: #242 +#2289 := [unit-resolution #608 #2288 #2286 #1611]: false +#2291 := [lemma #2289]: #2290 +#2295 := [unit-resolution #2291 #2294 #1611 #1457 #2293]: up_29 +#2296 := [unit-resolution #2149 #2295]: up_54 +#2297 := [unit-resolution #2155 #2295 #1605]: up_59 +#2298 := [unit-resolution #1458 #2297 #2296]: false +#2300 := [lemma #2298]: #2299 +#2301 := [unit-resolution #2300 #1457 #1611]: up_34 +#2302 := [unit-resolution #1136 #2301]: #202 +#2303 := [unit-resolution #1113 #2301]: #242 +#2304 := [unit-resolution #608 #2303 #1611]: up_24 +#2305 := [unit-resolution #1016 #2304]: #201 +#2306 := [unit-resolution #505 #2305 #1457 #2302]: up_20 +#2307 := [unit-resolution #965 #2306]: #185 +#2308 := [unit-resolution #2224 #2307 #2236]: up_19 +#2309 := [unit-resolution #957 #2306]: #194 +#2310 := [unit-resolution #1131 #2301]: #234 +#2311 := [unit-resolution #1074 #2308]: #226 +#2312 := [unit-resolution #579 #2311 #2310 #2309]: up_32 +#2313 := [unit-resolution #1245 #2312]: #267 +#2198 := (or #186 up_59 up_40) +#2193 := [unit-resolution #653 #2192 #1782 #2191]: up_41 +#2194 := [unit-resolution #1349 #2193]: #300 +#2195 := [unit-resolution #1344 #2193]: #309 +#2196 := [unit-resolution #819 #2195 #1929]: up_58 +#2197 := [unit-resolution #1781 #2196 #2194 #2190]: false +#2199 := [lemma #2197]: #2198 +#2314 := [unit-resolution #2199 #2313 #2308]: up_59 +#2315 := [unit-resolution #1448 #2314]: #345 +#2316 := [unit-resolution #1524 #2315]: up_56 +#2317 := [unit-resolution #1084 #2308]: #227 +#2318 := [unit-resolution #1089 #2308]: #218 +#2319 := [unit-resolution #2188 #2318]: up_28 +#2320 := [unit-resolution #1178 #2319]: #259 +#2321 := [unit-resolution #653 #2313 #2320 #2317]: up_41 +#2322 := [unit-resolution #1334 #2321]: #308 +#2323 := [unit-resolution #1240 #2312]: #276 +#2324 := [unit-resolution #1458 #2314]: #317 +#2325 := [unit-resolution #759 #2324 #2323 #2322]: up_53 +#2326 := [unit-resolution #1391 #2325 #2316]: false +#2328 := [lemma #2326]: #2327 +#2337 := [unit-resolution #2328 #1611]: up_12 +#2338 := [unit-resolution #878 #2337]: #165 +#2339 := [unit-resolution #430 #2338]: up_13 +#2340 := [unit-resolution #1021 #2339]: #209 +#2341 := [unit-resolution #608 #2340 #1611]: up_33 +#2342 := [unit-resolution #1113 #2341]: #243 +#2343 := [unit-resolution #2258 #2342]: #235 +#2344 := [unit-resolution #1121 #2341]: #234 +#2345 := [unit-resolution #2284 #2342]: #186 +#2346 := [unit-resolution #891 #2337]: #160 +#2335 := (or #194 up_9 up_19) +#2329 := [hypothesis]: #186 +#2330 := [unit-resolution #975 #1674]: #185 +#2331 := [unit-resolution #2224 #2330 #2329]: up_8 +#2332 := [hypothesis]: #160 +#2333 := [unit-resolution #410 #1678 #2332]: up_7 +#2334 := [unit-resolution #856 #2333 #2331]: false +#2336 := [lemma #2334]: #2335 +#2347 := [unit-resolution #2336 #2346 #2345]: #194 +#2348 := [unit-resolution #579 #2347 #2344 #2343]: up_29 +#2349 := [unit-resolution #2149 #2348]: up_54 +#2350 := [unit-resolution #2155 #2348 #2342]: up_59 +#2351 := [unit-resolution #1458 #2350 #2349]: false +#2352 := [lemma #2351]: up_35 +#2353 := [unit-resolution #1281 #2352]: #291 +#2354 := [unit-resolution #1516 #2353]: up_55 +#2355 := [unit-resolution #1373 #2354]: #325 +#2356 := [unit-resolution #1524 #2355]: up_60 +#2357 := [unit-resolution #1453 #2356]: #317 +#2358 := [unit-resolution #2149 #2357]: #226 +#2359 := [unit-resolution #1448 #2356]: #341 +#2217 := (or #308 up_59) +#2218 := [unit-resolution #1939 #2147]: #2217 +#2360 := [unit-resolution #2218 #2359]: #308 +#2361 := [unit-resolution #1381 #2354]: #316 +#2362 := [unit-resolution #759 #2357 #2361 #2360]: up_43 +#2363 := [unit-resolution #1235 #2362]: #267 +#2364 := [unit-resolution #2199 #2363 #2359]: #186 +#1145 := (or #209 #250) +#252 := (or #250 #209) +#1147 := (iff #252 #1145) +#1148 := [rewrite]: #1147 +#1144 := [asserted]: #252 +#1151 := [mp #1144 #1148]: #1145 +#2365 := [unit-resolution #1151 #2352]: #209 +#2230 := (or #234 up_54 up_24) +#2231 := [unit-resolution #2025 #2147]: #2230 +#2366 := [unit-resolution #2231 #2357 #2365]: #234 +#2367 := [unit-resolution #1240 #2362]: #235 +#2368 := [unit-resolution #579 #2367 #2366 #2358]: up_21 +#2369 := [unit-resolution #2336 #2368 #2364]: up_9 +#870 := (or #155 #160) +#163 := (or #160 #155) +#871 := (iff #163 #870) +#872 := [rewrite]: #871 +#868 := [asserted]: #163 +#875 := [mp #868 #872]: #870 +#2370 := [unit-resolution #875 #2369]: #155 +#2371 := [unit-resolution #891 #2369]: #166 +[unit-resolution #2159 #2371 #2370 #2358 #2357]: false +unsat diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/SMT.thy --- a/src/HOL/SMT/SMT.thy Tue Oct 20 08:10:47 2009 +0200 +++ b/src/HOL/SMT/SMT.thy Tue Oct 20 10:11:30 2009 +0200 @@ -2,50 +2,20 @@ Author: Sascha Boehme, TU Muenchen *) -header {* SMT method using external SMT solvers (CVC3, Yices, Z3) *} +header {* Bindings to several SMT solvers *} theory SMT -imports SMT_Definitions +imports SMT_Base Z3 uses - "Tools/smt_normalize.ML" - "Tools/smt_monomorph.ML" - "Tools/smt_translate.ML" - "Tools/smt_solver.ML" - "Tools/smtlib_interface.ML" "Tools/cvc3_solver.ML" "Tools/yices_solver.ML" - "Tools/z3_model.ML" - "Tools/z3_interface.ML" - "Tools/z3_solver.ML" begin -setup {* - SMT_Normalize.setup #> - SMT_Solver.setup #> - CVC3_Solver.setup #> - Yices_Solver.setup #> - Z3_Solver.setup -*} +setup {* CVC3_Solver.setup #> Yices_Solver.setup *} -ML {* -OuterSyntax.improper_command "smt_status" - "Show the available SMT solvers and the currently selected solver." - OuterKeyword.diag - (Scan.succeed (Toplevel.no_timing o Toplevel.keep (fn state => - SMT_Solver.print_setup (Context.Proof (Toplevel.context_of state))))) -*} - -method_setup smt = {* - let fun solver thms ctxt = SMT_Solver.smt_tac ctxt thms - in - Scan.optional (Scan.lift (Args.add -- Args.colon) |-- Attrib.thms) [] >> - (Method.SIMPLE_METHOD' oo solver) - end -*} "Applies an SMT solver to the current goal." - -declare [[ smt_solver = z3, smt_timeout = 20, smt_trace = false ]] +declare [[ smt_solver = z3, smt_timeout = 20 ]] declare [[ smt_unfold_defs = true ]] -declare [[ z3_proofs = false ]] +declare [[ smt_trace = false, smt_keep = "", smt_cert = "" ]] +declare [[ z3_proofs = false, z3_options = "" ]] end - diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/SMT_Base.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/SMT_Base.thy Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,131 @@ +(* Title: HOL/SMT/SMT_Base.thy + Author: Sascha Boehme, TU Muenchen +*) + +header {* SMT-specific definitions and basic tools *} + +theory SMT_Base +imports Real Word "~~/src/HOL/Decision_Procs/Dense_Linear_Order" +uses + ("Tools/smt_normalize.ML") + ("Tools/smt_monomorph.ML") + ("Tools/smt_translate.ML") + ("Tools/smt_solver.ML") + ("Tools/smtlib_interface.ML") +begin + +section {* Triggers for quantifier instantiation *} + +text {* +Some SMT solvers support triggers for quantifier instantiation. Each trigger +consists of one ore more patterns. A pattern may either be a list of positive +subterms (the first being tagged by "pat" and the consecutive subterms tagged +by "andpat"), or a list of negative subterms (the first being tagged by "nopat" +and the consecutive subterms tagged by "andpat"). +*} + +datatype pattern = Pattern + +definition pat :: "'a \ pattern" +where "pat _ = Pattern" + +definition nopat :: "bool \ pattern" +where "nopat _ = Pattern" + +definition andpat :: "pattern \ 'a \ pattern" (infixl "andpat" 60) +where "_ andpat _ = Pattern" + +definition trigger :: "pattern list \ bool \ bool" +where "trigger _ P = P" + + +section {* Arithmetic *} + +text {* +The sign of @{term "op mod :: int \ int \ int"} follows the sign of the +divisor. In contrast to that, the sign of the following operation is that of +the dividend. +*} + +definition rem :: "int \ int \ int" (infixl "rem" 70) +where "a rem b = + (if (a \ 0 \ b < 0) \ (a < 0 \ b \ 0) then - (a mod b) else a mod b)" + +text {* A decision procedure for linear real arithmetic: *} + +setup {* + Arith_Data.add_tactic "Ferrante-Rackoff" (K FerranteRackoff.dlo_tac) +*} + + +section {* Bitvectors *} + +text {* +The following definitions provide additional functions not found in HOL-Word. +*} + +definition sdiv :: "'a::len word \ 'a word \ 'a word" (infix "sdiv" 70) +where "w1 sdiv w2 = word_of_int (sint w1 div sint w2)" + +definition smod :: "'a::len word \ 'a word \ 'a word" (infix "smod" 70) + (* sign follows divisor *) +where "w1 smod w2 = word_of_int (sint w1 mod sint w2)" + +definition srem :: "'a::len word \ 'a word \ 'a word" (infix "srem" 70) + (* sign follows dividend *) +where "w1 srem w2 = word_of_int (sint w1 rem sint w2)" + +definition bv_shl :: "'a::len0 word \ 'a word \ 'a word" +where "bv_shl w1 w2 = (w1 << unat w2)" + +definition bv_lshr :: "'a::len0 word \ 'a word \ 'a word" +where "bv_lshr w1 w2 = (w1 >> unat w2)" + +definition bv_ashr :: "'a::len word \ 'a word \ 'a word" +where "bv_ashr w1 w2 = (w1 >>> unat w2)" + + +section {* Higher-Order Encoding *} + +definition "apply" where "apply f x = f x" + +lemmas array_rules = apply_def fun_upd_same fun_upd_other fun_upd_upd ext + + +section {* First-order logic *} + +text {* +Some SMT solver formats require a strict separation between formulas and terms. +The following marker symbols are used internally to separate those categories: +*} + +definition formula :: "bool \ bool" where "formula x = x" +definition "term" where "term x = x" + +text {* +Predicate symbols also occurring as function symbols are turned into function +symbols by translating atomic formulas into terms: +*} + +abbreviation holds :: "bool \ bool" where "holds \ (\P. term P = term True)" + +text {* +The following constant represents equivalence, to be treated differently than +the (polymorphic) equality predicate: +*} + +definition iff :: "bool \ bool \ bool" (infix "iff" 50) where + "(x iff y) = (x = y)" + + +section {* Setup *} + +use "Tools/smt_normalize.ML" +use "Tools/smt_monomorph.ML" +use "Tools/smt_translate.ML" +use "Tools/smt_solver.ML" +use "Tools/smtlib_interface.ML" + +setup {* SMT_Normalize.setup #> SMT_Solver.setup *} + +end diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/SMT_Definitions.thy --- a/src/HOL/SMT/SMT_Definitions.thy Tue Oct 20 08:10:47 2009 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,113 +0,0 @@ -(* Title: HOL/SMT/SMT_Definitions.thy - Author: Sascha Boehme, TU Muenchen -*) - -header {* SMT-specific definitions *} - -theory SMT_Definitions -imports Real Word "~~/src/HOL/Decision_Procs/Dense_Linear_Order" -begin - -section {* Triggers for quantifier instantiation *} - -text {* -Some SMT solvers support triggers for quantifier instantiation. Each trigger -consists of one ore more patterns. A pattern may either be a list of positive -subterms (the first being tagged by "pat" and the consecutive subterms tagged -by "andpat"), or a list of negative subterms (the first being tagged by "nopat" -and the consecutive subterms tagged by "andpat"). -*} - -datatype pattern = Pattern - -definition pat :: "'a \ pattern" -where "pat _ = Pattern" - -definition nopat :: "bool \ pattern" -where "nopat _ = Pattern" - -definition andpat :: "pattern \ 'a \ pattern" (infixl "andpat" 60) -where "_ andpat _ = Pattern" - -definition trigger :: "pattern list \ bool \ bool" -where "trigger _ P = P" - - -section {* Arithmetic *} - -text {* -The sign of @{term "op mod :: int \ int \ int"} follows the sign of the -divisor. In contrast to that, the sign of the following operation is that of -the dividend. -*} - -definition rem :: "int \ int \ int" (infixl "rem" 70) -where "a rem b = - (if (a \ 0 \ b < 0) \ (a < 0 \ b \ 0) then - (a mod b) else a mod b)" - -text {* A decision procedure for linear real arithmetic: *} - -setup {* - Arith_Data.add_tactic "Ferrante-Rackoff" (K FerranteRackoff.dlo_tac) -*} - - -section {* Bitvectors *} - -text {* -The following definitions provide additional functions not found in HOL-Word. -*} - -definition sdiv :: "'a::len word \ 'a word \ 'a word" (infix "sdiv" 70) -where "w1 sdiv w2 = word_of_int (sint w1 div sint w2)" - -definition smod :: "'a::len word \ 'a word \ 'a word" (infix "smod" 70) - (* sign follows divisor *) -where "w1 smod w2 = word_of_int (sint w1 mod sint w2)" - -definition srem :: "'a::len word \ 'a word \ 'a word" (infix "srem" 70) - (* sign follows dividend *) -where "w1 srem w2 = word_of_int (sint w1 rem sint w2)" - -definition bv_shl :: "'a::len0 word \ 'a word \ 'a word" -where "bv_shl w1 w2 = (w1 << unat w2)" - -definition bv_lshr :: "'a::len0 word \ 'a word \ 'a word" -where "bv_lshr w1 w2 = (w1 >> unat w2)" - -definition bv_ashr :: "'a::len word \ 'a word \ 'a word" -where "bv_ashr w1 w2 = (w1 >>> unat w2)" - - -section {* Higher-order encoding *} - -definition "apply" where "apply f x = f x" - - -section {* First-order logic *} - -text {* -Some SMT solver formats require a strict separation between formulas and terms. -The following marker symbols are used internally to separate those categories: -*} - -definition formula :: "bool \ bool" where "formula x = x" -definition "term" where "term x = x" - -text {* -Predicate symbols also occurring as function symbols are turned into function -symbols by translating atomic formulas into terms: -*} - -abbreviation holds :: "bool \ bool" where "holds \ (\P. term P = term True)" - -text {* -The following constant represents equivalence, to be treated differently than -the (polymorphic) equality predicate: -*} - -definition iff :: "bool \ bool \ bool" (infix "iff" 50) where - "(x iff y) = (x = y)" - -end - diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Tools/cvc3_solver.ML --- a/src/HOL/SMT/Tools/cvc3_solver.ML Tue Oct 20 08:10:47 2009 +0200 +++ b/src/HOL/SMT/Tools/cvc3_solver.ML Tue Oct 20 10:11:30 2009 +0200 @@ -27,10 +27,8 @@ fun raise_cex real ctxt recon ls = let - val start = String.isPrefix "%Satisfiable Variable Assignment: %" - val index = find_index start ls - val ls = if index > 0 then Library.drop (index + 1, ls) else [] - val p = Pretty.big_list (cex_kind real ^ " found:") (map Pretty.str ls) + 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 core_oracle (SMT_Solver.ProofData {context, output, recon, ...}) = @@ -47,9 +45,9 @@ fun smtlib_solver oracle _ = SMT_Solver.SolverConfig { - name = {env_var=env_var, remote_name=solver_name}, + command = {env_var=env_var, remote_name=solver_name}, + arguments = options, interface = SMTLIB_Interface.interface, - arguments = options, reconstruct = oracle } val setup = diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Tools/smt_normalize.ML --- a/src/HOL/SMT/Tools/smt_normalize.ML Tue Oct 20 08:10:47 2009 +0200 +++ b/src/HOL/SMT/Tools/smt_normalize.ML Tue Oct 20 10:11:30 2009 +0200 @@ -273,59 +273,78 @@ fun inst_meta cT = Thm.instantiate_cterm ([(meta_eqT, cT)], []) meta_eq fun mk_meta_eq ct cu = Thm.mk_binop (inst_meta (Thm.ctyp_of_term ct)) ct cu - fun lambda_conv conv = - let - fun sub_conv cvs ctxt ct = - (case Thm.term_of ct of - Const (@{const_name All}, _) $ Abs _ => quant_conv cvs ctxt - | Const (@{const_name Ex}, _) $ Abs _ => quant_conv cvs ctxt - | Const _ $ Abs _ => Conv.arg_conv (at_lambda_conv cvs ctxt) - | Const (@{const_name Let}, _) $ _ $ Abs _ => Conv.combination_conv - (Conv.arg_conv (sub_conv cvs ctxt)) (abs_conv cvs ctxt) - | Abs _ => at_lambda_conv cvs ctxt - | _ $ _ => Conv.comb_conv (sub_conv cvs ctxt) - | _ => Conv.all_conv) ct - and abs_conv cvs = Conv.abs_conv (fn (cv, cx) => sub_conv (cv::cvs) cx) - and quant_conv cvs ctxt = Conv.arg_conv (abs_conv cvs ctxt) - and at_lambda_conv cvs ctxt = abs_conv cvs ctxt then_conv conv cvs ctxt - in sub_conv [] end + val fresh_name = yield_singleton Name.variants fun used_vars cvs ct = let val lookup = AList.lookup (op aconv) (map (` Thm.term_of) cvs) val add = (fn (SOME ct) => insert (op aconvc) ct | _ => I) in Term.fold_aterms (add o lookup) (Thm.term_of ct) [] end + fun make_def cvs eq = Thm.symmetric (fold norm_meta_def cvs eq) + fun add_def ct thm = Termtab.update (Thm.term_of ct, (serial (), thm)) - val rev_int_fst_ord = rev_order o int_ord o pairself fst - fun ordered_values tab = - Termtab.fold (fn (_, x) => OrdList.insert rev_int_fst_ord x) tab [] - |> map snd + fun replace ctxt cvs ct (cx as (nctxt, defs)) = + let + val cvs' = used_vars cvs ct + val ct' = fold 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 + SOME (_, eq) => (make_def cvs' eq, cx) + | NONE => + let + val {t, T, ...} = Thm.rep_cterm ct' + val (n, nctxt') = fresh_name "" nctxt + val eq = Thm.assume (mk_meta_eq (cert ctxt (Free (n, T))) ct') + in (make_def cvs' eq, (nctxt', add_def ct' eq defs)) end) + end + + fun none ct cx = (Thm.reflexive ct, cx) + fun in_comb f g ct cx = + let val (cu1, cu2) = Thm.dest_comb ct + in cx |> f cu1 ||>> g cu2 |>> uncurry Thm.combination end + fun in_arg f = in_comb none f + fun in_abs f cvs ct (nctxt, defs) = + let + val (n, nctxt') = fresh_name Name.uu nctxt + val (cv, cu) = Thm.dest_abs (SOME n) ct + in f (cv :: cvs) cu (nctxt', defs) |>> Thm.abstract_rule n cv end + + fun replace_lambdas ctxt = + let + fun repl cvs ct = + (case Thm.term_of ct of + Const (@{const_name All}, _) $ Abs _ => in_arg (in_abs repl cvs) + | Const (@{const_name Ex}, _) $ Abs _ => in_arg (in_abs repl cvs) + | Const _ $ Abs _ => in_arg (at_lambda cvs) + | Const (@{const_name Let}, _) $ _ $ Abs _ => + in_comb (in_arg (repl cvs)) (in_abs repl cvs) + | Abs _ => at_lambda cvs + | _ $ _ => in_comb (repl cvs) (repl cvs) + | _ => none) ct + and at_lambda cvs ct cx = + let + val (thm1, cx') = in_abs repl cvs ct cx + val (thm2, cx'') = replace ctxt cvs (Thm.rhs_of thm1) cx' + in (Thm.transitive thm1 thm2, cx'') end + in repl [] end in fun lift_lambdas ctxt thms = let val declare_frees = fold (Thm.fold_terms Term.declare_term_frees) - val names = Unsynchronized.ref (declare_frees thms (Name.make_context [])) - val fresh_name = Unsynchronized.change_result names o yield_singleton Name.variants + fun rewrite f thm cx = + let val (thm', cx') = f (Thm.cprop_of thm) cx + in (Thm.equal_elim thm' thm, cx') end - val defs = Unsynchronized.ref (Termtab.empty : (int * thm) Termtab.table) - fun add_def t thm = Unsynchronized.change defs (Termtab.update (t, (serial (), thm))) - fun make_def cvs eq = Thm.symmetric (fold norm_meta_def cvs eq) - fun def_conv cvs ctxt ct = - let - val cvs' = used_vars cvs ct - val ct' = fold Thm.cabs cvs' ct - in - (case Termtab.lookup (!defs) (Thm.term_of ct') of - SOME (_, eq) => make_def cvs' eq - | NONE => - let - val {t, T, ...} = Thm.rep_cterm ct' - val eq = mk_meta_eq (cert ctxt (Free (fresh_name "", T))) ct' - val thm = Thm.assume eq - in (add_def t thm; make_def cvs' thm) end) - end - val thms' = map (Conv.fconv_rule (lambda_conv def_conv ctxt)) thms - val eqs = ordered_values (!defs) + val rev_int_fst_ord = rev_order o int_ord o pairself fst + fun ordered_values tab = + Termtab.fold (fn (_, x) => OrdList.insert rev_int_fst_ord x) tab [] + |> map snd + + val (thms', (_, defs)) = + (declare_frees thms (Name.make_context []), Termtab.empty) + |> fold_map (rewrite (replace_lambdas ctxt)) thms + val eqs = ordered_values defs in (maps (#hyps o Thm.crep_thm) eqs, map (normalize_rule ctxt) eqs @ thms') end diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Tools/smt_solver.ML --- a/src/HOL/SMT/Tools/smt_solver.ML Tue Oct 20 08:10:47 2009 +0200 +++ b/src/HOL/SMT/Tools/smt_solver.ML Tue Oct 20 10:11:30 2009 +0200 @@ -19,9 +19,9 @@ assms: thm list option } datatype solver_config = SolverConfig of { - name: {env_var: string, remote_name: string}, + command: {env_var: string, remote_name: string}, + arguments: string list, interface: interface, - arguments: string list, reconstruct: proof_data -> thm } (*options*) @@ -29,6 +29,8 @@ val with_timeout: Proof.context -> ('a -> 'b) -> 'a -> 'b val trace: bool Config.T val trace_msg: Proof.context -> ('a -> string) -> 'a -> unit + val keep: string Config.T + val cert: string Config.T (*solvers*) type solver = Proof.context -> thm list -> thm @@ -67,9 +69,9 @@ assms: thm list option } datatype solver_config = SolverConfig of { - name: {env_var: string, remote_name: string}, + command: {env_var: string, remote_name: string}, + arguments: string list, interface: interface, - arguments: string list, reconstruct: proof_data -> thm } @@ -86,17 +88,28 @@ fun trace_msg ctxt f x = if Config.get ctxt trace then tracing (f x) else () +val (keep, setup_keep) = Attrib.config_string "smt_keep" "" +val (cert, setup_cert) = Attrib.config_string "smt_cert" "" + (* interface to external solvers *) local -fun with_tmp_files f x = +fun with_files ctxt f x = let - fun tmp_path () = File.tmp_path (Path.explode ("smt-" ^ serial_string ())) - val in_path = tmp_path () and out_path = tmp_path () - val y = Exn.capture (f in_path out_path) x - val _ = try File.rm in_path and _ = try File.rm out_path + fun make_names n = (n, n ^ ".proof") + + val keep' = Config.get ctxt keep + val paths as (problem_path, proof_path) = + if keep' <> "" andalso File.exists (Path.dir (Path.explode keep')) + then pairself Path.explode (make_names keep') + else pairself (File.tmp_path o Path.explode) + (make_names ("smt-" ^ serial_string ())) + + val y = Exn.capture (f problem_path proof_path) x + + val _ = if keep' = "" then (pairself (try File.rm) paths; ()) else () in Exn.release y end fun run in_path out_path (ctxt, cmd, output) = @@ -115,12 +128,18 @@ fun run_solver ctxt {env_var, remote_name} args output = let val qf = File.shell_path and qq = File.shell_quote - val path = getenv env_var and remote = getenv "REMOTE_SMT_SOLVER" + val qs = qf o Path.explode + val local_name = getenv env_var + val cert_name = Config.get ctxt cert + val remote = qs (getenv "REMOTE_SMT_SOLVER") + val cert_script = qs (getenv "CERT_SMT_SOLVER") fun cmd f1 f2 = - if path <> "" - then map qq (path :: args) @ [qf f1, ">", qf f2] - else "perl -w" :: map qq (remote :: remote_name :: args) @ [qf f1, qf f2] - in with_tmp_files run (ctxt, space_implode " " oo cmd, output) end + if cert_name <> "" + then "perl -w" :: [cert_script, qs cert_name, qf f1, ">", qf f2] + else if local_name <> "" + then qs local_name :: map qq args @ [qf f1, ">", qf f2] + else "perl -w" :: remote :: map qq (remote_name :: args) @ [qf f1, qf f2] + in with_files ctxt run (ctxt, space_implode " " oo cmd, output) end end @@ -129,12 +148,12 @@ fun gen_solver solver ctxt prems = let - val SolverConfig {name, interface, arguments, reconstruct} = solver ctxt + val SolverConfig {command, arguments, interface, reconstruct} = solver ctxt val Interface {normalize=nc, translate=tc} = interface val thy = ProofContext.theory_of ctxt in SMT_Normalize.normalize nc ctxt prems - ||> run_solver ctxt name arguments o SMT_Translate.translate tc thy + ||> run_solver ctxt command arguments o SMT_Translate.translate tc thy ||> reconstruct o make_proof_data ctxt |-> fold SMT_Normalize.discharge_definition end @@ -196,8 +215,8 @@ fun pretty_counterex ctxt (real, ex) = let - val msg = if real then "Counterexample found:" - else "Potential counterexample found:" + 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 @@ -212,6 +231,11 @@ val smt_tac = smt_tac' false +val smt_method = + Scan.optional (Scan.lift (Args.add -- Args.colon) |-- Attrib.thms) [] >> + (fn thms => fn ctxt => METHOD (fn facts => + HEADGOAL (smt_tac ctxt (thms @ facts)))) + (* setup *) @@ -221,7 +245,11 @@ (Thm.declaration_attribute o K o select_solver)) "SMT solver configuration" #> setup_timeout #> - setup_trace + setup_trace #> + setup_keep #> + setup_cert #> + Method.setup (Binding.name "smt") smt_method + "Applies an SMT solver to the current goal." fun print_setup gen = let @@ -243,4 +271,10 @@ Pretty.big_list "Solver-specific settings:" infos]) end +val _ = OuterSyntax.improper_command "smt_status" + "Show the available SMT solvers and the currently selected solver." + OuterKeyword.diag + (Scan.succeed (Toplevel.no_timing o Toplevel.keep (fn state => + print_setup (Context.Proof (Toplevel.context_of state))))) + end diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Tools/yices_solver.ML --- a/src/HOL/SMT/Tools/yices_solver.ML Tue Oct 20 08:10:47 2009 +0200 +++ b/src/HOL/SMT/Tools/yices_solver.ML Tue Oct 20 10:11:30 2009 +0200 @@ -40,9 +40,9 @@ fun smtlib_solver oracle _ = SMT_Solver.SolverConfig { - name = {env_var=env_var, remote_name=solver_name}, + command = {env_var=env_var, remote_name=solver_name}, + arguments = options, interface = SMTLIB_Interface.interface, - arguments = options, reconstruct = oracle } val setup = diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Tools/z3_proof.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Tools/z3_proof.ML Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,247 @@ +(* Title: HOL/SMT/Tools/z3_proof.ML + Author: Sascha Boehme, TU Muenchen + +Proof reconstruction for proofs found by Z3. +*) + +signature Z3_PROOF = +sig + val reconstruct: Proof.context -> thm list option -> SMT_Translate.recon -> + string list -> thm +end + +structure Z3_Proof: Z3_PROOF = +struct + +structure T = Z3_Proof_Terms +structure R = Z3_Proof_Rules + +fun z3_exn msg = error ("Z3 proof reconstruction: " ^ msg) + + +fun lift f (x, y) = apsnd (pair x) (f y) +fun lift' f v (x, y) = apsnd (rpair y) (f v x) + +fun $$ s = lift (Scan.$$ s) +fun this s = lift (Scan.this_string s) + +fun blank s = lift (Scan.many1 Symbol.is_ascii_blank) s + +fun par scan = $$ "(" |-- scan --| $$ ")" +fun bra scan = $$ "[" |-- scan --| $$ "]" + +val digit = (fn + "0" => SOME 0 | "1" => SOME 1 | "2" => SOME 2 | "3" => SOME 3 | + "4" => SOME 4 | "5" => SOME 5 | "6" => SOME 6 | "7" => SOME 7 | + "8" => SOME 8 | "9" => SOME 9 | _ => NONE) + +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 :|-- + (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 + +datatype sym = Sym of string * sym list + +datatype context = Context of { + Ttab: typ Symtab.table, + ttab: Thm.cterm Symtab.table, + etab: T.preterm Inttab.table, + ptab: R.proof Inttab.table, + nctxt: Name.context } + +fun make_context (Ttab, ttab, etab, ptab, nctxt) = + Context {Ttab=Ttab, ttab=ttab, etab=etab, ptab=ptab, nctxt=nctxt} + +fun empty_context thy (SMT_Translate.Recon {typs, terms=ttab}) = + let + val ttab' = Symtab.map (fn @{term True} => @{term "~False"} | t => t) ttab + val ns = Symtab.fold (Term.add_free_names o snd) ttab' [] + val nctxt = Name.make_context ns + val tt = Symtab.map (Thm.cterm_of thy) ttab' + in make_context (typs, tt, Inttab.empty, Inttab.empty, nctxt) end + +fun map_context f (Context {Ttab, ttab, etab, ptab, nctxt}) = + make_context (f (Ttab, ttab, etab, ptab, nctxt)) + +fun map_type_tab f = map_context (fn (Ttab, ttab, etab, ptab, nctxt) => + (f Ttab, ttab, etab, ptab, nctxt)) + +fun map_term_tab f = map_context (fn (Ttab, ttab, etab, ptab, nctxt) => + (Ttab, f ttab, etab, ptab, nctxt)) + +fun map_expr_tab f = map_context (fn (Ttab, ttab, etab, ptab, nctxt) => + (Ttab, ttab, f etab, ptab, nctxt)) + +fun map_proof_tab f = map_context (fn (Ttab, ttab, etab, ptab, nctxt) => + (Ttab, ttab, etab, f ptab, nctxt)) + +val free_prefix = "f" + +fun fresh_name (cx as Context {nctxt, ...}) = + let val (n, nctxt') = yield_singleton Name.variants free_prefix nctxt + in + (n, map_context (fn (Ttab, ttab, etab, ptab, _) => + (Ttab, ttab, etab, ptab, nctxt')) cx) + end + +fun typ_of_sort name (cx as Context {Ttab, ...}) = + (case Symtab.lookup Ttab name of + SOME T => (T, cx) + | _ => cx |> fresh_name |-> (fn n => + let val T = TFree ("'" ^ n, @{sort type}) + in pair T o map_type_tab (Symtab.update (name, T)) end)) + +fun lookup_expr id (cx as Context {etab, ...}) = + (case Inttab.lookup etab id of + SOME e => (e, cx) + | _ => z3_exn ("unknown term id: " ^ quote (string_of_int id))) + +fun add_expr k t = map_expr_tab (Inttab.update (k, t)) + +fun add_proof thy k ((r, ps), t) (cx as Context {nctxt, ...}) = + let val p = R.make_proof r ps (T.compile thy nctxt t) + in (k, map_proof_tab (Inttab.update (k, p)) cx) end + +fun mk_app app (cx as Context {ttab, ...}) = + let + val mk = + (fn + (Sym ("true", _), _) => T.mk_true + | (Sym ("false", _), _) => T.mk_false + | (Sym ("=", _), [t, u]) => T.mk_eq t u + | (Sym ("distinct", _), ts) => T.mk_distinct ts + | (Sym ("ite", _), [s, t, u]) => T.mk_if s t u + | (Sym ("and", _), ts) => T.mk_and ts + | (Sym ("or", _), ts) => T.mk_or ts + | (Sym ("iff", _), [t, u]) => T.mk_iff t u + | (Sym ("xor", _), [t, u]) => T.mk_not (T.mk_iff t u) + | (Sym ("not", _), [t]) => T.mk_not t + | (Sym ("implies", _), [t, u]) => T.mk_implies t u + | (Sym ("~", _), [t, u]) => T.mk_eq t u + | (Sym ("<", _), [t, u]) => T.mk_lt t u + | (Sym ("<=", _), [t, u]) => T.mk_le t u + | (Sym (">", _), [t, u]) => T.mk_lt u t + | (Sym (">=", _), [t, u]) => T.mk_le u t + | (Sym ("+", _), [t, u]) => T.mk_add t u + | (Sym ("-", _), [t, u]) => T.mk_sub t u + | (Sym ("-", _), [t]) => T.mk_uminus t + | (Sym ("*", _), [t, u]) => T.mk_mul t u + | (Sym ("/", _), [t, u]) => T.mk_real_div t u + | (Sym ("div", _), [t, u]) => T.mk_int_div t u + | (Sym ("mod", _), [t, u]) => T.mk_mod t u + | (Sym ("rem", _), [t, u]) => T.mk_rem t u + | (Sym ("select", _), [m, k]) => T.mk_access m k + | (Sym ("store", _), [m, k, v]) => T.mk_update m k v + | (Sym ("pattern", _), _) => T.mk_true + | (Sym (n, _), ts) => + (case Symtab.lookup ttab n of + SOME ct => T.mk_fun ct ts + | NONE => z3_exn ("unknown function: " ^ quote n))) + in (mk app, cx) end + +fun add_decl thy (n, T) (cx as Context {ttab, ...}) = + (case Symtab.lookup ttab n of + SOME _ => cx + | _ => cx |> fresh_name |-> (fn n' => + map_term_tab (Symtab.update (n, Thm.cterm_of thy (Free (n', T)))))) + + +fun sep scan_sep scan = scan ::: Scan.repeat (scan_sep |-- scan) +fun bsep scan = Scan.repeat (blank |-- scan) +fun bsep1 scan = Scan.repeat1 (blank |-- scan) + +val id = Scan.$$ "#" |-- int_num + +fun sym s = + (lift name -- Scan.optional (bra (sep ($$ ":") sym)) [] >> Sym) s + +fun sort st = Scan.first [ + this "bool" >> K @{typ bool}, + this "int" >> K @{typ int}, + this "real" >> K @{typ real}, + this "bv" |-- bra (lift int_num) >> T.wordT, + this "array" |-- bra (sort --| $$ ":" -- sort) >> (op -->), + par (this "->" |-- bsep1 sort) >> ((op --->) o split_last), + lift name #-> lift' typ_of_sort] st + +fun bound thy = + par (this ":var" -- blank |-- lift int_num --| blank -- sort) >> + uncurry (T.mk_bound thy) + +val number = + int_num -- Scan.option (Scan.$$ "/" |-- int_num) --| + Scan.this_string "::" :|-- (fn num as (n, _) => + Scan.this_string "int" >> K (T.mk_int_num n) || + Scan.this_string "real" >> K (T.mk_real_frac_num num)) + +fun bv_number thy = + this "bv" |-- bra (lift (int_num --| Scan.$$ ":" -- int_num)) >> + uncurry (T.mk_bv_num thy) + +val constant = sym #-> lift' (mk_app o rpair []) + +fun arg thy = Scan.first [lift id #-> lift' lookup_expr, + lift number, bv_number thy, constant] + +fun application thy = + par (sym -- bsep1 (arg thy)) #-> lift' mk_app + +val variables = + par (this "vars" |-- bsep1 (par ((lift name >> K "x") --| blank -- sort))) +val patterns = bsep (par ((this ":pat" || this ":nopat") |-- bsep1 (lift id))) +val quant_kind = + this "forall" >> K T.mk_forall || this "exists" >> K T.mk_exists +fun quantifier thy = par (quant_kind --| blank -- + variables --| patterns --| blank -- arg thy) >> + (fn ((q, vs), body) => fold_rev (q thy) vs body) + +fun expr thy k = Scan.first [bound thy, quantifier thy, application thy, + lift number, bv_number thy, constant] #-> apfst o add_expr k + +fun rule_name name = + (case R.rule_of_string name of + SOME r => r + | NONE => z3_exn ("unknown proof rule: " ^ quote name)) + +fun rule thy k = + bra (lift (name >> rule_name) -- bsep (lift id)) --| + ($$ ":" -- blank) -- arg thy #-> lift' (add_proof thy k) + +fun decl thy = ((this "decl" -- blank) |-- lift name --| + (blank -- this "::" -- blank) -- sort) #-> apfst o add_decl thy + +fun def st = (lift id --| (blank -- this ":=" -- blank)) st + +fun node thy = + decl thy #> pair NONE || + 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 parse_line thy l (st as (stop, line_no, cx)) = + if is_some stop then st + else + (cx, explode l) + |> handle_errors line_no (Scan.finite' Symbol.stopper (node thy)) + |> (fn (stop', (cx', _)) => (stop', line_no + 1, cx')) + +fun reconstruct ctxt assms recon output = + let + val _ = T.var_prefix <> free_prefix orelse error "Same prefixes" + + val thy = ProofContext.theory_of ctxt + in + (case fold (parse_line thy) output (NONE, 1, empty_context thy recon) of + (SOME p, _, Context {ptab, ...}) => R.prove ctxt assms ptab p + | _ => z3_exn "bad proof") + end + +end diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Tools/z3_proof_rules.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Tools/z3_proof_rules.ML Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,1319 @@ +(* Title: HOL/SMT/Tools/z3_proof_rules.ML + Author: Sascha Boehme, TU Muenchen + +Z3 proof rules and their reconstruction. +*) + +signature Z3_PROOF_RULES = +sig + (*proof rule names*) + type rule + val rule_of_string: string -> rule option + val string_of_rule: rule -> string + + (*proof reconstruction*) + type proof + val make_proof: rule -> int list -> Thm.cterm * Thm.cterm list -> proof + val prove: Proof.context -> thm list option -> proof Inttab.table -> int -> + thm + + (*setup*) + val setup: theory -> theory +end + +structure Z3_Proof_Rules: Z3_PROOF_RULES = +struct + +structure T = Z3_Proof_Terms + +fun z3_exn msg = error ("Z3 proof reconstruction: " ^ msg) + + +(* proof rule names *) + +datatype rule = TrueAxiom | Asserted | Goal | ModusPonens | Reflexivity | + Symmetry | Transitivity | TransitivityStar | Monotonicity | QuantIntro | + Distributivity | AndElim | NotOrElim | Rewrite | RewriteStar | PullQuant | + PullQuantStar | PushQuant | ElimUnusedVars | DestEqRes | QuantInst | + Hypothesis | Lemma | UnitResolution | IffTrue | IffFalse | Commutativity | + DefAxiom | IntroDef | ApplyDef | IffOeq | NnfPos | NnfNeg | NnfStar | + CnfStar | Skolemize | ModusPonensOeq | ThLemma + +val rule_names = Symtab.make [ + ("true-axiom", TrueAxiom), + ("asserted", Asserted), + ("goal", Goal), + ("mp", ModusPonens), + ("refl", Reflexivity), + ("symm", Symmetry), + ("trans", Transitivity), + ("trans*", TransitivityStar), + ("monotonicity", Monotonicity), + ("quant-intro", QuantIntro), + ("distributivity", Distributivity), + ("and-elim", AndElim), + ("not-or-elim", NotOrElim), + ("rewrite", Rewrite), + ("rewrite*", RewriteStar), + ("pull-quant", PullQuant), + ("pull-quant*", PullQuantStar), + ("push-quant", PushQuant), + ("elim-unused", ElimUnusedVars), + ("der", DestEqRes), + ("quant-inst", QuantInst), + ("hypothesis", Hypothesis), + ("lemma", Lemma), + ("unit-resolution", UnitResolution), + ("iff-true", IffTrue), + ("iff-false", IffFalse), + ("commutativity", Commutativity), + ("def-axiom", DefAxiom), + ("intro-def", IntroDef), + ("apply-def", ApplyDef), + ("iff~", IffOeq), + ("nnf-pos", NnfPos), + ("nnf-neg", NnfNeg), + ("nnf*", NnfStar), + ("cnf*", CnfStar), + ("sk", Skolemize), + ("mp~", ModusPonensOeq), + ("th-lemma", ThLemma)] + +val rule_of_string = Symtab.lookup rule_names +fun string_of_rule r = + let fun fit (s, r') = if r = r' then SOME s else NONE + in the (Symtab.get_first NONE fit rule_names) end + + +(* proof representation *) + +datatype theorem = + Thm of thm | + MetaEq of thm | + Literals of thm * thm Termtab.table + +fun thm_of (Thm thm) = thm + | thm_of (MetaEq thm) = thm COMP @{thm meta_eq_to_obj_eq} + | thm_of (Literals (thm, _)) = thm + +fun meta_eq_of (MetaEq thm) = thm + | meta_eq_of p = thm_of p COMP @{thm eq_reflection} + +datatype proof = + Unproved of { + rule: rule, + subs: int list, + prop: Thm.cterm, + vars: Thm.cterm list } | + Sequent of { + hyps: Thm.cterm list, + vars: Thm.cterm list, + thm: theorem } + +fun make_proof r ps (ct, cvs) = Unproved {rule=r, subs=ps, prop=ct, vars=cvs} + + +(* proof reconstruction utilities *) + +fun try_apply ctxt name nfs ct = + let + val trace = SMT_Solver.trace_msg ctxt I + + fun first [] = z3_exn (name ^ " failed") + | first ((n, f) :: nfs) = + (case try f ct of + SOME thm => (trace (n ^ " succeeded"); thm) + | NONE => (trace (n ^ " failed"); first nfs)) + in first nfs end + +fun prop_of thm = (case Thm.prop_of thm of @{term Trueprop} $ t => t | t => t) + +fun as_meta_eq ct = uncurry T.mk_meta_eq (Thm.dest_binop ct) + +fun by_tac' tac ct = Goal.norm_result (Goal.prove_internal [] ct (K (tac 1))) +fun by_tac tac ct = by_tac' tac (T.mk_prop ct) + +fun match_instantiate' f ct thm = + Thm.instantiate (Thm.match (f (Thm.cprop_of thm), ct)) thm +val match_instantiate = match_instantiate' I + +local + fun maybe_instantiate ct thm = + try Thm.first_order_match (Thm.cprop_of thm, ct) + |> Option.map (fn inst => Thm.instantiate inst thm) +in +fun thm_net_of thms = + let fun insert thm = Net.insert_term (K false) (Thm.prop_of thm, thm) + in fold insert thms Net.empty end + +fun first_of thms ct = get_first (maybe_instantiate ct) thms +fun net_instance net ct = first_of (Net.match_term net (Thm.term_of ct)) ct +end + +fun certify ctxt = Thm.cterm_of (ProofContext.theory_of ctxt) +fun certify_var ctxt idx T = certify ctxt (Var (("x", idx), T)) + +fun varify ctxt = + let + fun varify1 cv thm = + let + val T = Thm.typ_of (Thm.ctyp_of_term cv) + val v = certify_var ctxt (Thm.maxidx_of thm + 1) T + in SMT_Normalize.instantiate_free (cv, v) thm end + in fold varify1 end + +fun under_assumption f ct = + let val ct' = T.mk_prop ct + in Thm.implies_intr ct' (f (Thm.assume ct')) end + +fun with_conv conv prove ct = + let val eq = Thm.symmetric (conv ct) + in Thm.equal_elim eq (prove (Thm.lhs_of eq)) end + +fun list2 (x, y) = [x, y] + +fun precompose f rule = (f (Thm.cprem_of rule 1), f, rule) + +fun discharge p pq = Thm.implies_elim pq p + +fun compose (cvs, f, rule) thm = + let fun inst thm = Thm.instantiate ([], cvs ~~ f (Thm.cprop_of thm)) + in discharge thm (inst thm rule) end + +fun make_hyp_def thm = (* |- c x == t x ==> P (c x) ~~> c == t |- P (c x) *) + let + val (lhs, rhs) = Thm.dest_binop (Thm.cprem_of thm 1) + val (cf, cvs) = Drule.strip_comb lhs + val eq = T.mk_meta_eq cf (fold_rev Thm.cabs cvs rhs) + fun apply cv th = + Thm.combination th (Thm.reflexive cv) + |> Conv.fconv_rule (Conv.arg_conv (Thm.beta_conversion false)) + in ([eq], Thm.implies_elim thm (fold apply cvs (Thm.assume eq))) end + +val true_thm = @{lemma "~False" by simp} + +val is_neg = (fn @{term Not} $ _ => true | _ => false) +fun is_neg' f = (fn @{term Not} $ t => f t | _ => false) +val is_conj = (fn @{term "op &"} $ _ $ _ => true | _ => false) +val is_disj = (fn @{term "op |"} $ _ $ _ => true | _ => false) + +(** explosion of conjunctions and disjunctions **) + +local + val dest_conj_term = (fn @{term "op &"} $ t $ u => SOME (t, u) | _ => NONE) + + val negate_term = (fn @{term Not} $ t => t | t => @{term Not} $ t) + fun dest_disj_term' f = (fn + @{term Not} $ (@{term "op |"} $ t $ u) => SOME (f t, f u) + | _ => NONE) + val dest_disj_term = dest_disj_term' negate_term + + fun destc ct = list2 (Thm.dest_binop (Thm.dest_arg ct)) + val dest_conj1 = precompose destc @{thm conjunct1} + val dest_conj2 = precompose destc @{thm conjunct2} + fun dest_conj_rules t = + dest_conj_term t |> Option.map (K (dest_conj1, dest_conj2)) + + fun destd f ct = list2 (f (Thm.dest_binop (Thm.dest_arg (Thm.dest_arg ct)))) + val dn1 = apfst Thm.dest_arg and dn2 = apsnd Thm.dest_arg + val dest_disj1 = precompose (destd I) @{lemma "~(P | Q) ==> ~P" by fast} + and dest_disj2 = precompose (destd dn1) @{lemma "~(~P | Q) ==> P" by fast} + and dest_disj3 = precompose (destd I) @{lemma "~(P | Q) ==> ~Q" by fast} + and dest_disj4 = precompose (destd dn2) @{lemma "~(P | ~Q) ==> Q" by fast} + + val is_neg = (fn @{term Not} $ _ => true | _ => false) + fun dest_disj_rules t = + (case dest_disj_term' is_neg t of + SOME (true, true) => SOME (dest_disj2, dest_disj4) + | SOME (true, false) => SOME (dest_disj2, dest_disj3) + | SOME (false, true) => SOME (dest_disj1, dest_disj4) + | SOME (false, false) => SOME (dest_disj1, dest_disj3) + | NONE => NONE) + + val is_dneg = is_neg' is_neg + fun destn ct = [Thm.dest_arg (Thm.dest_arg (Thm.dest_arg ct))] + val dneg_rule = precompose destn @{thm notnotD} +in +fun exists_lit is_conj P = + let + val dest = if is_conj then dest_conj_term else dest_disj_term + fun exists t = P t orelse + (case dest t of + SOME (t1, t2) => exists t1 orelse exists t2 + | NONE => false) + in exists end + +fun explode_term is_conj keep_intermediate = + let + val dest = if is_conj then dest_conj_term else dest_disj_term + val dest_rules = if is_conj then dest_conj_rules else dest_disj_rules + fun explode1 rules t = + (case dest t of + SOME (t1, t2) => + let val (rule1, rule2) = the (dest_rules t) + in + explode1 (rule1 :: rules) t1 #> + explode1 (rule2 :: rules) t2 #> + keep_intermediate ? cons (t, rev rules) + end + | NONE => cons (t, rev rules)) + fun explode0 (@{term Not} $ (@{term Not} $ t)) = [(t, [dneg_rule])] + | explode0 t = explode1 [] t [] + in explode0 end + +fun extract_lit thm rules = fold compose rules thm + +fun explode_thm is_conj full keep_intermediate stop_lits = + let + val dest_rules = if is_conj then dest_conj_rules else dest_disj_rules + val tab = fold (Termtab.update o rpair ()) stop_lits Termtab.empty + + fun explode1 thm = + if Termtab.defined tab (prop_of thm) then cons thm + else + (case dest_rules (prop_of thm) of + SOME (rule1, rule2) => explode2 rule1 thm #> explode2 rule2 thm #> + keep_intermediate ? cons thm + | NONE => cons thm) + and explode2 dest_rule thm = + if full orelse exists_lit is_conj (Termtab.defined tab) (prop_of thm) + then explode1 (compose dest_rule thm) + else cons (compose dest_rule thm) + fun explode0 thm = + if not is_conj andalso is_dneg (prop_of thm) then [compose dneg_rule thm] + else explode1 thm [] + in explode0 end +end + +(** joining of literals to conjunctions or disjunctions **) + +local + fun precomp2 f g thm = + (f (Thm.cprem_of thm 1), g (Thm.cprem_of thm 2), f, g, thm) + fun comp2 (cv1, cv2, f, g, rule) thm1 thm2 = + let val inst = [(cv1, f (Thm.cprop_of thm1)), (cv2, g (Thm.cprop_of thm2))] + in Thm.instantiate ([], inst) rule |> discharge thm1 |> discharge thm2 end + + fun d1 ct = Thm.dest_arg ct and d2 ct = Thm.dest_arg (Thm.dest_arg ct) + + val conj_rule = precomp2 d1 d1 @{thm conjI} + fun comp_conj ((_, thm1), (_, thm2)) = comp2 conj_rule thm1 thm2 + + val disj1 = precomp2 d2 d2 @{lemma "~P ==> ~Q ==> ~(P | Q)" by fast} + val disj2 = precomp2 d2 d1 @{lemma "~P ==> Q ==> ~(P | ~Q)" by fast} + val disj3 = precomp2 d1 d2 @{lemma "P ==> ~Q ==> ~(~P | Q)" by fast} + val disj4 = precomp2 d1 d1 @{lemma "P ==> Q ==> ~(~P | ~Q)" by fast} + + fun comp_disj ((false, thm1), (false, thm2)) = comp2 disj1 thm1 thm2 + | comp_disj ((false, thm1), (true, thm2)) = comp2 disj2 thm1 thm2 + | comp_disj ((true, thm1), (false, thm2)) = comp2 disj3 thm1 thm2 + | comp_disj ((true, thm1), (true, thm2)) = comp2 disj4 thm1 thm2 + + fun dest_conj (@{term "op &"} $ t $ u) = ((false, t), (false, u)) + | dest_conj t = raise TERM ("dest_conj", [t]) + + val neg = (fn @{term Not} $ t => (true, t) | t => (false, @{term Not} $ t)) + fun dest_disj (@{term Not} $ (@{term "op |"} $ t $ u)) = (neg t, neg u) + | dest_disj t = raise TERM ("dest_disj", [t]) + + val dnegE = precompose (single o d2 o d1) @{thm notnotD} + val dnegI = precompose (single o d1) @{lemma "P ==> ~~P" by fast} + fun as_dneg f t = f (@{term Not} $ (@{term Not} $ t)) + + fun dni f = list2 o apsnd f o Thm.dest_binop o f o d1 + val negIffE = precompose (dni d1) @{lemma "~(P = (~Q)) ==> Q = P" by fast} + val negIffI = precompose (dni I) @{lemma "P = Q ==> ~(Q = (~P))" by fast} + val iff_const = @{term "op = :: bool => _"} + fun as_negIff f (@{term "op = :: bool => _"} $ t $ u) = + f (@{term Not} $ (iff_const $ u $ (@{term Not} $ t))) + | as_negIff _ _ = NONE +in +fun make_lit_tab thms = fold (Termtab.update o ` prop_of) thms Termtab.empty + +fun join is_conj tab t = + let + val comp = if is_conj then comp_conj else comp_disj + val dest = if is_conj then dest_conj else dest_disj + + val lookup_lit = Termtab.lookup tab + fun lookup_lit' t = + (case t of + @{term Not} $ (@{term Not} $ t) => (compose dnegI, lookup_lit t) + | @{term Not} $ (@{term "op = :: bool => _"} $ t $ (@{term Not} $ u)) => + (compose negIffI, lookup_lit (iff_const $ u $ t)) + | @{term Not} $ ((eq as Const (@{const_name "op ="}, _)) $ t $ u) => + let fun rewr lit = lit COMP @{thm not_sym} + in (rewr, lookup_lit (@{term Not} $ (eq $ u $ t))) end + | _ => + (case as_dneg lookup_lit t of + NONE => (compose negIffE, as_negIff lookup_lit t) + | x => (compose dnegE, x))) + fun join1 (s, t) = + (case lookup_lit t of + SOME lit => (s, lit) + | NONE => + (case lookup_lit' t of + (rewrite, SOME lit) => (s, rewrite lit) + | (_, NONE) => (s, comp (pairself join1 (dest t))))) + in snd (join1 (if is_conj then (false, t) else (true, t))) end +end + +(** proving equality of conjunctions or disjunctions **) + +fun iff_intro thm1 thm2 = thm2 COMP (thm1 COMP @{thm iffI}) + +local + val cp1 = @{lemma "(~P) = (~Q) ==> P = Q" by simp} + val cp2 = @{lemma "(~P) = Q ==> P = (~Q)" by fastsimp} + val cp3 = @{lemma "P = (~Q) ==> (~P) = Q" by simp} + val neg = Thm.capply @{cterm Not} +in +fun contrapos1 prove (ct, cu) = prove (neg ct, neg cu) COMP cp1 +fun contrapos2 prove (ct, cu) = prove (neg ct, Thm.dest_arg cu) COMP cp2 +fun contrapos3 prove (ct, cu) = prove (Thm.dest_arg ct, neg cu) COMP cp3 +end + +local + fun prove_eq l r (cl, cr) = + let + fun explode is_conj = explode_thm is_conj true (l <> r) [] + fun make_tab is_conj thm = make_lit_tab (true_thm :: explode is_conj thm) + fun prove is_conj ct tab = join is_conj tab (Thm.term_of ct) + + val thm1 = under_assumption (prove r cr o make_tab l) cl + val thm2 = under_assumption (prove l cl o make_tab r) cr + in iff_intro thm1 thm2 end + + datatype conj_disj = CONJ | DISJ | NCON | NDIS + fun kind_of t = + if is_conj t then CONJ + else if is_disj t then DISJ + else if is_neg' is_conj t then NCON + else if is_neg' is_disj t then NDIS + else CONJ (*allows to prove equalities with single literals on each side*) +in +fun prove_conj_disj_eq ct = + let val cp = Thm.dest_binop ct + in + (case pairself (kind_of o Thm.term_of) cp of + (CONJ, CONJ) => prove_eq true true cp + | (CONJ, NDIS) => prove_eq true false cp + | (DISJ, DISJ) => contrapos1 (prove_eq false false) cp + | (DISJ, NCON) => contrapos2 (prove_eq false true) cp + | (NCON, NCON) => contrapos1 (prove_eq true true) cp + | (NCON, DISJ) => contrapos3 (prove_eq true false) cp + | (NDIS, NDIS) => prove_eq false false cp + | (NDIS, CONJ) => prove_eq false true cp) + end +end + +(** unfolding of distinct **) + +local + val distinct1 = @{lemma "distinct [] == ~False" by simp} + val distinct2 = @{lemma "distinct [x] == ~False" by simp} + val distinct3 = @{lemma "distinct (x # xs) == x ~: set xs & distinct xs" + by simp} + + val set1 = @{lemma "x ~: set [] == ~False" by simp} + val set2 = @{lemma "x ~: set [y] == x ~= y" by simp} + val set3 = @{lemma "x ~: set (y # ys) == x ~= y & x ~: set ys" by simp} + + fun binop_conv cv1 cv2 = Conv.combination_conv (Conv.arg_conv cv1) cv2 + + fun unfold_conv rule1 rule2 rule3 sub_conv = + let + fun uconv ct = + (Conv.rewr_conv rule1 else_conv + Conv.rewr_conv rule2 else_conv + (Conv.rewr_conv rule3 then_conv binop_conv sub_conv uconv)) ct + in uconv end + + val set_conv = unfold_conv set1 set2 set3 Conv.all_conv +in +val unfold_distinct_conv = unfold_conv distinct1 distinct2 distinct3 set_conv +end + + +(* core proof rules *) + +datatype assms = Some of thm list | Many of thm Net.net + +val true_false = @{lemma "True == ~ False" by simp} + +local + val TT_eq = @{lemma "(P = (~False)) == P" by simp} + val remove_trigger = @{lemma "trigger t p == p" + by (rule eq_reflection, rule trigger_def)} + val remove_iff = @{lemma "p iff q == p = q" + by (rule eq_reflection, rule iff_def)} + + fun with_context simpset ctxt = Simplifier.context ctxt simpset + + val prep_ss = with_context (Simplifier.empty_ss addsimps + [@{thm Let_def}, remove_trigger, remove_iff, true_false, TT_eq]) + + val TT_eq_conv = Conv.rewr_conv TT_eq + val norm_conv = More_Conv.bottom_conv (K (Conv.try_conv TT_eq_conv)) + + val threshold = 10 + + val lookup = (fn Some thms => first_of thms | Many net => net_instance net) + fun lookup_assm ctxt assms ct = + (case lookup assms ct of + SOME thm => thm + | _ => z3_exn ("not asserted: " ^ + quote (Syntax.string_of_term ctxt (Thm.term_of ct)))) +in +fun prepare_assms ctxt assms = + let + val rewrite = Conv.fconv_rule (Simplifier.rewrite (prep_ss ctxt)) + val thms = map rewrite assms + in if length assms < threshold then Some thms else Many (thm_net_of thms) end + +fun asserted _ NONE ct = Thm (Thm.assume (T.mk_prop ct)) + | asserted ctxt (SOME assms) ct = + Thm (with_conv (norm_conv ctxt) (lookup_assm ctxt assms) (T.mk_prop ct)) +end + + +(** P ==> P = Q ==> Q or P ==> P --> Q ==> Q **) +local + val meta_iffD1 = @{lemma "P == Q ==> P ==> (Q::bool)" by simp} + val meta_iffD1_c = precompose (list2 o Thm.dest_binop) meta_iffD1 + + val iffD1_c = precompose (list2 o Thm.dest_binop o Thm.dest_arg) @{thm iffD1} + val mp_c = precompose (list2 o Thm.dest_binop o Thm.dest_arg) @{thm mp} +in +fun mp (MetaEq thm) p = Thm (Thm.implies_elim (compose meta_iffD1_c thm) p) + | mp p_q p = + let + val pq = thm_of p_q + val thm = compose iffD1_c pq handle THM _ => compose mp_c pq + in Thm (Thm.implies_elim thm p) end +end + + +(** and_elim: P1 & ... & Pn ==> Pi **) +(** not_or_elim: ~(P1 | ... | Pn) ==> ~Pi **) +local + fun get_lit conj t (l, thm) = + let val is_sublit_of = exists_lit conj (fn u => u aconv t) + in if is_sublit_of (prop_of thm) then SOME (l, thm) else NONE end + + fun derive conj t lits idx ptab = + let + val (l, lit) = the (Termtab.get_first NONE (get_lit conj t) lits) + val ls = explode_thm conj false false [t] lit + val lits' = fold (Termtab.update o ` prop_of) ls (Termtab.delete l lits) + fun upd (Sequent {hyps, vars, thm}) = + Sequent {hyps=hyps, vars=vars, thm = Literals (thm_of thm, lits')} + | upd p = p + in (the (Termtab.lookup lits' t), Inttab.map_entry idx upd ptab) end + + val mk_tab = make_lit_tab o single + val literals_of = (fn Literals (_, lits) => lits | p => mk_tab (thm_of p)) + fun lit_elim conj (p, idx) ct ptab = + let val lits = literals_of p + in + (case Termtab.lookup lits (Thm.term_of ct) of + SOME lit => (Thm lit, ptab) + | NONE => apfst Thm (derive conj (Thm.term_of ct) lits idx ptab)) + end +in +val and_elim = lit_elim true +val not_or_elim = lit_elim false +end + + +(** P1 ... Pn |- False ==> |- ~P1 | ... | ~Pn **) +local + fun step lit thm = + Thm.implies_elim (Thm.implies_intr (Thm.cprop_of lit) thm) lit + val explode_disj = explode_thm false false false + fun intro hyps thm th = fold step (explode_disj hyps th) thm + + fun dest_ccontr ct = [Thm.dest_arg (Thm.dest_arg (Thm.dest_arg1 ct))] + val ccontr = precompose dest_ccontr @{thm ccontr} +in +fun lemma thm ct = + let + val cu = Thm.capply @{cterm Not} ct + val hyps = map_filter (try HOLogic.dest_Trueprop) (#hyps (Thm.rep_thm thm)) + in Thm (compose ccontr (under_assumption (intro hyps thm) cu)) end +end + + +(** \/{P1, ..., Pn, Q1, ..., Qn} & ~P1 & ... & ~Pn ==> \/{Q1, ..., Qn} **) +local + val explode_disj = explode_thm false true false and join_disj = join false + fun unit thm thms th = + let val t = @{term Not} $ prop_of thm and ts = map prop_of thms + in join_disj (make_lit_tab (thms @ explode_disj ts th)) t end + + fun dest_arg2 ct = Thm.dest_arg (Thm.dest_arg ct) + fun dest ct = list2 (pairself dest_arg2 (Thm.dest_binop ct)) + val contrapos = precompose dest @{lemma "(~P ==> ~Q) ==> Q ==> P" by fast} +in +fun unit_resolution thm thms ct = + under_assumption (unit thm thms) (Thm.capply @{cterm Not} ct) + |> Thm o discharge thm o compose contrapos +end + + +local + val iff1 = @{lemma "P ==> P == (~ False)" by simp} + val iff2 = @{lemma "~P ==> P == False" by simp} +in +fun iff_true thm = MetaEq (thm COMP iff1) +fun iff_false thm = MetaEq (thm COMP iff2) +end + + +(** distributivity of | over & **) +val distributivity = Thm o by_tac (Classical.fast_tac HOL_cs) + + +(** Tseitin-like axioms **) +local + val disjI1 = @{lemma "(P ==> Q) ==> ~P | Q" by fast} + val disjI2 = @{lemma "(~P ==> Q) ==> P | Q" by fast} + val disjI3 = @{lemma "(~Q ==> P) ==> P | Q" by fast} + val disjI4 = @{lemma "(Q ==> P) ==> P | ~Q" by fast} + + fun prove' conj1 conj2 ct2 thm = + let val tab = + make_lit_tab (true_thm :: explode_thm conj1 true (conj1 <> conj2) [] thm) + in join conj2 tab (Thm.term_of ct2) end + + fun prove rule (ct1, conj1) (ct2, conj2) = + under_assumption (prove' conj1 conj2 ct2) ct1 COMP rule + + fun prove_def_axiom ct = + let val (ct1, ct2) = Thm.dest_binop ct + in + (case Thm.term_of ct1 of + @{term Not} $ (@{term "op &"} $ _ $ _) => + prove disjI1 (Thm.dest_arg ct1, true) (ct2, true) + | @{term "op &"} $ _ $ _ => + prove disjI3 (Thm.capply @{cterm Not} ct2, false) (ct1, true) + | @{term Not} $ (@{term "op |"} $ _ $ _) => + prove disjI3 (Thm.capply @{cterm Not} ct2, false) (ct1, false) + | @{term "op |"} $ _ $ _ => + prove disjI2 (Thm.capply @{cterm Not} ct1, false) (ct2, true) + | Const (@{const_name distinct}, _) $ _ => + let + fun dis_conv cv = Conv.arg_conv (Conv.arg1_conv cv) + fun prv cu = + let val (cu1, cu2) = Thm.dest_binop (Thm.dest_arg cu) + in prove disjI4 (Thm.dest_arg cu2, true) (cu1, true) end + in with_conv (dis_conv unfold_distinct_conv) prv (T.mk_prop ct) end + | @{term Not} $ (Const (@{const_name distinct}, _) $ _) => + let + fun dis_conv cv = Conv.arg_conv (Conv.arg1_conv (Conv.arg_conv cv)) + fun prv cu = + let val (cu1, cu2) = Thm.dest_binop (Thm.dest_arg cu) + in prove disjI1 (Thm.dest_arg cu1, true) (cu2, true) end + in with_conv (dis_conv unfold_distinct_conv) prv (T.mk_prop ct) end + | _ => raise CTERM ("prove_def_axiom", [ct])) + end + + val ifI = @{lemma "(P ==> Q1) ==> (~P ==> Q2) ==> if P then Q1 else Q2" + by simp} + val ifE = @{lemma + "(if P then Q1 else Q2) ==> (P --> Q1 ==> ~P --> Q2 ==> R) ==> R" by simp} + val claset = HOL_cs addIs [ifI] addEs [ifE] +in +fun def_axiom ctxt ct = + Thm (try_apply ctxt "def_axiom" [ + ("conj/disj", prove_def_axiom), + ("fast", by_tac (Classical.fast_tac claset)), + ("simp+fast", by_tac (Simplifier.simp_tac HOL_ss THEN_ALL_NEW + Classical.fast_tac claset))] ct) +end + + +(** local definitions **) +local + val intro_rules = [ + @{lemma "n == P ==> (~n | P) & (n | ~P)" by simp}, + @{lemma "n == (if P then s else t) ==> (~P | n = s) & (P | n = t)" + by simp}, + @{lemma "n == P ==> n = P" by (rule meta_eq_to_obj_eq)} ] + + val apply_rules = [ + @{lemma "(~n | P) & (n | ~P) ==> P == n" by (atomize(full)) fast}, + @{lemma "(~P | n = s) & (P | n = t) ==> (if P then s else t) == n" + by (atomize(full)) fastsimp} ] + + val inst_rule = match_instantiate' Thm.dest_arg + + fun apply_rule ct = + (case get_first (try (inst_rule (T.mk_prop ct))) intro_rules of + SOME thm => thm + | NONE => raise CTERM ("intro_def", [ct])) +in +fun intro_def ct = apsnd Thm (make_hyp_def (apply_rule ct)) + +fun apply_def thm = + get_first (try (fn rule => MetaEq (thm COMP rule))) apply_rules + |> the_default (Thm thm) +end + + +local + val quant_rules1 = ([ + @{lemma "(!!x. P x == Q) ==> ALL x. P x == Q" by simp}, + @{lemma "(!!x. P x == Q) ==> EX x. P x == Q" by simp}], [ + @{lemma "(!!x. P x == Q x) ==> ALL x. P x == ALL x. Q x" by simp}, + @{lemma "(!!x. P x == Q x) ==> EX x. P x == EX x. Q x" by simp}]) + + val quant_rules2 = ([ + @{lemma "(!!x. ~P x == Q) ==> ~(ALL x. P x) == Q" by simp}, + @{lemma "(!!x. ~P x == Q) ==> ~(EX x. P x) == Q" by simp}], [ + @{lemma "(!!x. ~P x == Q x) ==> ~(ALL x. P x) == EX x. Q x" by simp}, + @{lemma "(!!x. ~P x == Q x) ==> ~(EX x. P x) == ALL x. Q x" by simp}]) + + fun nnf_quant_tac thm (qs as (qs1, qs2)) i st = ( + Tactic.rtac thm ORELSE' + (Tactic.match_tac qs1 THEN' nnf_quant_tac thm qs) ORELSE' + (Tactic.match_tac qs2 THEN' nnf_quant_tac thm qs)) i st + + fun nnf_quant ctxt qs (p, (vars, _)) ct = + as_meta_eq ct + |> by_tac' (nnf_quant_tac (varify ctxt vars (meta_eq_of p)) qs) + + val nnf_rules = thm_net_of [@{thm not_not}] + + fun prove_nnf ctxt = + try_apply ctxt "nnf" [ + ("conj/disj", prove_conj_disj_eq o Thm.dest_arg), + ("rule", the o net_instance nnf_rules), + ("tactic", by_tac' (Classical.best_tac HOL_cs))] +in +fun nnf ctxt ps ct = + (case Thm.term_of ct of + _ $ (l as Const _ $ Abs _) $ (r as Const _ $ Abs _) => + if l aconv r then MetaEq (Thm.reflexive (Thm.dest_arg ct)) + else MetaEq (nnf_quant ctxt quant_rules1 (hd ps) ct) + | _ $ (@{term Not} $ (Const _ $ Abs _)) $ (Const _ $ Abs _) => + MetaEq (nnf_quant ctxt quant_rules2 (hd ps) ct) + | _ => + let + val eqs = map (Thm.symmetric o meta_eq_of o fst) ps + val nnf_rewr_conv = Conv.arg_conv (Conv.arg_conv + (More_Conv.top_sweep_conv (K (More_Conv.rewrs_conv eqs)) ctxt)) + in Thm (with_conv nnf_rewr_conv (prove_nnf ctxt) (T.mk_prop ct)) end) +end + + +(* equality proof rules *) + +(** t = t **) +fun refl ct = MetaEq (Thm.reflexive (Thm.dest_arg ct)) + + +(** s = t ==> t = s **) +local + val symm_rule = @{lemma "s = t ==> t == s" by simp} +in +fun symm (MetaEq thm) = MetaEq (Thm.symmetric thm) + | symm p = MetaEq (thm_of p COMP symm_rule) +end + + +(** s = t ==> t = u ==> s = u **) +local + val trans_rule = @{lemma "s = t ==> t = u ==> s == u" by simp} +in +fun trans (MetaEq thm) q = MetaEq (Thm.transitive thm (meta_eq_of q)) + | trans p (MetaEq thm) = MetaEq (Thm.transitive (meta_eq_of p) thm) + | trans p q = MetaEq (thm_of q COMP (thm_of p COMP trans_rule)) +end + + +(** t1 = s1 & ... & tn = sn ==> f t1 ... tn = f s1 .. sn + (reflexive antecendents are droppped) **) +local + exception MONO + + fun prove_refl (ct, _) = Thm.reflexive ct + fun prove_comb f g cp = + let val ((ct1, ct2), (cu1, cu2)) = pairself Thm.dest_comb cp + in Thm.combination (f (ct1, cu1)) (g (ct2, cu2)) end + fun prove_arg f = prove_comb prove_refl f + + fun prove f cp = prove_comb (prove f) f cp handle CTERM _ => prove_refl cp + + fun prove_nary is_comb f = + let + fun prove (cp as (ct, _)) = f cp handle MONO => + if is_comb (Thm.term_of ct) + then prove_comb (prove_arg prove) prove cp + else prove_refl cp + in prove end + + fun prove_list f n cp = + if n = 0 then prove_refl cp + else prove_comb (prove_arg f) (prove_list f (n-1)) cp + + fun with_length f (cp as (cl, _)) = + f (length (HOLogic.dest_list (Thm.term_of cl))) cp + + fun prove_distinct f = prove_arg (with_length (prove_list f)) + + fun prove_eq exn lookup cp = + (case lookup (Logic.mk_equals (pairself Thm.term_of cp)) of + SOME eq => eq + | NONE => if exn then raise MONO else prove_refl cp) + val prove_eq_exn = prove_eq true and prove_eq_safe = prove_eq false + + fun mono f (cp as (cl, _)) = + (case Term.head_of (Thm.term_of cl) of + @{term "op &"} => prove_nary is_conj (prove_eq_exn f) + | @{term "op |"} => prove_nary is_disj (prove_eq_exn f) + | Const (@{const_name distinct}, _) => prove_distinct (prove_eq_safe f) + | _ => prove (prove_eq_safe f)) cp +in +fun monotonicity eqs ct = + let + val tab = map (` Thm.prop_of o meta_eq_of) eqs + val lookup = AList.lookup (op aconv) tab + val cp = Thm.dest_binop ct + in MetaEq (prove_eq_exn lookup cp handle MONO => mono lookup cp) end +end + + +(** f a b = f b a **) +local + val rule = @{lemma "a = b == b = a" by (atomize(full)) (rule eq_commute)} +in +fun commutativity ct = MetaEq (match_instantiate (as_meta_eq ct) rule) +end + + +(* quantifier proof rules *) + +(** P ?x = Q ?x ==> (ALL x. P x) = (ALL x. Q x) + P ?x = Q ?x ==> (EX x. P x) = (EX x. Q x) **) +local + val rules = [ + @{lemma "(!!x. P x == Q x) ==> (EX x. P x) == (EX x. Q x)" by simp}, + @{lemma "(!!x. P x == Q x) ==> (ALL x. P x) == (ALL x. Q x)" by simp}] +in +fun quant_intro ctxt (p, (vars, _)) ct = + let + val rules' = varify ctxt vars (meta_eq_of p) :: rules + val cu = as_meta_eq ct + in MetaEq (by_tac' (REPEAT_ALL_NEW (Tactic.match_tac rules')) cu) end +end + + +(** |- ((ALL x. P x) | Q) = (ALL x. P x | Q) **) +val pull_quant = + Thm o by_tac (Tactic.rtac @{thm refl} ORELSE' Simplifier.simp_tac HOL_ss) + + +(** |- (ALL x. P x & Q x) = ((ALL x. P x) & (ALL x. Q x)) **) +val push_quant = + Thm o by_tac (Tactic.rtac @{thm refl} ORELSE' Simplifier.simp_tac HOL_ss) + + +(** + |- (ALL x1 ... xn y1 ... yn. P x1 ... xn) = (ALL x1 ... xn. P x1 ... xn) +**) +local + val elim_all = @{lemma "ALL x. P == P" by simp} + val elim_ex = @{lemma "EX x. P == P" by simp} + + val rule = (fn @{const_name All} => elim_all | _ => elim_ex) + + fun collect xs tp = + if (op aconv) tp then rev xs + else + (case tp of + (Const (q, _) $ Abs (_, _, l), r' as Const _ $ Abs (_, _, r)) => + if l aconv r then rev xs + else if Term.loose_bvar1 (l, 0) then collect (NONE :: xs) (l, r) + else collect (SOME (rule q) :: xs) (Term.incr_bv (~1, 0, l), r') + | (Const (q, _) $ Abs (_, _, l), r) => + collect (SOME (rule q) :: xs) (Term.incr_bv (~1, 0, l), r) + | (l, r) => raise TERM ("elim_unused", [l, r])) + + fun elim _ [] ct = Conv.all_conv ct + | elim ctxt (x::xs) ct = + (case x of + SOME rule => Conv.rewr_conv rule then_conv elim ctxt xs + | _ => Conv.arg_conv (Conv.abs_conv (fn (_,cx) => elim cx xs) ctxt)) ct +in +fun elim_unused_vars ctxt ct = + let val (lhs, rhs) = Thm.dest_binop ct + in MetaEq (elim ctxt (collect [] (Thm.term_of lhs, Thm.term_of rhs)) lhs) end +end + + +(** + |- (ALL x1 ... xn. ~(x1 = t1 & ... xn = tn) | P x1 ... xn) = P t1 ... tn +**) +val dest_eq_res = Thm o by_tac (Simplifier.simp_tac HOL_ss) + + +(** |- ~(ALL x1...xn. P x1...xn) | P a1...an **) +local + val rule = @{lemma "~ P x | Q ==> ~(ALL x. P x) | Q" by fast} +in +val quant_inst = Thm o by_tac ( + REPEAT_ALL_NEW (Tactic.match_tac [rule]) + THEN' Tactic.rtac @{thm excluded_middle}) +end + + +(** c = SOME x. P x |- (EX x. P x) = P c + c = SOME x. ~ P x |- ~(ALL x. P x) = ~ P c **) +local + val elim_ex = @{lemma "EX x. P == P" by simp} + val elim_all = @{lemma "~ (ALL x. P) == ~P" by simp} + val sk_ex = @{lemma "c == SOME x. P x ==> EX x. P x == P c" + by simp (intro eq_reflection some_eq_ex[symmetric])} + val sk_all = @{lemma "c == SOME x. ~ P x ==> ~(ALL x. P x) == ~ P c" + by (simp only: not_all) (intro eq_reflection some_eq_ex[symmetric])} + val sk_ex_rule = ((sk_ex, I), elim_ex) + and sk_all_rule = ((sk_all, Thm.dest_arg), elim_all) + + fun dest f sk_rule = + Thm.dest_comb (f (Thm.dest_arg (Thm.dest_arg (Thm.cprop_of sk_rule)))) + fun type_of f sk_rule = Thm.ctyp_of_term (snd (dest f sk_rule)) + fun inst_sk (sk_rule, f) p c = + Thm.instantiate ([(type_of f sk_rule, Thm.ctyp_of_term c)], []) sk_rule + |> (fn sk' => Thm.instantiate ([], (list2 (dest f sk') ~~ [p, c])) sk') + |> Conv.fconv_rule (Thm.beta_conversion true) + + fun kind (Const (q as @{const_name Ex}, _) $ _) = (sk_ex_rule, q, I, I) + | kind (@{term Not} $ (Const (q as @{const_name All}, _) $ _)) = + (sk_all_rule, q, Thm.dest_arg, Thm.capply @{cterm Not}) + | kind _ = z3_exn "skolemize: no quantifier" + + fun bodies_of ctxt ct = + let + val (rule, q, dest, make) = kind (Thm.term_of ct) + + fun inst_abs idx T cbs ct = + let + val cv = certify_var ctxt idx T + val cu = Drule.beta_conv (Thm.dest_arg ct) cv + in dest_body (idx + 1) ((cv, Thm.dest_arg ct) :: cbs) cu end + and dest_body idx cbs ct = + (case Thm.term_of ct of + Const (qname, _) $ Abs (_, T, _) => + if q = qname then inst_abs idx T cbs ct else (make ct, rev cbs) + | _ => (make ct, rev cbs)) + in (rule, dest_body (#maxidx (Thm.rep_cterm ct) + 1) [] (dest ct)) end + + fun transitive f thm = Thm.transitive thm (f (Thm.rhs_of thm)) + + fun sk_step (rule, elim) (cv, mct, cb) (is, thm) = + (case mct of + SOME ct => + make_hyp_def (inst_sk rule (Thm.instantiate_cterm ([], is) cb) ct) + |> apsnd (pair ((cv, ct) :: is) o Thm.transitive thm) + | NONE => ([], (is, transitive (Conv.rewr_conv elim) thm))) +in +fun skolemize ctxt ct = + let + val (lhs, rhs) = Thm.dest_binop ct + val (rule, (cu, cbs)) = bodies_of ctxt lhs + val ctab = snd (Thm.first_order_match (cu, rhs)) + fun lookup_var (cv, cb) = (cv, AList.lookup (op aconvc) ctab cv, cb) + in + ([], Thm.reflexive lhs) + |> fold_map (sk_step rule) (map lookup_var cbs) + |> apfst (rev o flat) o apsnd (MetaEq o snd) + end +end + + +(* theory proof rules *) + +(** prove linear arithmetic problems via generalization **) +local + val is_numeral = can HOLogic.dest_number + fun is_number (Const (@{const_name uminus}, _) $ t) = is_numeral t + | is_number t = is_numeral t + + local + val int_distrib = @{lemma "n * (x + y) == n * x + n * (y::int)" + by (simp add: int_distrib)} + val real_distrib = @{lemma "n * (x + y) == n * x + n * (y::real)" + by (simp add: mult.add_right)} + val int_assoc = @{lemma "n * (m * x) == (n * m) * (x::int)" by linarith} + val real_assoc = @{lemma "n * (m * x) == (n * m) * (x::real)" by linarith} + + val number_of_cong = @{lemma + "number_of x * number_of y == (number_of (x * y) :: int)" + "number_of x * number_of y == (number_of (x * y) :: real)" + by simp_all} + val reduce_ss = HOL_ss addsimps @{thms mult_bin_simps} + addsimps @{thms add_bin_simps} addsimps @{thms succ_bin_simps} + addsimps @{thms minus_bin_simps} addsimps @{thms pred_bin_simps} + addsimps number_of_cong + val reduce_conv = Simplifier.rewrite reduce_ss + + fun apply_conv distrib assoc u ct = + ((case u of + Const (@{const_name times}, _) $ n $ _ => + if is_number n + then Conv.rewr_conv assoc then_conv Conv.arg1_conv reduce_conv + else Conv.rewr_conv distrib + | _ => Conv.rewr_conv distrib) + then_conv Conv.binop_conv (Conv.try_conv distrib_conv)) ct + + and distrib_conv ct = + (case Thm.term_of ct of + @{term "op * :: int => _"} $ n $ u => + if is_number n then apply_conv int_distrib int_assoc u + else Conv.no_conv + | @{term "op * :: real => _"} $ n $ u => + if is_number n then apply_conv real_distrib real_assoc u + else Conv.no_conv + | _ => Conv.no_conv) ct + in + val all_distrib_conv = More_Conv.top_sweep_conv (K distrib_conv) + end + + local + fun make_ctxt ctxt = (ctxt, Ctermtab.empty, 1) + fun fresh ct (cx as (ctxt, tab, idx)) = + (case Ctermtab.lookup tab ct of + SOME cv => (cv, cx) + | NONE => + let val cv = certify_var ctxt idx (#T (Thm.rep_cterm ct)) + in (cv, (ctxt, Ctermtab.update (ct, cv) tab, idx + 1)) end) + + fun fold_map_op f ct = + let val (cf, cu) = Thm.dest_comb ct + in f cu #>> Thm.capply cf end + + fun fold_map_binop f1 f2 ct = + let val ((cf, cu1), cu2) = apfst Thm.dest_comb (Thm.dest_comb ct) + in f1 cu1 ##>> f2 cu2 #>> uncurry (Thm.mk_binop cf) end + + fun mult f1 f2 ct t u = + if is_number t + then if is_number u then pair ct else fold_map_binop f1 f2 ct + else fresh ct + + fun poly ct = + (case Thm.term_of ct of + Const (@{const_name plus}, _) $ _ $ _ => fold_map_binop poly poly ct + | Const (@{const_name minus}, _) $ _ $ _ => fold_map_binop poly poly ct + | Const (@{const_name times}, _) $ t $ u => mult pair fresh ct t u + | Const (@{const_name div}, _) $ t $ u => mult fresh pair ct t u + | Const (@{const_name mod}, _) $ t $ u => mult fresh pair ct t u + | t => if is_number t then pair ct else fresh ct) + + val ineq_ops = [@{term "op = :: int => _"}, @{term "op < :: int => _"}, + @{term "op <= :: int => _"}, @{term "op = :: real => _"}, + @{term "op < :: real => _"}, @{term "op <= :: real => _"}] + fun ineq ct = + (case Thm.term_of ct of + t $ _ $ _ => + if member (op =) ineq_ops t then fold_map_binop poly poly ct + else raise CTERM ("arith_lemma", [ct]) + | @{term Not} $ (t $ _ $ _) => + if member (op =) ineq_ops t + then fold_map_op (fold_map_binop poly poly) ct + else raise CTERM ("arith_lemma", [ct]) + | _ => raise CTERM ("arith_lemma", [ct])) + + fun conj ct = + (case Thm.term_of ct of + @{term "op &"} $ _ $ _ => fold_map_binop conj conj ct + | @{term "~False"} => pair ct + | _ => ineq ct) + + fun disj ct = + (case Thm.term_of ct of + @{term "op |"} $ _ $ _ => fold_map_binop disj disj ct + | @{term False} => pair ct + | _ => conj ct) + in + fun prove_arith ctxt thms ct = + let + val (goal, (_, tab, _)) = + make_ctxt ctxt + |> fold_map (fold_map_op ineq o Thm.cprop_of) thms + ||>> fold_map_op disj ct + |>> uncurry (fold_rev (Thm.mk_binop @{cterm "op ==>"})) + in + Goal.prove_internal [] goal (fn _ => Arith_Data.arith_tac ctxt 1) + |> Thm.instantiate ([], map swap (Ctermtab.dest tab)) + |> fold (fn th1 => fn th2 => Thm.implies_elim th2 th1) thms + end + end +in +fun arith_lemma ctxt thms ct = + let val thms' = map (Conv.fconv_rule (all_distrib_conv ctxt)) thms + in with_conv (all_distrib_conv ctxt) (prove_arith ctxt thms') ct end +end + +(** theory lemmas: linear arithmetic, arrays **) +local + val array_ss = HOL_ss addsimps @{thms array_rules} + fun array_tac thms = + Tactic.cut_facts_tac thms + THEN' Simplifier.asm_full_simp_tac array_ss + + fun full_arith_tac ctxt thms = + Tactic.cut_facts_tac thms + THEN' Arith_Data.arith_tac ctxt +in +fun th_lemma ctxt thms ct = + Thm (try_apply ctxt "th-lemma" [ + ("abstract arith", arith_lemma ctxt thms), + ("array", by_tac' (array_tac thms)), + ("full arith", by_tac' (full_arith_tac ctxt thms))] (T.mk_prop ct)) +end + + +(** rewriting: prove equalities: + * ACI of conjunction/disjunction + * contradiction, excluded middle + * logical rewriting rules (for negation, implication, equivalence, + distinct) + * normal forms for polynoms (integer/real arithmetic) + * quantifier elimination over linear arithmetic + * ... ? **) +structure Z3_Rewrite_Rules = +struct + val name = "z3_rewrite" + val descr = "Z3 rewrite rules used in proof reconstruction" + + structure Data = GenericDataFun + ( + type T = thm Net.net + val empty = Net.empty + val extend = I + fun merge _ = Net.merge Thm.eq_thm_prop + ) + val get = Data.get o Context.Proof + + val entry = ` Thm.prop_of o Simplifier.rewrite_rule [true_false] + val eq = Thm.eq_thm_prop + val ins = Net.insert_term eq o entry and del = Net.delete_term eq o entry + fun insert thm net = ins thm net handle Net.INSERT => net + fun delete thm net = del thm net handle Net.DELETE => net + + val add = Thm.declaration_attribute (Data.map o insert) + val del = Thm.declaration_attribute (Data.map o delete) + val setup = Attrib.setup (Binding.name name) (Attrib.add_del add del) descr +end + +local + val contra_rule = @{lemma "P ==> ~P ==> False" by (rule notE)} + fun contra_left conj thm = + let + fun make_tab xs = fold Termtab.update xs Termtab.empty + val tab = make_tab (explode_term conj true (prop_of thm)) + fun pnlits (t, nrs) = + (case t of + @{term Not} $ u => Termtab.lookup tab u |> Option.map (pair nrs) + | _ => NONE) + in + (case Termtab.lookup tab @{term False} of + SOME rs => extract_lit thm rs + | NONE => + pairself (extract_lit thm) (the (Termtab.get_first NONE pnlits tab)) + |> (fn (nlit, plit) => nlit COMP (plit COMP contra_rule))) + end + val falseE_v = Thm.dest_arg (Thm.dest_arg (Thm.cprop_of @{thm FalseE})) + fun contra_right ct = Thm.instantiate ([], [(falseE_v, ct)]) @{thm FalseE} + fun contradiction conj ct = + iff_intro (under_assumption (contra_left conj) ct) (contra_right ct) + + fun conj_disj ct = + let + val cp as (cl, _) = Thm.dest_binop (Thm.dest_arg ct) + val (lhs, rhs) = pairself Thm.term_of cp + in + if is_conj lhs andalso rhs = @{term False} + then contradiction true cl + else if is_disj lhs andalso rhs = @{term "~False"} + then contrapos2 (contradiction false o fst) cp + else prove_conj_disj_eq (Thm.dest_arg ct) + end + + val distinct = + let val try_unfold = Conv.try_conv unfold_distinct_conv + in with_conv (Conv.arg_conv (Conv.binop_conv try_unfold)) conj_disj end + + val nnf_neg_rule = @{lemma "~~P == P" by fastsimp} + val nnf_cd_rules = @{lemma "~(P | Q) == ~P & ~Q" "~(P & Q) == ~P | ~Q" + by fastsimp+} + + fun nnf_conv ct = Conv.try_conv ( + (Conv.rewr_conv nnf_neg_rule then_conv nnf_conv) else_conv + (More_Conv.rewrs_conv nnf_cd_rules then_conv Conv.binop_conv nnf_conv)) ct + val iffI_rule = @{lemma "~P | Q ==> ~Q | P ==> P = Q" by fast} + fun arith_tac ctxt = CSUBGOAL (fn (goal, i) => + let val prep_then = with_conv (Conv.arg_conv (Conv.binop_conv nnf_conv)) + in Tactic.rtac (prep_then (arith_lemma ctxt []) goal) i end) + fun arith_eq_tac ctxt = + Tactic.rtac iffI_rule THEN_ALL_NEW arith_tac ctxt + ORELSE' arith_tac ctxt + + val simpset = HOL_ss addsimps @{thms array_rules} + addsimps @{thms ring_distribs} addsimps @{thms field_eq_simps} + addsimps @{thms arith_special} addsimps @{thms less_bin_simps} + addsimps @{thms le_bin_simps} addsimps @{thms eq_bin_simps} + addsimps @{thms add_bin_simps} addsimps @{thms succ_bin_simps} + addsimps @{thms minus_bin_simps} addsimps @{thms pred_bin_simps} + addsimps @{thms mult_bin_simps} addsimps @{thms iszero_simps} + addsimprocs [ + Simplifier.simproc @{theory} "fast_int_arith" [ + "(m::int) < n", "(m::int) <= n", "(m::int) = n"] (K Lin_Arith.simproc), + Simplifier.simproc @{theory} "fast_real_arith" [ + "(m::real) < n", "(m::real) <= n", "(m::real) = n"] + (K Lin_Arith.simproc)] + val simp_tac = CHANGED o Simplifier.simp_tac simpset + ORELSE' Classical.best_tac HOL_cs +in +fun rewrite ctxt thms ct = + let val rules_net = Z3_Rewrite_Rules.get ctxt + in + Thm (try_apply ctxt "rewrite" [ + ("schematic rule", the o net_instance rules_net), + ("conj/disj", conj_disj), + ("distinct", distinct), + ("arith", by_tac' (arith_eq_tac ctxt)), + ("classical", by_tac' (Classical.best_tac HOL_cs)), + ("simp", by_tac' simp_tac), + ("full arith", by_tac' (Arith_Data.arith_tac ctxt))] (T.mk_prop ct)) + end +end + + +(* tracing and debugging *) + +fun check idx r ct ((_, p), _) = + let val thm = thm_of p |> tap (Thm.join_proofs o single) + in + if (Thm.cprop_of thm) aconvc (T.mk_prop ct) then () + else z3_exn ("proof step failed: " ^ quote (string_of_rule r) ^ + " (#" ^ string_of_int idx ^ ")") + end + +local + fun trace_before ctxt idx (r, ps, ct) = + Pretty.string_of ( + Pretty.big_list ("#" ^ string_of_int idx ^ ": " ^ string_of_rule r) [ + Pretty.big_list "assumptions:" + (map (Display.pretty_thm ctxt o thm_of o fst) ps), + Pretty.block [Pretty.str "goal: ", + Syntax.pretty_term ctxt (Thm.term_of ct)]]) + + fun trace_after ctxt ((_, p), _) = Pretty.string_of (Pretty.block + [Pretty.str "result: ", Display.pretty_thm ctxt (thm_of p)]) +in +fun trace_rule ctxt idx prove r ps ct ptab = + let + val _ = SMT_Solver.trace_msg ctxt (trace_before ctxt idx) (r, ps, ct) + val result = prove r ps ct ptab + val _ = SMT_Solver.trace_msg ctxt (trace_after ctxt) result + in result end +end + + +(* overall reconstruction procedure *) + +fun not_supported r = + z3_exn ("proof rule not implemented: " ^ quote (string_of_rule r)) + +fun prove ctxt assms = + let + val prems = Option.map (prepare_assms ctxt) assms + + fun step r ps ct ptab = + (case (r, ps) of + (* core rules *) + (TrueAxiom, _) => (([], Thm true_thm), ptab) + | (Asserted, _) => (([], asserted ctxt prems ct), ptab) + | (Goal, _) => (([], asserted ctxt prems ct), ptab) + | (ModusPonens, [(p, _), (q, _)]) => (([], mp q (thm_of p)), ptab) + | (ModusPonensOeq, [(p, _), (q, _)]) => (([], mp q (thm_of p)), ptab) + | (AndElim, [(p, (_, i))]) => apfst (pair []) (and_elim (p, i) ct ptab) + | (NotOrElim, [(p, (_, i))]) => + apfst (pair []) (not_or_elim (p, i) ct ptab) + | (Hypothesis, _) => (([], Thm (Thm.assume (T.mk_prop ct))), ptab) + | (Lemma, [(p, _)]) => (([], lemma (thm_of p) ct), ptab) + | (UnitResolution, (p, _) :: ps) => + (([], unit_resolution (thm_of p) (map (thm_of o fst) ps) ct), ptab) + | (IffTrue, [(p, _)]) => (([], iff_true (thm_of p)), ptab) + | (IffFalse, [(p, _)]) => (([], iff_false (thm_of p)), ptab) + | (Distributivity, _) => (([], distributivity ct), ptab) + | (DefAxiom, _) => (([], def_axiom ctxt ct), ptab) + | (IntroDef, _) => (intro_def ct, ptab) + | (ApplyDef, [(p, _)]) => (([], apply_def (thm_of p)), ptab) + | (IffOeq, [(p, _)]) => (([], p), ptab) + | (NnfPos, _) => (([], nnf ctxt ps ct), ptab) + | (NnfNeg, _) => (([], nnf ctxt ps ct), ptab) + + (* equality rules *) + | (Reflexivity, _) => (([], refl ct), ptab) + | (Symmetry, [(p, _)]) => (([], symm p), ptab) + | (Transitivity, [(p, _), (q, _)]) => (([], trans p q), ptab) + | (Monotonicity, _) => (([], monotonicity (map fst ps) ct), ptab) + | (Commutativity, _) => (([], commutativity ct), ptab) + + (* quantifier rules *) + | (QuantIntro, [p]) => (([], quant_intro ctxt p ct), ptab) + | (PullQuant, _) => (([], pull_quant ct), ptab) + | (PushQuant, _) => (([], push_quant ct), ptab) + | (ElimUnusedVars, _) => (([], elim_unused_vars ctxt ct), ptab) + | (DestEqRes, _) => (([], dest_eq_res ct), ptab) + | (QuantInst, _) => (([], quant_inst ct), ptab) + | (Skolemize, _) => (skolemize ctxt ct, ptab) + + (* theory rules *) + | (ThLemma, _) => (([], th_lemma ctxt (map (thm_of o fst) ps) ct), ptab) + | (Rewrite, _) => (([], rewrite ctxt [] ct), ptab) + | (RewriteStar, ps) => + (([], rewrite ctxt (map (thm_of o fst) ps) ct), ptab) + + | (NnfStar, _) => not_supported r + | (CnfStar, _) => not_supported r + | (TransitivityStar, _) => not_supported r + | (PullQuantStar, _) => not_supported r + + | _ => z3_exn ("Proof rule " ^ quote (string_of_rule r) ^ + " has an unexpected number of arguments.")) + + fun eq_hyp (ct, cu) = Thm.dest_arg1 ct aconvc Thm.dest_arg1 cu + + fun conclude idx rule prop ((hypss, ps), ptab) = + trace_rule ctxt idx step rule ps prop ptab + |> Config.get ctxt SMT_Solver.trace ? tap (check idx rule prop) + |>> apfst (distinct eq_hyp o fold append hypss) + + fun add_sequent idx vars (hyps, thm) ptab = + let val s = Sequent {hyps=hyps, vars=vars, thm=thm} + in ((hyps, (thm, vars)), Inttab.update (idx, s) ptab) end + + fun lookup idx ptab = + (case Inttab.lookup ptab idx of + SOME (Unproved {rule, subs, vars, prop}) => + fold_map lookup subs ptab + |>> split_list + |>> apsnd (map2 (fn idx => fn (p, vs) => (p, (vs, idx))) subs) + |> conclude idx rule prop + |-> add_sequent idx vars + | SOME (Sequent {hyps, vars, thm}) => ((hyps, (thm, vars)), ptab) + | NONE => z3_exn ("unknown proof id: " ^ quote (string_of_int idx))) + + fun result (hyps, (thm, _)) = + fold SMT_Normalize.discharge_definition hyps (thm_of thm) + + in (fn ptab => fn idx => result (fst (lookup idx ptab))) end + +val setup = Z3_Rewrite_Rules.setup + +end diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Tools/z3_proof_terms.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Tools/z3_proof_terms.ML Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,246 @@ +(* Title: HOL/SMT/Tools/z3_proof_terms.ML + Author: Sascha Boehme, TU Muenchen + +Reconstruction functions for terms occurring in Z3 proofs. +*) + +signature Z3_PROOF_TERMS = +sig + val mk_prop: Thm.cterm -> Thm.cterm + val mk_meta_eq: Thm.cterm -> Thm.cterm -> Thm.cterm + + type preterm + + val compile: theory -> Name.context -> preterm -> Thm.cterm * Thm.cterm list + + val mk_bound: theory -> int -> typ -> preterm + val mk_fun: Thm.cterm -> preterm list -> preterm + val mk_forall: theory -> string * typ -> preterm -> preterm + val mk_exists: theory -> string * typ -> preterm -> preterm + + val mk_true: preterm + val mk_false: preterm + val mk_not: preterm -> preterm + val mk_and: preterm list -> preterm + val mk_or: preterm list -> preterm + val mk_implies: preterm -> preterm -> preterm + val mk_iff: preterm -> preterm -> preterm + + val mk_eq: preterm -> preterm -> preterm + val mk_if: preterm -> preterm -> preterm -> preterm + val mk_distinct: preterm list -> preterm + + val mk_pat: preterm list -> preterm + val mk_nopat: preterm list -> preterm + val mk_trigger: preterm list -> preterm -> preterm + + val mk_access: preterm -> preterm -> preterm + val mk_update: preterm -> preterm -> preterm -> preterm + + val mk_int_num: int -> preterm + val mk_real_frac_num: int * int option -> preterm + val mk_uminus: preterm -> preterm + val mk_add: preterm -> preterm -> preterm + val mk_sub: preterm -> preterm -> preterm + val mk_mul: preterm -> preterm -> preterm + val mk_int_div: preterm -> preterm -> preterm + val mk_real_div: preterm -> preterm -> preterm + val mk_rem: preterm -> preterm -> preterm + val mk_mod: preterm -> preterm -> preterm + val mk_lt: preterm -> preterm -> preterm + val mk_le: preterm -> preterm -> preterm + + val wordT : int -> typ + val mk_bv_num : theory -> int -> int -> preterm + + val var_prefix: string +end + +structure Z3_Proof_Terms: Z3_PROOF_TERMS = +struct + +fun instTs cUs (cTs, ct) = Thm.instantiate_cterm (cTs ~~ cUs, []) ct +fun instT cU (cT, ct) = instTs [cU] ([cT], ct) +fun mk_inst_pair destT cpat = (destT (Thm.ctyp_of_term cpat), cpat) +val destT1 = hd o Thm.dest_ctyp +val destT2 = hd o tl o Thm.dest_ctyp + + +val mk_prop = Thm.capply @{cterm Trueprop} + +val meta_eq = mk_inst_pair destT1 @{cpat "op =="} +fun mk_meta_eq ct = Thm.mk_binop (instT (Thm.ctyp_of_term ct) meta_eq) ct + + +datatype preterm = Preterm of { + cterm: Thm.cterm, + vars: (int * Thm.cterm) list } + +fun mk_preterm (ct, vs) = Preterm {cterm=ct, vars=vs} +fun dest_preterm (Preterm {cterm, vars}) = (cterm, vars) +fun ctyp_of_preterm (Preterm {cterm, ...}) = Thm.ctyp_of_term cterm + +fun instT' e = instT (ctyp_of_preterm e) + +val maxidx_of = #maxidx o Thm.rep_cterm + +val var_prefix = "v" + +local +fun mk_inst nctxt cert vs = + let + val max = fold (curry Int.max o fst) vs 0 + val names = fst (Name.variants (replicate (max + 1) var_prefix) nctxt) + fun mk (i, v) = (v, cert (Free (nth names i, #T (Thm.rep_cterm v)))) + in map mk vs end + +fun fix_vars _ _ ct [] = (ct, []) + | fix_vars thy nctxt ct vs = + let + val cert = Thm.cterm_of thy + val inst = mk_inst nctxt cert vs + in (Thm.instantiate_cterm ([], inst) ct, map snd inst) end +in +fun compile thy nctxt (Preterm {cterm, vars}) = fix_vars thy nctxt cterm vars +end + +local +fun app e (ct1, vs1) = + let + fun part (var as (i, v)) (inst, vs) = + (case AList.lookup (op =) vs1 i of + NONE => (inst, var :: vs) + | SOME v' => ((v, v') :: inst, vs)) + + val (ct2, vs2) = dest_preterm e + val incr = + if maxidx_of ct1 < 0 orelse maxidx_of ct2 < 0 then I + else Thm.incr_indexes_cterm (maxidx_of ct1 + 1) + + val (inst, vs) = fold (part o apsnd incr) vs2 ([], vs1) + val ct2' = Thm.instantiate_cterm ([], inst) (incr ct2) + in (Thm.capply ct1 ct2', vs) end +in +fun mk_fun ct es = mk_preterm (fold app es (ct, [])) +fun mk_binop f t u = mk_fun f [t, u] +fun mk_nary _ e [] = e + | mk_nary ct _ es = uncurry (fold_rev (mk_binop ct)) (split_last es) +end + +fun mk_bound thy i T = + let val ct = Thm.cterm_of thy (Var ((Name.uu, 0), T)) + in mk_preterm (ct, [(i, ct)]) end + +local +fun mk_quant q thy (n, T) e = + let + val (ct, vs) = dest_preterm e + val cv = + (case AList.lookup (op =) vs 0 of + SOME cv => cv + | _ => Thm.cterm_of thy (Var ((Name.uu, maxidx_of ct + 1), T))) + val cq = instT (Thm.ctyp_of_term cv) q + fun dec (i, v) = if i = 0 then NONE else SOME (i - 1, v) + in mk_preterm (Thm.capply cq (Thm.cabs cv ct), map_filter dec vs) end +in +val mk_forall = mk_quant (mk_inst_pair (destT1 o destT1) @{cpat All}) +val mk_exists = mk_quant (mk_inst_pair (destT1 o destT1) @{cpat Ex}) +end + + +val mk_false = mk_fun @{cterm False} [] +val mk_not = mk_fun @{cterm Not} o single +val mk_true = mk_not mk_false +val mk_and = mk_nary @{cterm "op &"} mk_true +val mk_or = mk_nary @{cterm "op |"} mk_false +val mk_implies = mk_binop @{cterm "op -->"} +val mk_iff = mk_binop @{cterm "op = :: bool => _"} + +val eq = mk_inst_pair destT1 @{cpat "op ="} +fun mk_eq t u = mk_binop (instT' t eq) t u + +val if_term = mk_inst_pair (destT1 o destT2) @{cpat If} +fun mk_if c t u = mk_fun (instT' t if_term) [c, t, u] + +val nil_term = mk_inst_pair destT1 @{cpat Nil} +val cons_term = mk_inst_pair destT1 @{cpat Cons} +fun mk_list cT es = + fold_rev (mk_binop (instT cT cons_term)) es (mk_fun (instT cT nil_term) []) + +val distinct = mk_inst_pair (destT1 o destT1) @{cpat distinct} +fun mk_distinct [] = mk_true + | mk_distinct (es as (e :: _)) = + mk_fun (instT' e distinct) [mk_list (ctyp_of_preterm e) es] + +val pat = mk_inst_pair destT1 @{cpat pat} +val nopat = mk_inst_pair destT1 @{cpat nopat} +val andpat = mk_inst_pair (destT1 o destT2) @{cpat "op andpat"} +fun mk_gen_pat _ [] = raise TERM ("mk_gen_pat: empty pattern", []) + | mk_gen_pat pat (e :: es) = + let fun mk t p = mk_fun (instT' t andpat) [p, t] + in fold mk es (mk_fun (instT' e pat) [e]) end +val mk_pat = mk_gen_pat pat +val mk_nopat = mk_gen_pat nopat + +fun mk_trigger es e = mk_fun @{cterm trigger} [mk_list @{ctyp pattern} es, e] + + +val access = mk_inst_pair (Thm.dest_ctyp o destT1) @{cpat apply} +fun mk_access array index = + let val cTs = Thm.dest_ctyp (ctyp_of_preterm array) + in mk_fun (instTs cTs access) [array, index] end + +val update = mk_inst_pair (Thm.dest_ctyp o destT1) @{cpat fun_upd} +fun mk_update array index value = + let val cTs = Thm.dest_ctyp (ctyp_of_preterm array) + in mk_fun (instTs cTs update) [array, index, value] end + + +fun mk_int_num i = mk_fun (Numeral.mk_cnumber @{ctyp int} i) [] +fun mk_real_num i = mk_fun (Numeral.mk_cnumber @{ctyp real} i) [] + +fun mk_real_frac_num (e, NONE) = mk_real_num e + | mk_real_frac_num (e, SOME d) = + mk_binop @{cterm "op / :: real => _"} (mk_real_num e) (mk_real_num d) + +fun has_int_type e = (Thm.typ_of (ctyp_of_preterm e) = @{typ int}) +fun choose e i r = if has_int_type e then i else r + +val uminus_i = @{cterm "uminus :: int => _"} +val uminus_r = @{cterm "uminus :: real => _"} +fun mk_uminus e = mk_fun (choose e uminus_i uminus_r) [e] + +fun arith_op int_op real_op t u = mk_binop (choose t int_op real_op) t u + +val mk_add = arith_op @{cterm "op + :: int => _"} @{cterm "op + :: real => _"} +val mk_sub = arith_op @{cterm "op - :: int => _"} @{cterm "op - :: real => _"} +val mk_mul = arith_op @{cterm "op * :: int => _"} @{cterm "op * :: real => _"} +val mk_int_div = mk_binop @{cterm "op div :: int => _"} +val mk_real_div = mk_binop @{cterm "op / :: real => _"} +val mk_rem = mk_binop @{cterm "op rem :: int => _"} +val mk_mod = mk_binop @{cterm "op mod :: int => _"} +val mk_lt = arith_op @{cterm "op < :: int => _"} @{cterm "op < :: real => _"} +val mk_le = arith_op @{cterm "op <= :: int => _"} @{cterm "op <= :: real => _"} + +fun binT size = + let + fun bitT i T = + if i = 0 + then Type (@{type_name "Numeral_Type.bit0"}, [T]) + else Type (@{type_name "Numeral_Type.bit1"}, [T]) + + fun binT i = + if i = 0 then @{typ "Numeral_Type.num0"} + else if i = 1 then @{typ "Numeral_Type.num1"} + else let val (q, r) = Integer.div_mod i 2 in bitT r (binT q) end + in + if size >= 0 then binT size + else raise TYPE ("mk_binT: " ^ string_of_int size, [], []) + end + +fun wordT size = Type (@{type_name "word"}, [binT size]) + +fun mk_bv_num thy num size = + mk_fun (Numeral.mk_cnumber (Thm.ctyp_of thy (wordT size)) num) [] + +end diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Tools/z3_solver.ML --- a/src/HOL/SMT/Tools/z3_solver.ML Tue Oct 20 08:10:47 2009 +0200 +++ b/src/HOL/SMT/Tools/z3_solver.ML Tue Oct 20 10:11:30 2009 +0200 @@ -8,7 +8,6 @@ sig val proofs: bool Config.T val options: string Config.T - val setup: theory -> theory end @@ -57,20 +56,18 @@ check_unsat recon output |> K @{cprop False} -(* FIXME fun prover (SMT_Solver.ProofData {context, output, recon, assms}) = check_unsat recon output |> Z3_Proof.reconstruct context assms recon -*) fun solver oracle ctxt = let val with_proof = Config.get ctxt proofs in SMT_Solver.SolverConfig { - name = {env_var=env_var, remote_name=solver_name}, + command = {env_var=env_var, remote_name=solver_name}, + arguments = cmdline_options ctxt, interface = Z3_Interface.interface, - arguments = cmdline_options ctxt, - reconstruct = (*FIXME:if with_proof then prover else*) oracle } + reconstruct = if with_proof then prover else oracle } end val setup = diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/Z3.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Z3.thy Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,24 @@ +(* Title: HOL/SMT/Z3.thy + Author: Sascha Boehme, TU Muenchen +*) + +header {* Binding to the SMT solver Z3, with proof reconstruction *} + +theory Z3 +imports SMT_Base +uses + "Tools/z3_proof_terms.ML" + "Tools/z3_proof_rules.ML" + "Tools/z3_proof.ML" + "Tools/z3_model.ML" + "Tools/z3_interface.ML" + "Tools/z3_solver.ML" +begin + +setup {* Z3_Proof_Rules.setup #> Z3_Solver.setup *} + +lemmas [z3_rewrite] = + refl eq_commute conj_commute disj_commute simp_thms nnf_simps + ring_distribs field_eq_simps + +end diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/etc/settings --- a/src/HOL/SMT/etc/settings Tue Oct 20 08:10:47 2009 +0200 +++ b/src/HOL/SMT/etc/settings Tue Oct 20 10:11:30 2009 +0200 @@ -1,9 +1,11 @@ ISABELLE_SMT="$COMPONENT" -REMOTE_SMT_SOLVER="$ISABELLE_SMT/lib/scripts/remote_smt.pl" +REMOTE_SMT_SOLVER="$COMPONENT/lib/scripts/remote_smt.pl" REMOTE_SMT_URL="http://www4.in.tum.de/smt/smt" +CERT_SMT_SOLVER="$COMPONENT/lib/scripts/cert_smt.pl" + # # Paths to local SMT solvers: # diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/lib/scripts/cert_smt.pl --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/lib/scripts/cert_smt.pl Tue Oct 20 10:11:30 2009 +0200 @@ -0,0 +1,33 @@ +# +# Author: Sascha Boehme, TU Muenchen +# +# Fake SMT solver: check that input matches previously computed input and +# and return previously computed output. +# + +use strict; +use File::Compare; + + +# arguments + +my $cert_path = $ARGV[0]; +my $new_problem = $ARGV[1]; + + +# check content of new problem file against old problem file + +my $old_problem = $cert_path; +my $old_proof = $cert_path . ".proof"; + +if (-e $old_problem and compare($old_problem, $new_problem) == 0) { + if (-e $old_proof) { + open FILE, "<$old_proof"; + foreach () { + print $_; + } + close FILE; + } + else { print "ERROR: unable to open proof file\n"; } +} +else { print "ERROR: bad problem\n"; } diff -r b0ff69f0a248 -r 39f73a59e855 src/HOL/SMT/lib/scripts/remote_smt.pl --- a/src/HOL/SMT/lib/scripts/remote_smt.pl Tue Oct 20 08:10:47 2009 +0200 +++ b/src/HOL/SMT/lib/scripts/remote_smt.pl Tue Oct 20 10:11:30 2009 +0200 @@ -1,7 +1,8 @@ # -# Script to invoke remote SMT solvers. # Author: Sascha Boehme, TU Muenchen # +# Invoke remote SMT solvers. +# use strict; use LWP;