| author | paulson <lp15@cam.ac.uk> | 
| Wed, 23 Aug 2017 23:46:35 +0100 | |
| changeset 66497 | 18a6478a574c | 
| parent 66453 | cc19f7ca2ed6 | 
| child 67006 | b1278ed3cd46 | 
| permissions | -rw-r--r-- | 
| 13880 | 1  | 
(* Title: HOL/ex/PresburgerEx.thy  | 
2  | 
Author: Amine Chaieb, TU Muenchen  | 
|
| 17388 | 3  | 
*)  | 
| 13880 | 4  | 
|
| 61343 | 5  | 
section \<open>Some examples for Presburger Arithmetic\<close>  | 
| 13880 | 6  | 
|
| 23462 | 7  | 
theory PresburgerEx  | 
| 
66453
 
cc19f7ca2ed6
session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
 
wenzelm 
parents: 
61945 
diff
changeset
 | 
8  | 
imports HOL.Presburger  | 
| 23462 | 9  | 
begin  | 
| 23323 | 10  | 
|
| 25801 | 11  | 
lemma "\<And>m n ja ia. \<lbrakk>\<not> m \<le> j; \<not> (n::nat) \<le> i; (e::nat) \<noteq> 0; Suc j \<le> ja\<rbrakk> \<Longrightarrow> \<exists>m. \<forall>ja ia. m \<le> ja \<longrightarrow> (if j = ja \<and> i = ia then e else 0) = 0" by presburger  | 
| 23323 | 12  | 
lemma "(0::nat) < emBits mod 8 \<Longrightarrow> 8 + emBits div 8 * 8 - emBits = 8 - emBits mod 8"  | 
13  | 
by presburger  | 
|
14  | 
lemma "(0::nat) < emBits mod 8 \<Longrightarrow> 8 + emBits div 8 * 8 - emBits = 8 - emBits mod 8"  | 
|
15  | 
by presburger  | 
|
| 13880 | 16  | 
|
| 
14353
 
79f9fbef9106
Added lemmas to Ring_and_Field with slightly modified simplification rules
 
paulson 
parents: 
13880 
diff
changeset
 | 
17  | 
theorem "(\<forall>(y::int). 3 dvd y) ==> \<forall>(x::int). b < x --> a \<le> x"  | 
| 13880 | 18  | 
by presburger  | 
19  | 
||
20  | 
theorem "!! (y::int) (z::int) (n::int). 3 dvd z ==> 2 dvd (y::int) ==>  | 
|
| 
14353
 
79f9fbef9106
Added lemmas to Ring_and_Field with slightly modified simplification rules
 
paulson 
parents: 
13880 
diff
changeset
 | 
21  | 
(\<exists>(x::int). 2*x = y) & (\<exists>(k::int). 3*k = z)"  | 
| 13880 | 22  | 
by presburger  | 
23  | 
||
24  | 
theorem "!! (y::int) (z::int) n. Suc(n::nat) < 6 ==> 3 dvd z ==>  | 
|
| 
14353
 
79f9fbef9106
Added lemmas to Ring_and_Field with slightly modified simplification rules
 
paulson 
parents: 
13880 
diff
changeset
 | 
25  | 
2 dvd (y::int) ==> (\<exists>(x::int). 2*x = y) & (\<exists>(k::int). 3*k = z)"  | 
| 13880 | 26  | 
by presburger  | 
27  | 
||
| 15075 | 28  | 
theorem "\<forall>(x::nat). \<exists>(y::nat). (0::nat) \<le> 5 --> y = 5 + x "  | 
| 13880 | 29  | 
by presburger  | 
30  | 
||
| 61343 | 31  | 
text\<open>Slow: about 7 seconds on a 1.6GHz machine.\<close>  | 
| 15075 | 32  | 
theorem "\<forall>(x::nat). \<exists>(y::nat). y = 5 + x | x div 6 + 1= 2"  | 
33  | 
by presburger  | 
|
34  | 
||
35  | 
theorem "\<exists>(x::int). 0 < x"  | 
|
| 13880 | 36  | 
by presburger  | 
37  | 
||
| 15075 | 38  | 
theorem "\<forall>(x::int) y. x < y --> 2 * x + 1 < 2 * y"  | 
39  | 
by presburger  | 
|
| 13880 | 40  | 
|
| 15075 | 41  | 
theorem "\<forall>(x::int) y. 2 * x + 1 \<noteq> 2 * y"  | 
42  | 
by presburger  | 
|
| 13880 | 43  | 
|
| 15075 | 44  | 
theorem "\<exists>(x::int) y. 0 < x & 0 \<le> y & 3 * x - 5 * y = 1"  | 
45  | 
by presburger  | 
|
| 13880 | 46  | 
|
| 
14353
 
79f9fbef9106
Added lemmas to Ring_and_Field with slightly modified simplification rules
 
paulson 
parents: 
13880 
diff
changeset
 | 
47  | 
theorem "~ (\<exists>(x::int) (y::int) (z::int). 4*x + (-6::int)*y = 1)"  | 
| 13880 | 48  | 
by presburger  | 
49  | 
||
| 
14353
 
79f9fbef9106
Added lemmas to Ring_and_Field with slightly modified simplification rules
 
paulson 
parents: 
13880 
diff
changeset
 | 
50  | 
theorem "\<forall>(x::int). b < x --> a \<le> x"  | 
| 23323 | 51  | 
apply (presburger elim)  | 
| 13880 | 52  | 
oops  | 
53  | 
||
| 
14353
 
79f9fbef9106
Added lemmas to Ring_and_Field with slightly modified simplification rules
 
paulson 
parents: 
13880 
diff
changeset
 | 
