src/HOL/ex/PresburgerEx.thy
author wenzelm
Wed, 09 Jun 2004 18:52:42 +0200
changeset 14898 a25550451b51
parent 14758 af3b71a46a1c
child 14981 e73f8140af78
permissions -rw-r--r--
Url.File;
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
4f7f30f68926 Added examples for Presburger arithmetic.
berghofe
parents:
diff changeset
     4
    License:    GPL (GNU GENERAL PUBLIC LICENSE)
4f7f30f68926 Added examples for Presburger arithmetic.
berghofe
parents:
diff changeset
     5
4f7f30f68926 Added examples for Presburger arithmetic.
berghofe
parents:
diff changeset
     6
Some examples for Presburger Arithmetic
4f7f30f68926 Added examples for Presburger arithmetic.
berghofe
parents:
diff changeset
     7
*)
4f7f30f68926 Added examples for Presburger arithmetic.
berghofe
parents:
diff changeset
     8
4f7f30f68926 Added examples for Presburger arithmetic.
berghofe
parents:
diff changeset
     9
theory PresburgerEx = Main:
4f7f30f68926 Added examples for Presburger arithmetic.
berghofe
parents:
diff changeset
    10
14353
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 13880
diff changeset
    11
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
    12
  by presburger
4f7f30f68926 Added examples for Presburger arithmetic.
berghofe
parents:
diff changeset
    13
4f7f30f68926 Added examples for Presburger arithmetic.
berghofe
parents:
diff changeset
    14
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
    15
  (\<exists>(x::int).  2*x =  y) & (\<exists>(k::int). 3*k = z)"
13880
4f7f30f68926 Added examples for Presburger arithmetic.
berghofe
parents:
diff changeset
    16
  by presburger
4f7f30f68926 Added examples for Presburger arithmetic.
berghofe
parents:
diff changeset
    17
4f7f30f68926 Added examples for Presburger arithmetic.
berghofe
parents:
diff changeset
    18
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
    19
  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
    20
  by presburger
4f7f30f68926 Added examples for Presburger arithmetic.
berghofe
parents:
diff changeset
    21
14353
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 13880
diff changeset
    22
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
    23
  by presburger
4f7f30f68926 Added examples for Presburger arithmetic.
berghofe
parents:
diff changeset
    24
14353
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 13880
diff changeset
    25
text{*Very slow: about 55 seconds on a 1.8GHz machine.*}
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 13880
diff changeset
    26
theorem "\<forall>(x::nat). \<exists>(y::nat). y = 5 + x | x div 6 + 1= 2";
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
14353
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 13880
diff changeset
    29
theorem "\<exists>(x::int). 0 < x" by presburger
13880
4f7f30f68926 Added examples for Presburger arithmetic.
berghofe
parents:
diff changeset
    30
14353
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 13880
diff changeset
    31
theorem "\<forall>(x::int) y. x < y --> 2 * x + 1 < 2 * y" by presburger
13880
4f7f30f68926 Added examples for Presburger arithmetic.
berghofe
parents:
diff changeset
    32
 
14353
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 13880
diff changeset
    33
theorem "\<forall>(x::int) y. 2 * x + 1 \<noteq> 2 * y" by presburger
13880
4f7f30f68926 Added examples for Presburger arithmetic.
berghofe
parents:
diff changeset
    34
 
4f7f30f68926 Added examples for Presburger arithmetic.
berghofe
parents:
diff changeset
    35
theorem
14353
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 13880
diff changeset
    36
   "\<exists>(x::int) y. 0 < x  & 0 \<le> y  & 3 * x - 5 * y = 1" by presburger
13880
4f7f30f68926 Added examples for Presburger arithmetic.
berghofe
parents:
diff changeset
    37
14353
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 13880
diff changeset
    38
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
    39
  by presburger
4f7f30f68926 Added examples for Presburger arithmetic.
berghofe
parents:
diff changeset
    40
14353
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 13880
diff changeset
    41
theorem "\<forall>(x::int). b < x --> a \<le> x"
14758
af3b71a46a1c A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
chaieb
parents: 14353
diff changeset
    42
  apply (presburger (no_quantify))
13880
4f7f30f68926 Added examples for Presburger arithmetic.
berghofe
parents:
diff changeset
    43
  oops
4f7f30f68926 Added examples for Presburger arithmetic.
berghofe
parents:
diff changeset
    44
14353
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 13880
diff changeset
    45
theorem "\<forall>(x::int). b < x --> a \<le> x"
14758
af3b71a46a1c A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
chaieb
parents: 14353
diff changeset
    46
  apply (presburger (no_quantify))
13880
4f7f30f68926 Added examples for Presburger arithmetic.
berghofe
parents:
diff changeset
    47
  oops
4f7f30f68926 Added examples for Presburger arithmetic.
berghofe
parents:
diff changeset
    48
14353
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 13880
diff changeset
    49
