src/HOL/ex/PresburgerEx.thy
author huffman
Tue, 15 Jan 2008 02:18:01 +0100
changeset 25913 e1b6521c1f94
parent 25801 331d8ce79ee2
child 29705 a1ecdd8cf81c
permissions -rw-r--r--
declare cpair_strict [simp]
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
13880
4f7f30f68926 Added examples for Presburger arithmetic.
berghofe
parents:
diff changeset
     1
(*  Title:      HOL/ex/PresburgerEx.thy
4f7f30f68926 Added examples for Presburger arithmetic.
berghofe
parents:
diff changeset
     2
    ID:         $Id$
4f7f30f68926 Added examples for Presburger arithmetic.
berghofe
parents:
diff changeset
     3
    Author:     Amine Chaieb, TU Muenchen
17388
495c799df31d tuned headers etc.;
wenzelm
parents: 16417
diff changeset
     4
*)
13880
4f7f30f68926 Added examples for Presburger arithmetic.
berghofe
parents:
diff changeset
     5
17388
495c799df31d tuned headers etc.;
wenzelm
parents: 16417
diff changeset
     6
header {* Some examples for Presburger Arithmetic *}
13880
4f7f30f68926 Added examples for Presburger arithmetic.
berghofe
parents:
diff changeset
     7
23462
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents: 23323
diff changeset
     8
theory PresburgerEx
24402
382f67ffbda5 imports Presburger; no need for Main
chaieb
parents: 23462
diff changeset
     9
imports Presburger
23462
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents: 23323
diff changeset
    10
begin
23323
2274edb9a8b2 Added more examples
chaieb
parents: 20663
diff changeset
    11
25801
331d8ce79ee2 Tuned (type information in Lemmas)
chaieb
parents: 24402
diff changeset
    12
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
2274edb9a8b2 Added more examples
chaieb
parents: 20663
diff changeset
    13
lemma "(0::nat) < emBits mod 8 \<Longrightarrow> 8 + emBits div 8 * 8 - emBits = 8 - emBits mod 8" 
2274edb9a8b2 Added more examples
chaieb
parents: 20663
diff changeset
    14
by presburger
2274edb9a8b2 Added more examples
chaieb
parents: 20663
diff changeset
    15
lemma "(0::nat) < emBits mod 8 \<Longrightarrow> 8 + emBits div 8 * 8 - emBits = 8 - emBits mod 8" 
2274edb9a8b2 Added more examples
chaieb
parents: 20663
diff changeset
    16
by presburger
13880
4f7f30f68926 Added examples for Presburger arithmetic.
berghofe
parents:
diff changeset
    17
14353
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 13880
diff changeset
    18
theorem "(\<forall>(y::int). 3 dvd y) ==> \<forall>(x::int). b < x --> a \<le> x"
13880
4f7f30f68926 Added examples for Presburger arithmetic.
berghofe
parents:
diff changeset
    19
  by presburger
4f7f30f68926 Added examples for Presburger arithmetic.
berghofe
parents:
diff changeset
    20
4f7f30f68926 Added examples for Presburger arithmetic.
berghofe
parents:
diff changeset
    21
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
    22
  (\<exists>(x::int).  2*x =  y) & (\<exists>(k::int). 3*k = z)"
13880
4f7f30f68926 Added examples for Presburger arithmetic.
berghofe
parents:
diff changeset
    23
  by presburger
4f7f30f68926 Added examples for Presburger arithmetic.
berghofe
parents:
diff changeset
    24
4f7f30f68926 Added examples for Presburger arithmetic.
berghofe
parents:
diff changeset
    25
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
    26
  2 dvd (y::int) ==> (\<exists>(x::int).  2*x =  y) & (\<exists>(k::int). 3*k = z)"
13880
4f7f30f68926 Added examples for Presburger arithmetic.
berghofe
parents:
diff changeset
    27
  by presburger
4f7f30f68926 Added examples for Presburger arithmetic.
berghofe
parents:
diff changeset
    28
15075
a6cd1a454751 minor formatting fixes
webertj
parents: 14981
diff changeset
    29
theorem "\<forall>(x::nat). \<exists>(y::nat). (0::nat) \<le> 5 --> y = 5 + x "
13880
4f7f30f68926 Added examples for Presburger arithmetic.
berghofe
parents:
diff changeset
    30
  by presburger
4f7f30f68926 Added examples for Presburger arithmetic.
berghofe
parents:
diff changeset
    31
20663
2024d9f7df9c updated timings;
wenzelm
parents: 19824
diff changeset
    32
text{*Slow: about 7 seconds on a 1.6GHz machine.*}
15075
a6cd1a454751 minor formatting fixes
webertj
parents: 14981
diff changeset
    33
theorem "\<forall>(x::nat). \<exists>(y::nat). y = 5 + x | x div 6 + 1= 2"
a6cd1a454751 minor formatting fixes
webertj
parents: 14981
diff changeset
    34
  by presburger
a6cd1a454751 minor formatting fixes
webertj
parents: 14981
diff changeset
    35