54  | 
theorem "~ (\<exists>(x::int). False)"  | 
| 13880 | 55  | 
by presburger  | 
56  | 
||
| 
14353
 
79f9fbef9106
Added lemmas to Ring_and_Field with slightly modified simplification rules
 
paulson 
parents: 
13880 
diff
changeset
 | 
57  | 
theorem "\<forall>(x::int). (a::int) < 3 * x --> b < 3 * x"  | 
| 23323 | 58  | 
apply (presburger elim)  | 
| 13880 | 59  | 
oops  | 
60  | 
||
| 15075 | 61  | 
theorem "\<forall>(x::int). (2 dvd x) --> (\<exists>(y::int). x = 2*y)"  | 
62  | 
by presburger  | 
|
| 13880 | 63  | 
|
| 15075 | 64  | 
theorem "\<forall>(x::int). (2 dvd x) --> (\<exists>(y::int). x = 2*y)"  | 
65  | 
by presburger  | 
|
| 13880 | 66  | 
|
| 15075 | 67  | 
theorem "\<forall>(x::int). (2 dvd x) = (\<exists>(y::int). x = 2*y)"  | 
68  | 
by presburger  | 
|
| 
14353
 
79f9fbef9106
Added lemmas to Ring_and_Field with slightly modified simplification rules
 
paulson 
parents: 
13880 
diff
changeset
 | 
69  | 
|
| 15075 | 70  | 
theorem "\<forall>(x::int). ((2 dvd x) = (\<forall>(y::int). x \<noteq> 2*y + 1))"  | 
71  | 
by presburger  | 
|
| 13880 | 72  | 
|
| 
14353
 
79f9fbef9106
Added lemmas to Ring_and_Field with slightly modified simplification rules
 
paulson 
parents: 
13880 
diff
changeset
 | 
73  | 
theorem "~ (\<forall>(x::int).  | 
| 
 
79f9fbef9106
Added lemmas to Ring_and_Field with slightly modified simplification rules
 
paulson 
parents: 
13880 
diff
changeset
 | 
74  | 
((2 dvd x) = (\<forall>(y::int). x \<noteq> 2*y+1) |  | 
| 
 
79f9fbef9106
Added lemmas to Ring_and_Field with slightly modified simplification rules
 
paulson 
parents: 
13880 
diff
changeset
 | 
75  | 
(\<exists>(q::int) (u::int) i. 3*i + 2*q - u < 17)  | 
| 
 
79f9fbef9106
Added lemmas to Ring_and_Field with slightly modified simplification rules
 
paulson 
parents: 
13880 
diff
changeset
 | 
76  | 
--> 0 < x | ((~ 3 dvd x) &(x + 8 = 0))))"  | 
| 13880 | 77  | 
by presburger  | 
78  | 
||
| 15075 | 79  | 
theorem "~ (\<forall>(i::int). 4 \<le> i --> (\<exists>x y. 0 \<le> x & 0 \<le> y & 3 * x + 5 * y = i))"  | 
| 13880 | 80  | 
by presburger  | 
81  | 
||
| 15075 | 82  | 
theorem "\<forall>(i::int). 8 \<le> i --> (\<exists>x y. 0 \<le> x & 0 \<le> y & 3 * x + 5 * y = i)"  | 
| 
14353
 
79f9fbef9106
Added lemmas to Ring_and_Field with slightly modified simplification rules
 
paulson 
parents: 
13880 
diff
changeset
 | 
83  | 
by presburger  | 
| 13880 | 84  | 
|
| 15075 | 85  | 
theorem "\<exists>(j::int). \<forall>i. j \<le> i --> (\<exists>x y. 0 \<le> x & 0 \<le> y & 3 * x + 5 * y = i)"  | 
86  | 
by presburger  | 
|
87  | 
||
88  | 
theorem "~ (\<forall>j (i::int). j \<le> i --> (\<exists>x y. 0 \<le> x & 0 \<le> y & 3 * x + 5 * y = i))"  | 
|
| 13880 | 89  | 
by presburger  | 
90  | 
||
| 61343 | 91  | 
text\<open>Slow: about 5 seconds on a 1.6GHz machine.\<close>  | 
| 15075 | 92  | 
theorem "(\<exists>m::nat. n = 2 * m) --> (n + 1) div 2 = n div 2"  | 
93  | 
by presburger  | 
|
| 13880 | 94  | 
|
| 61343 | 95  | 
text\<open>This following theorem proves that all solutions to the  | 
| 19824 | 96  | 
recurrence relation $x_{i+2} = |x_{i+1}| - x_i$ are periodic with
 | 
97  | 
period 9. The example was brought to our attention by John  | 
|
98  | 
Harrison. It does does not require Presburger arithmetic but merely  | 
|
99  | 
quantifier-free linear arithmetic and holds for the rationals as well.  | 
|
100  | 
||
| 61343 | 101  | 
Warning: it takes (in 2006) over 4.2 minutes!\<close>  | 
| 19824 | 102  | 
|
| 61945 | 103  | 
lemma "\<lbrakk> x3 = \<bar>x2\<bar> - x1; x4 = \<bar>x3\<bar> - x2; x5 = \<bar>x4\<bar> - x3;  | 
104  | 
x6 = \<bar>x5\<bar> - x4; x7 = \<bar>x6\<bar> - x5; x8 = \<bar>x7\<bar> - x6;  | 
|
105  | 
x9 = \<bar>x8\<bar> - x7; x10 = \<bar>x9\<bar> - x8; x11 = \<bar>x10\<bar> - x9 \<rbrakk>  | 
|
| 19824 | 106  | 
\<Longrightarrow> x1 = x10 & x2 = (x11::int)"  | 
107  | 
by arith  | 
|
108  | 
||
| 15075 | 109  | 
end  |