theorem "~ (\<exists>(x::int). False)"
13880
4f7f30f68926 Added examples for Presburger arithmetic.
berghofe
parents:
diff changeset
    50
  by presburger
4f7f30f68926 Added examples for Presburger arithmetic.
berghofe
parents:
diff changeset
    51
14353
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 13880
diff changeset
    52
theorem "\<forall>(x::int). (a::int) < 3 * x --> b < 3 * x"
14758
af3b71a46a1c A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
chaieb
parents: 14353
diff changeset
    53
  apply (presburger (no_quantify))
13880
4f7f30f68926 Added examples for Presburger arithmetic.
berghofe
parents:
diff changeset
    54
  oops
4f7f30f68926 Added examples for Presburger arithmetic.
berghofe
parents:
diff changeset
    55
14353
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 13880
diff changeset
    56
theorem "\<forall>(x::int). (2 dvd x) --> (\<exists>(y::int). x = 2*y)" by presburger 
13880
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). (2 dvd x) --> (\<exists>(y::int). x = 2*y)" by presburger 
13880
4f7f30f68926 Added examples for Presburger arithmetic.
berghofe
parents:
diff changeset
    59
14353
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 13880
diff changeset
    60
theorem "\<forall>(x::int). (2 dvd x) = (\<exists>(y::int). x = 2*y)" by presburger 
13880
4f7f30f68926 Added examples for Presburger arithmetic.
berghofe
parents:
diff changeset
    61
  
14353
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 13880
diff changeset
    62
theorem "\<forall>(x::int). ((2 dvd x) = (\<forall>(y::int). x \<noteq> 2*y + 1))" by presburger 
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 13880
diff changeset
    63
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 13880
diff changeset
    64
theorem "\<forall>(x::int). ((2 dvd x) = (\<forall>(y::int). x \<noteq> 2*y + 1))" by presburger 
13880
4f7f30f68926 Added examples for Presburger arithmetic.
berghofe
parents:
diff changeset
    65
14353
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 13880
diff changeset
    66
theorem "~ (\<forall>(x::int). 
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 13880
diff changeset
    67
            ((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
    68
             (\<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
    69
             --> 0 < x | ((~ 3 dvd x) &(x + 8 = 0))))"
13880
4f7f30f68926 Added examples for Presburger arithmetic.
berghofe
parents:
diff changeset
    70
  by presburger
4f7f30f68926 Added examples for Presburger arithmetic.
berghofe
parents:
diff changeset
    71
 
4f7f30f68926 Added examples for Presburger arithmetic.
berghofe
parents:
diff changeset
    72
theorem 
14353
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 13880
diff changeset
    73
   "~ (\<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
    74
  by presburger
4f7f30f68926 Added examples for Presburger arithmetic.
berghofe
parents:
diff changeset
    75
4f7f30f68926 Added examples for Presburger arithmetic.
berghofe
parents:
diff changeset
    76
theorem
14353
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 13880
diff changeset
    77
    "\<forall>(i::int). 8 \<le> i --> (\<exists>x y. 0 \<le> x & 0 \<le> y & 3 * x + 5 * y = i)" 
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 13880
diff changeset
    78
  by presburger
13880
4f7f30f68926 Added examples for Presburger arithmetic.
berghofe
parents:
diff changeset
    79
   
4f7f30f68926 Added examples for Presburger arithmetic.
berghofe
parents:
diff changeset
    80
theorem
14353
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 13880
diff changeset
    81
   "\<exists>(j::int). \<forall>i. j \<le> i --> (\<exists>x y. 0 \<le> x & 0 \<le> y & 3 * x + 5 * y = i)"
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 13880
diff changeset
    82
  by presburger
13880
4f7f30f68926 Added examples for Presburger arithmetic.
berghofe
parents:
diff changeset
    83
4f7f30f68926 Added examples for Presburger arithmetic.
berghofe
parents:
diff changeset
    84
theorem
14353
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 13880
diff changeset
    85
   "~ (\<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
    86
  by presburger
4f7f30f68926 Added examples for Presburger arithmetic.
berghofe
parents:
diff changeset
    87
14353
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 13880
diff changeset
    88
text{*Very slow: about 80 seconds on a 1.8GHz machine.*}
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 13880
diff changeset
    89
theorem "(\<exists>m::nat. n = 2 * m) --> (n + 1) div 2 = n div 2" by presburger
13880
4f7f30f68926 Added examples for Presburger arithmetic.
berghofe
parents:
diff changeset
    90
14353
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 13880
diff changeset
    91
theorem "(\<exists>m::int. n = 2 * m) --> (n + 1) div 2 = n div 2" by presburger
13880
4f7f30f68926 Added examples for Presburger arithmetic.
berghofe
parents:
diff changeset
    92
4f7f30f68926 Added examples for Presburger arithmetic.
berghofe
parents:
diff changeset
    93
end