a6cd1a454751 minor formatting fixes
webertj
parents: 14981
diff changeset
    36
theorem "\<exists>(x::int). 0 < x"
13880
4f7f30f68926 Added examples for Presburger arithmetic.
berghofe
parents:
diff changeset
    37
  by presburger
4f7f30f68926 Added examples for Presburger arithmetic.
berghofe
parents:
diff changeset
    38
15075
a6cd1a454751 minor formatting fixes
webertj
parents: 14981
diff changeset
    39
theorem "\<forall>(x::int) y. x < y --> 2 * x + 1 < 2 * y"
a6cd1a454751 minor formatting fixes
webertj
parents: 14981
diff changeset
    40
  by presburger
13880
4f7f30f68926 Added examples for Presburger arithmetic.
berghofe
parents:
diff changeset
    41
 
15075
a6cd1a454751 minor formatting fixes
webertj
parents: 14981
diff changeset
    42
theorem "\<forall>(x::int) y. 2 * x + 1 \<noteq> 2 * y"
a6cd1a454751 minor formatting fixes
webertj
parents: 14981
diff changeset
    43
  by presburger
13880
4f7f30f68926 Added examples for Presburger arithmetic.
berghofe
parents:
diff changeset
    44
 
15075
a6cd1a454751 minor formatting fixes
webertj
parents: 14981
diff changeset
    45
theorem "\<exists>(x::int) y. 0 < x  & 0 \<le> y  & 3 * x - 5 * y = 1"
a6cd1a454751 minor formatting fixes
webertj
parents: 14981
diff changeset
    46
  by presburger
13880
4f7f30f68926 Added examples for Presburger arithmetic.
berghofe
parents:
diff changeset
    47
14353
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 13880
diff changeset
    48
theorem "~ (\<exists>(x::int) (y::int) (z::int). 4*x + (-6::int)*y = 1)"
13880
4f7f30f68926 Added examples for Presburger arithmetic.
berghofe
parents:
diff changeset
    49
  by presburger
4f7f30f68926 Added examples for Presburger arithmetic.
berghofe
parents:
diff changeset
    50
14353
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 13880
diff changeset
    51
theorem "\<forall>(x::int). b < x --> a \<le> x"
23323
2274edb9a8b2 Added more examples
chaieb
parents: 20663
diff changeset
    52
  apply (presburger elim)
13880
4f7f30f68926 Added examples for Presburger arithmetic.
berghofe
parents:
diff changeset
    53
  oops
4f7f30f68926 Added examples for Presburger arithmetic.
berghofe
parents:
diff changeset
    54
14353
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 13880
diff changeset
    55
theorem "~ (\<exists>(x::int). False)"
13880
4f7f30f68926 Added examples for Presburger arithmetic.
berghofe
parents:
diff changeset
    56
  by presburger
4f7f30f68926 Added examples for Presburger arithmetic.
berghofe
parents:
diff changeset
    57
14353
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 13880
diff changeset
    58
theorem "\<forall>(x::int). (a::int) < 3 * x --> b < 3 * x"
23323
2274edb9a8b2 Added more examples
chaieb
parents: 20663
diff changeset
    59
  apply (presburger elim)
13880
4f7f30f68926 Added examples for Presburger arithmetic.
berghofe
parents:
diff changeset
    60
  oops
4f7f30f68926 Added examples for Presburger arithmetic.
berghofe
parents:
diff changeset
    61
15075
a6cd1a454751 minor formatting fixes
webertj
parents: 14981
diff changeset
    62
theorem "\<forall>(x::int). (2 dvd x) --> (\<exists>(y::int). x = 2*y)"
a6cd1a454751 minor formatting fixes
webertj
parents: 14981
diff changeset
    63
  by presburger 
13880
4f7f30f68926 Added examples for Presburger arithmetic.
berghofe
parents:
diff changeset
    64
15075
a6cd1a454751 minor formatting fixes
webertj
parents: 14981
diff changeset
    65
theorem "\<forall>(x::int). (2 dvd x) --> (\<exists>(y::int). x = 2*y)"
a6cd1a454751 minor formatting fixes
webertj
parents: 14981
diff changeset
    66
  by presburger 
13880
4f7f30f68926 Added examples for Presburger arithmetic.
berghofe
parents:
diff changeset
    67
15075
a6cd1a454751 minor formatting fixes
webertj
parents: 14981
diff changeset
    68
theorem "\<forall>(x::int). (2 dvd x) = (\<exists>(y::int). x = 2*y)"
a6cd1a454751 minor formatting fixes
webertj
parents: 14981
diff changeset
    69
  by presburger 
14353
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 13880
diff changeset
    70
15075
a6cd1a454751 minor formatting fixes
webertj
parents: 14981
diff changeset
    71
theorem "\<forall>(x::int). ((2 dvd x) = (\<forall>(y::int). x \<noteq> 2*y + 1))"
a6cd1a454751 minor formatting fixes
webertj
parents: 14981
diff changeset
    72
  by presburger 
13880
4f7f30f68926 Added examples for Presburger arithmetic.
berghofe
parents:
diff changeset
    73
