(* Title: HOL/ex/PresburgerEx.thy ID: $Id: PresburgerEx.thy,v 1.13 2008/01/03 09:27:40 chaieb Exp $ Author: Amine Chaieb, TU Muenchen *) header {* Some examples for Presburger Arithmetic *} theory PresburgerEx imports Presburger begin lemma "!!m n ja ia. [|¬ m ≤ j; ¬ (n::nat) ≤ i; (e::nat) ≠ 0; Suc j ≤ ja|] ==> ∃m. ∀ja ia. m ≤ ja --> (if j = ja ∧ i = ia then e else 0) = 0" by presburger lemma "(0::nat) < emBits mod 8 ==> 8 + emBits div 8 * 8 - emBits = 8 - emBits mod 8" by presburger lemma "(0::nat) < emBits mod 8 ==> 8 + emBits div 8 * 8 - emBits = 8 - emBits mod 8" by presburger theorem "(∀(y::int). 3 dvd y) ==> ∀(x::int). b < x --> a ≤ x" by presburger theorem "!! (y::int) (z::int) (n::int). 3 dvd z ==> 2 dvd (y::int) ==> (∃(x::int). 2*x = y) & (∃(k::int). 3*k = z)" by presburger theorem "!! (y::int) (z::int) n. Suc(n::nat) < 6 ==> 3 dvd z ==> 2 dvd (y::int) ==> (∃(x::int). 2*x = y) & (∃(k::int). 3*k = z)" by presburger theorem "∀(x::nat). ∃(y::nat). (0::nat) ≤ 5 --> y = 5 + x " by presburger text{*Slow: about 7 seconds on a 1.6GHz machine.*} theorem "∀(x::nat). ∃(y::nat). y = 5 + x | x div 6 + 1= 2" by presburger theorem "∃(x::int). 0 < x" by presburger theorem "∀(x::int) y. x < y --> 2 * x + 1 < 2 * y" by presburger theorem "∀(x::int) y. 2 * x + 1 ≠ 2 * y" by presburger theorem "∃(x::int) y. 0 < x & 0 ≤ y & 3 * x - 5 * y = 1" by presburger theorem "~ (∃(x::int) (y::int) (z::int). 4*x + (-6::int)*y = 1)" by presburger theorem "∀(x::int). b < x --> a ≤ x" apply (presburger elim) oops theorem "~ (∃(x::int). False)" by presburger theorem "∀(x::int). (a::int) < 3 * x --> b < 3 * x" apply (presburger elim) oops theorem "∀(x::int). (2 dvd x) --> (∃(y::int). x = 2*y)" by presburger theorem "∀(x::int). (2 dvd x) --> (∃(y::int). x = 2*y)" by presburger theorem "∀(x::int). (2 dvd x) = (∃(y::int). x = 2*y)" by presburger theorem "∀(x::int). ((2 dvd x) = (∀(y::int). x ≠ 2*y + 1))" by presburger theorem "~ (∀(x::int). ((2 dvd x) = (∀(y::int). x ≠ 2*y+1) | (∃(q::int) (u::int) i. 3*i + 2*q - u < 17) --> 0 < x | ((~ 3 dvd x) &(x + 8 = 0))))" by presburger theorem "~ (∀(i::int). 4 ≤ i --> (∃x y. 0 ≤ x & 0 ≤ y & 3 * x + 5 * y = i))" by presburger theorem "∀(i::int). 8 ≤ i --> (∃x y. 0 ≤ x & 0 ≤ y & 3 * x + 5 * y = i)" by presburger theorem "∃(j::int). ∀i. j ≤ i --> (∃x y. 0 ≤ x & 0 ≤ y & 3 * x + 5 * y = i)" by presburger theorem "~ (∀j (i::int). j ≤ i --> (∃x y. 0 ≤ x & 0 ≤ y & 3 * x + 5 * y = i))" by presburger text{*Slow: about 5 seconds on a 1.6GHz machine.*} theorem "(∃m::nat. n = 2 * m) --> (n + 1) div 2 = n div 2" by presburger text{* This following theorem proves that all solutions to the recurrence relation $x_{i+2} = |x_{i+1}| - x_i$ are periodic with period 9. The example was brought to our attention by John Harrison. It does does not require Presburger arithmetic but merely quantifier-free linear arithmetic and holds for the rationals as well. Warning: it takes (in 2006) over 4.2 minutes! *} 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 arith end
lemma
[| ¬ m ≤ j; ¬ n ≤ i; e ≠ 0; Suc j ≤ ja |]
==> ∃m. ∀ja ia. m ≤ ja --> (if j = ja ∧ i = ia then e else 0) = 0
lemma
0 < emBits mod 8 ==> 8 + emBits div 8 * 8 - emBits = 8 - emBits mod 8
lemma
0 < emBits mod 8 ==> 8 + emBits div 8 * 8 - emBits = 8 - emBits mod 8
theorem
∀y. 3 dvd y ==> ∀x>b. a ≤ x
theorem
[| 3 dvd z; 2 dvd y |] ==> (∃x. 2 * x = y) ∧ (∃k. 3 * k = z)
theorem
[| Suc n < 6; 3 dvd z; 2 dvd y |] ==> (∃x. 2 * x = y) ∧ (∃k. 3 * k = z)
theorem
∀x. ∃y. 0 ≤ 5 --> y = 5 + x
theorem
∀x. ∃y. y = 5 + x ∨ x div 6 + 1 = 2
theorem
∃x. 0 < x
theorem
∀x y. x < y --> 2 * x + 1 < 2 * y
theorem
∀x y. 2 * x + 1 ≠ 2 * y
theorem
∃x y. 0 < x ∧ 0 ≤ y ∧ 3 * x - 5 * y = 1
theorem
¬ (∃x y z. 4 * x + -6 * y = 1)
theorem
¬ (∃x. False)
theorem
∀x. 2 dvd x --> (∃y. x = 2 * y)
theorem
∀x. 2 dvd x --> (∃y. x = 2 * y)
theorem
∀x. (2 dvd x) = (∃y. x = 2 * y)
theorem
∀x. (2 dvd x) = (∀y. x ≠ 2 * y + 1)
theorem
¬ (∀x. (2 dvd x) = (∀y. x ≠ 2 * y + 1) ∨ (∃q u i. 3 * i + 2 * q - u < 17) -->
0 < x ∨ ¬ 3 dvd x ∧ x + 8 = 0)
theorem
¬ (∀i≥4. ∃x y. 0 ≤ x ∧ 0 ≤ y ∧ 3 * x + 5 * y = i)
theorem
∀i≥8. ∃x y. 0 ≤ x ∧ 0 ≤ y ∧ 3 * x + 5 * y = i
theorem
∃j. ∀i≥j. ∃x y. 0 ≤ x ∧ 0 ≤ y ∧ 3 * x + 5 * y = i
theorem
¬ (∀j i. j ≤ i --> (∃x y. 0 ≤ x ∧ 0 ≤ y ∧ 3 * x + 5 * y = i))
theorem
(∃m. n = 2 * m) --> (n + 1) div 2 = n div 2
lemma
[| x3.0 = ¦x2.0¦ - x1.0; x4.0 = ¦x3.0¦ - x2.0; x5.0 = ¦x4.0¦ - x3.0;
x6.0 = ¦x5.0¦ - x4.0; x7.0 = ¦x6.0¦ - x5.0; x8.0 = ¦x7.0¦ - x6.0;
x9.0 = ¦x8.0¦ - x7.0; x10.0 = ¦x9.0¦ - x8.0; x11.0 = ¦x10.0¦ - x9.0 |]
==> x1.0 = x10.0 ∧ x2.0 = x11.0