diff -r 624faeda77b5 -r 175ac95720d4 src/HOL/SMT_Examples/SMT_Examples.thy --- a/src/HOL/SMT_Examples/SMT_Examples.thy Thu Mar 13 13:18:13 2014 +0100 +++ b/src/HOL/SMT_Examples/SMT_Examples.thy Thu Mar 13 13:18:13 2014 +0100 @@ -11,38 +11,40 @@ declare [[smt_certificates = "SMT_Examples.certs"]] declare [[smt_read_only_certificates = true]] +declare [[smt2_certificates = "SMT_Examples.certs2"]] +declare [[smt2_read_only_certificates = true]] section {* Propositional and first-order logic *} -lemma "True" by smt +lemma "True" by smt2 -lemma "p \ \p" by smt +lemma "p \ \p" by smt2 -lemma "(p \ True) = p" by smt +lemma "(p \ True) = p" by smt2 -lemma "(p \ q) \ \p \ q" by smt +lemma "(p \ q) \ \p \ q" by smt2 lemma "(a \ b) \ (c \ d) \ (a \ b) \ (c \ d)" - by smt + by smt2 -lemma "(p1 \ p2) \ p3 \ (p1 \ (p3 \ p2) \ (p1 \ p3)) \ p1" by smt +lemma "(p1 \ p2) \ p3 \ (p1 \ (p3 \ p2) \ (p1 \ p3)) \ p1" by smt2 -lemma "P=P=P=P=P=P=P=P=P=P" by smt +lemma "P = P = P = P = P = P = P = P = P = P" by smt2 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))))" + 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 assms by smt + using assms by smt2 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 symm_f) +lemma "a = a \ symm_f a b = symm_f b a" by (smt2 symm_f) (* Taken from ~~/src/HOL/ex/SAT_Examples.thy. @@ -53,253 +55,253 @@ 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" + 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 by smt + using assms by smt (* smt2 FIXME: THM 0 *) lemma "\x::int. P x \ (\y::int. P x \ P y)" - by smt + by smt2 lemma assumes "(\x y. P x y = x)" shows "(\y. P x y) = P x c" - using assms by smt + using assms by smt (* smt2 FIXME: Option *) 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 by smt + using assms by smt (* smt2 FIXME: Option *) lemma assumes "if P x then \(\y. P y) else (\y. \P y)" shows "P x \ P y" - using assms by smt + using assms by smt2 section {* Arithmetic *} subsection {* Linear arithmetic over integers and reals *} -lemma "(3::int) = 3" by smt +lemma "(3::int) = 3" by smt2 -lemma "(3::real) = 3" by smt +lemma "(3::real) = 3" by smt2 -lemma "(3 :: int) + 1 = 4" by smt +lemma "(3 :: int) + 1 = 4" by smt2 -lemma "x + (y + z) = y + (z + (x::int))" by smt +lemma "x + (y + z) = y + (z + (x::int))" by smt2 -lemma "max (3::int) 8 > 5" by smt +lemma "max (3::int) 8 > 5" by smt2 -lemma "abs (x :: real) + abs y \ abs (x + y)" by smt +lemma "abs (x :: real) + abs y \ abs (x + y)" by smt2 -lemma "P ((2::int) < 3) = P True" by smt +lemma "P ((2::int) < 3) = P True" by smt2 -lemma "x + 3 \ 4 \ x < (1::int)" by smt +lemma "x + 3 \ 4 \ x < (1::int)" by smt2 lemma assumes "x \ (3::int)" and "y = x + 4" shows "y - x > 0" - using assms by smt + using assms by smt2 -lemma "let x = (2 :: int) in x + x \ 5" by smt +lemma "let x = (2 :: int) in x + x \ 5" by smt2 lemma fixes x :: real assumes "3 * x + 7 * a < 4" and "3 < 2 * x" shows "a < 0" - using assms by smt + using assms by smt2 -lemma "(0 \ y + -1 * x \ \ 0 \ x \ 0 \ (x::int)) = (\ False)" by smt +lemma "(0 \ y + -1 * x \ \ 0 \ x \ 0 \ (x::int)) = (\ False)" by smt2 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))" - by smt + (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))" + by smt2 text{* The following example was taken from HOL/ex/PresburgerEx.thy, where it says: @@ -320,175 +322,173 @@ 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)" - by smt + \ x1 = x10 \ x2 = (x11::int)" + by smt2 -lemma "let P = 2 * x + 1 > x + (x::real) in P \ False \ P" by smt +lemma "let P = 2 * x + 1 > x + (x::real) in P \ False \ P" by smt2 lemma "x + (let y = x mod 2 in 2 * y + 1) \ x + (1::int)" - using [[z3_with_extensions]] - by smt + using [[z3_new_extensions]] + by smt2 lemma "x + (let y = x mod 2 in y + y) < x + (3::int)" - using [[z3_with_extensions]] - by smt + using [[z3_new_extensions]] + by smt2 lemma assumes "x \ (0::real)" - shows "x + x \ (let P = (abs x > 1) in if P \ \P then 4 else 2) * x" - using assms by smt + shows "x + x \ (let P = (abs x > 1) in if P \ \ P then 4 else 2) * x" + using assms [[z3_new_extensions]] by smt2 lemma assumes "(n + m) mod 2 = 0" and "n mod 4 = 3" - shows "n mod 2 = 1 & m mod 2 = (1::int)" - using assms [[z3_with_extensions]] by smt - + shows "n mod 2 = 1 \ m mod 2 = (1::int)" + using assms [[z3_new_extensions]] by smt2 subsection {* Linear arithmetic with quantifiers *} -lemma "~ (\x::int. False)" by smt +lemma "~ (\x::int. False)" by smt2 -lemma "~ (\x::real. False)" by smt +lemma "~ (\x::real. False)" by smt2 lemma "\x::int. 0 < x" - using [[smt_oracle=true]] (* no Z3 proof *) - by smt + using [[smt2_oracle=true]] (* no Z3 proof *) + by smt2 lemma "\x::real. 0 < x" - using [[smt_oracle=true]] (* no Z3 proof *) - by smt + using [[smt2_oracle=true]] (* no Z3 proof *) + by smt2 lemma "\x::int. \y. y > x" - using [[smt_oracle=true]] (* no Z3 proof *) - by smt + using [[smt2_oracle=true]] (* no Z3 proof *) + by smt2 -lemma "\x y::int. (x = 0 \ y = 1) \ x \ y" by smt +lemma "\x y::int. (x = 0 \ y = 1) \ x \ y" by smt2 -lemma "\x::int. \y. x < y \ y < 0 \ y >= 0" by smt +lemma "\x::int. \y. x < y \ y < 0 \ y >= 0" by smt2 -lemma "\x y::int. x < y \ (2 * x + 1) < (2 * y)" by smt +lemma "\x y::int. x < y \ (2 * x + 1) < (2 * y)" by smt2 -lemma "\x y::int. (2 * x + 1) \ (2 * y)" by smt +lemma "\x y::int. (2 * x + 1) \ (2 * y)" by smt2 -lemma "\x y::int. x + y > 2 \ x + y = 2 \ x + y < 2" by smt +lemma "\x y::int. x + y > 2 \ x + y = 2 \ x + y < 2" by smt2 -lemma "\x::int. if x > 0 then x + 1 > 0 else 1 > x" by smt +lemma "\x::int. if x > 0 then x + 1 > 0 else 1 > x" by smt2 -lemma "if (ALL x::int. x < 0 \ x > 0) then False else True" by smt +lemma "if (ALL x::int. x < 0 \ x > 0) then False else True" by smt2 -lemma "(if (ALL x::int. x < 0 \ x > 0) then -1 else 3) > (0::int)" by smt +lemma "(if (ALL x::int. x < 0 \ x > 0) then -1 else 3) > (0::int)" by smt2 -lemma "~ (\x y z::int. 4 * x + -6 * y = (1::int))" by smt +lemma "~ (\x y z::int. 4 * x + -6 * y = (1::int))" by smt2 -lemma "\x::int. \x y. 0 < x \ 0 < y \ (0::int) < x + y" by smt +lemma "\x::int. \x y. 0 < x \ 0 < y \ (0::int) < x + y" by smt2 -lemma "\u::int. \(x::int) y::real. 0 < x \ 0 < y \ -1 < x" by smt +lemma "\u::int. \(x::int) y::real. 0 < x \ 0 < y \ -1 < x" by smt2 -lemma "\x::int. (\y. y \ x \ y > 0) \ x > 0" by smt +lemma "\x::int. (\y. y \ x \ y > 0) \ x > 0" by smt2 -lemma "\x::int. SMT.trigger [[SMT.pat x]] (x < a \ 2 * x < 2 * a)" by smt +lemma "\x::int. SMT2.trigger [[SMT2.pat x]] (x < a \ 2 * x < 2 * a)" by smt2 -lemma "\(a::int) b::int. 0 < b \ b < 1" by smt +lemma "\(a::int) b::int. 0 < b \ b < 1" by smt2 subsection {* Non-linear arithmetic over integers and reals *} lemma "a > (0::int) \ a*b > 0 \ b > 0" - using [[smt_oracle, z3_with_extensions]] - by smt + using [[smt2_oracle, z3_new_extensions]] + by smt2 lemma "(a::int) * (x + 1 + y) = a * x + a * (y + 1)" - using [[z3_with_extensions]] - by smt + using [[z3_new_extensions]] + by smt2 lemma "((x::real) * (1 + y) - x * (1 - y)) = (2 * x * y)" - using [[z3_with_extensions]] - by smt + using [[z3_new_extensions]] + by smt2 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 [[z3_with_extensions]] - by smt + using [[z3_new_extensions]] + by smt2 -lemma [z3_rule]: +lemma [z3_rule, z3_new_rule]: fixes x :: "int" assumes "x * y \ 0" and "\ y \ 0" and "\ x \ 0" shows False using assms by (metis mult_le_0_iff) lemma "x * y \ (0 :: int) \ x \ 0 \ y \ 0" - using [[z3_with_extensions]] - by smt - + using [[z3_with_extensions]] [[z3_new_extensions]] + by smt (* smt2 FIXME: "th-lemma" tactic fails *) subsection {* Linear arithmetic for natural numbers *} -lemma "2 * (x::nat) ~= 1" by smt +lemma "2 * (x::nat) ~= 1" by smt2 -lemma "a < 3 \ (7::nat) > 2 * a" by smt +lemma "a < 3 \ (7::nat) > 2 * a" by smt2 -lemma "let x = (1::nat) + y in x - y > 0 * x" by smt +lemma "let x = (1::nat) + y in x - y > 0 * x" by smt2 lemma "let x = (1::nat) + y in let P = (if x > 0 then True else False) in False \ P = (x - 1 = y) \ (\P \ False)" - by smt + by smt2 -lemma "int (nat \x::int\) = \x\" by smt +lemma "int (nat \x::int\) = \x\" by smt2 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)" by (smt prime_nat_def) +lemma "prime_nat (4*m + 1) \ m \ (1::nat)" by (smt2 prime_nat_def) section {* Pairs *} lemma "fst (x, y) = a \ x = a" using fst_conv - by smt + by smt2 lemma "p1 = (x, y) \ p2 = (y, x) \ fst p1 = snd p2" using fst_conv snd_conv - by smt + by smt2 section {* Higher-order problems and recursion *} lemma "i \ i1 \ i \ i2 \ (f (i1 := v1, i2 := v2)) i = f i" using fun_upd_same fun_upd_apply - by smt + by smt2 lemma "(f g (x::'a::type) = (g x \ True)) \ (f g x = True) \ (g x = True)" - by smt + by smt2 -lemma "id x = x \ id True = True" by (smt id_def) +lemma "id x = x \ id True = True" by (smt id_def) (* smt2 FIXME: Option *) lemma "i \ i1 \ i \ i2 \ ((f (i1 := v1)) (i2 := v2)) i = f i" using fun_upd_same fun_upd_apply - by smt + by smt2 lemma "f (\x. g x) \ True" "f (\x. g x) \ True" - by smt+ + by smt2+ -lemma True using let_rsp by smt +lemma True using let_rsp by smt2 -lemma "le = op \ \ le (3::int) 42" by smt +lemma "le = op \ \ le (3::int) 42" by smt2 -lemma "map (\i::nat. i + 1) [0, 1] = [1, 2]" by (smt list.map) +lemma "map (\i::nat. i + 1) [0, 1] = [1, 2]" by (smt2 list.map) -lemma "(ALL x. P x) | ~ All P" by smt +lemma "(ALL x. P x) \ ~ All P" by smt2 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 dec_10.simps) +lemma "dec_10 (4 * dec_10 4) = 6" by (smt2 dec_10.simps) axiomatization @@ -505,35 +505,36 @@ (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_oracle=true]] (*FIXME*) - using [[z3_with_extensions]] - by (smt eval_dioph_mod[where n=2] eval_dioph_div_mult[where n=2]) + using [[smt2_oracle=true]] (*FIXME*) + using [[z3_new_extensions]] + by (smt2 eval_dioph_mod[where n=2] eval_dioph_div_mult[where n=2]) context complete_lattice begin lemma - assumes "Sup { a | i::bool . True } \ Sup { b | i::bool . True }" - and "Sup { b | i::bool . True } \ Sup { a | i::bool . True }" - shows "Sup { a | i::bool . True } \ Sup { a | i::bool . True }" - using assms by (smt order_trans) + assumes "Sup {a | i::bool. True} \ Sup {b | i::bool. True}" + and "Sup {b | i::bool. True} \ Sup {a | i::bool. True}" + shows "Sup {a | i::bool. True} \ Sup {a | i::bool. True}" + using assms by (smt2 order_trans) end - section {* Monomorphization examples *} definition Pred :: "'a \ bool" where "Pred x = True" -lemma poly_Pred: "Pred x \ (Pred [x] \ \Pred[x])" by (simp add: Pred_def) -lemma "Pred (1::int)" by (smt poly_Pred) + +lemma poly_Pred: "Pred x \ (Pred [x] \ \ Pred [x])" by (simp add: Pred_def) +lemma "Pred (1::int)" by (smt2 poly_Pred) axiomatization g :: "'a \ nat" axiomatization where g1: "g (Some x) = g [x]" and g2: "g None = g []" and g3: "g xs = length xs" -lemma "g (Some (3::int)) = g (Some True)" by (smt g1 g2 g3 list.size) + +lemma "g (Some (3::int)) = g (Some True)" by (smt g1 g2 g3 list.size) (* smt2 FIXME: Option *) end