14353
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 13880
diff changeset
    74
theorem "~ (\<forall>(x::int). 
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 13880
diff changeset
    75
            ((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
    76
             (\<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
    77
             --> 0 < x | ((~ 3 dvd x) &(x + 8 = 0))))"
13880
4f7f30f68926 Added examples for Presburger arithmetic.
berghofe
parents:
diff changeset
    78
  by presburger
4f7f30f68926 Added examples for Presburger arithmetic.
berghofe
parents:
diff changeset
    79
 
15075
a6cd1a454751 minor formatting fixes
webertj
parents: 14981
diff changeset
    80
theorem "~ (\<forall>(i::int). 4 \<le> i --> (\<exists>x y. 0 \<le> x & 0 \<le> y & 3 * x + 5 * y = i))"
13880
4f7f30f68926 Added examples for Presburger arithmetic.
berghofe
parents:
diff changeset
    81
  by presburger
4f7f30f68926 Added examples for Presburger arithmetic.
berghofe
parents:
diff changeset
    82
15075
a6cd1a454751 minor formatting fixes
webertj
parents: 14981
diff changeset
    83
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
    84
  by presburger
13880
4f7f30f68926 Added examples for Presburger arithmetic.
berghofe
parents:
diff changeset
    85
15075
a6cd1a454751 minor formatting fixes
webertj
parents: 14981
diff changeset
    86
theorem "\<exists>(j::int). \<forall>i. j \<le> i --> (\<exists>x y. 0 \<le> x & 0 \<le> y & 3 * x + 5 * y = i)"
a6cd1a454751 minor formatting fixes
webertj
parents: 14981
diff changeset
    87
  by presburger
a6cd1a454751 minor formatting fixes
webertj
parents: 14981
diff changeset
    88
a6cd1a454751 minor formatting fixes
webertj
parents: 14981
diff changeset
    89
theorem "~ (\<forall>j (i::int). j \<le> i --> (\<exists>x y. 0 \<le> x & 0 \<le> y & 3 * x + 5 * y = i))"
13880
4f7f30f68926 Added examples for Presburger arithmetic.
berghofe
parents:
diff changeset
    90
  by presburger
4f7f30f68926 Added examples for Presburger arithmetic.
berghofe
parents:
diff changeset
    91
20663
2024d9f7df9c updated timings;
wenzelm
parents: 19824
diff changeset
    92
text{*Slow: about 5 seconds on a 1.6GHz machine.*}
15075
a6cd1a454751 minor formatting fixes
webertj
parents: 14981
diff changeset
    93
theorem "(\<exists>m::nat. n = 2 * m) --> (n + 1) div 2 = n div 2"
a6cd1a454751 minor formatting fixes
webertj
parents: 14981
diff changeset
    94
  by presburger
13880
4f7f30f68926 Added examples for Presburger arithmetic.
berghofe
parents:
diff changeset
    95
19824
fafceecebef0 added John's example
nipkow
parents: 17388
diff changeset
    96
text{* This following theorem proves that all solutions to the
fafceecebef0 added John's example
nipkow
parents: 17388
diff changeset
    97
recurrence relation $x_{i+2} = |x_{i+1}| - x_i$ are periodic with
fafceecebef0 added John's example
nipkow
parents: 17388
diff changeset
    98
period 9.  The example was brought to our attention by John
fafceecebef0 added John's example
nipkow
parents: 17388
diff changeset
    99
Harrison. It does does not require Presburger arithmetic but merely
fafceecebef0 added John's example
nipkow
parents: 17388
diff changeset
   100
quantifier-free linear arithmetic and holds for the rationals as well.
fafceecebef0 added John's example
nipkow
parents: 17388
diff changeset
   101
20663
2024d9f7df9c updated timings;
wenzelm
parents: 19824
diff changeset
   102
Warning: it takes (in 2006) over 4.2 minutes! *}
19824
fafceecebef0 added John's example
nipkow
parents: 17388
diff changeset
   103
fafceecebef0 added John's example
nipkow
parents: 17388
diff changeset
   104
lemma "\<lbrakk> x3 = abs x2 - x1; x4 = abs x3 - x2; x5 = abs x4 - x3;
fafceecebef0 added John's example
nipkow
parents: 17388
diff changeset
   105
         x6 = abs x5 - x4; x7 = abs x6 - x5; x8 = abs x7 - x6;
fafceecebef0 added John's example
nipkow
parents: 17388
diff changeset
   106
         x9 = abs x8 - x7; x10 = abs x9 - x8; x11 = abs x10 - x9 \<rbrakk>
fafceecebef0 added John's example
nipkow
parents: 17388
diff changeset
   107
 \<Longrightarrow> x1 = x10 & x2 = (x11::int)"
fafceecebef0 added John's example
nipkow
parents: 17388
diff changeset
   108
by arith
fafceecebef0 added John's example
nipkow
parents: 17388
diff changeset
   109
15075
a6cd1a454751 minor formatting fixes
webertj
parents: 14981
diff changeset
   110
end