| 
23193
 | 
     1  | 
(*  Title:  HOL/ex/Arith_Examples.thy
  | 
| 
 | 
     2  | 
    ID:     $Id$
  | 
| 
 | 
     3  | 
    Author: Tjark Weber
  | 
| 
 | 
     4  | 
*)
  | 
| 
 | 
     5  | 
  | 
| 
23218
 | 
     6  | 
header {* Arithmetic *}
 | 
| 
23193
 | 
     7  | 
  | 
| 
 | 
     8  | 
theory Arith_Examples imports Main begin
  | 
| 
 | 
     9  | 
  | 
| 
 | 
    10  | 
text {*
 | 
| 
23218
 | 
    11  | 
  The @{text arith} method is used frequently throughout the Isabelle
 | 
| 
23193
 | 
    12  | 
  distribution.  This file merely contains some additional tests and special
  | 
| 
 | 
    13  | 
  corner cases.  Some rather technical remarks:
  | 
| 
 | 
    14  | 
  | 
| 
23218
 | 
    15  | 
  @{ML fast_arith_tac} is a very basic version of the tactic.  It performs no
 | 
| 
23193
 | 
    16  | 
  meta-to-object-logic conversion, and only some splitting of operators.
  | 
| 
23218
 | 
    17  | 
  @{ML simple_arith_tac} performs meta-to-object-logic conversion, full
 | 
| 
 | 
    18  | 
  splitting of operators, and NNF normalization of the goal.  The @{text arith}
 | 
| 
 | 
    19  | 
  method combines them both, and tries other methods (e.g.~@{text presburger})
 | 
| 
23193
 | 
    20  | 
  as well.  This is the one that you should use in your proofs!
  | 
| 
 | 
    21  | 
  | 
| 
23218
 | 
    22  | 
  An @{text arith}-based simproc is available as well
 | 
| 
 | 
    23  | 
  (see @{ML Fast_Arith.lin_arith_prover}),
 | 
| 
23211
 | 
    24  | 
  which---for performance reasons---however
  | 
| 
23218
 | 
    25  | 
  does even less splitting than @{ML fast_arith_tac} at the moment (namely
 | 
| 
23193
 | 
    26  | 
  inequalities only).  (On the other hand, it does take apart conjunctions,
  | 
| 
23218
 | 
    27  | 
  which @{ML fast_arith_tac} currently does not do.)
 | 
| 
23193
 | 
    28  | 
*}
  | 
| 
 | 
    29  | 
  | 
| 
23196
 | 
    30  | 
(*
  | 
| 
23193
 | 
    31  | 
ML {* set trace_arith; *}
 | 
| 
23196
 | 
    32  | 
*)
  | 
| 
23193
 | 
    33  | 
  | 
| 
23218
 | 
    34  | 
subsection {* Splitting of Operators: @{term max}, @{term min}, @{term abs},
 | 
| 
23193
 | 
    35  | 
           @{term HOL.minus}, @{term nat}, @{term Divides.mod},
 | 
| 
 | 
    36  | 
           @{term Divides.div} *}
 | 
| 
 | 
    37  | 
  | 
| 
 | 
    38  | 
lemma "(i::nat) <= max i j"
  | 
| 
 | 
    39  | 
  by (tactic {* fast_arith_tac 1 *})
 | 
| 
 | 
    40  | 
  | 
| 
 | 
    41  | 
lemma "(i::int) <= max i j"
  | 
| 
 | 
    42  | 
  by (tactic {* fast_arith_tac 1 *})
 | 
| 
 | 
    43  | 
  | 
| 
 | 
    44  | 
lemma "min i j <= (i::nat)"
  | 
| 
 | 
    45  | 
  by (tactic {* fast_arith_tac 1 *})
 | 
| 
 | 
    46  | 
  | 
| 
 | 
    47  | 
lemma "min i j <= (i::int)"
  | 
| 
 | 
    48  | 
  by (tactic {* fast_arith_tac 1 *})
 | 
| 
 | 
    49  | 
  | 
| 
 | 
    50  | 
lemma "min (i::nat) j <= max i j"
  | 
| 
 | 
    51  | 
  by (tactic {* fast_arith_tac 1 *})
 | 
| 
 | 
    52  | 
  | 
| 
 | 
    53  | 
lemma "min (i::int) j <= max i j"
  | 
| 
 | 
    54  | 
  by (tactic {* fast_arith_tac 1 *})
 | 
| 
 | 
    55  | 
  | 
| 
23208
 | 
    56  | 
lemma "min (i::nat) j + max i j = i + j"
  | 
| 
 | 
    57  | 
  by (tactic {* fast_arith_tac 1 *})
 | 
| 
 | 
    58  | 
  | 
| 
 | 
    59  | 
lemma "min (i::int) j + max i j = i + j"
  | 
| 
 | 
    60  | 
  by (tactic {* fast_arith_tac 1 *})
 | 
| 
 | 
    61  | 
  | 
| 
23193
 | 
    62  | 
lemma "(i::nat) < j ==> min i j < max i j"
  | 
| 
 | 
    63  | 
  by (tactic {* fast_arith_tac 1 *})
 | 
| 
 | 
    64  | 
  | 
| 
 | 
    65  | 
lemma "(i::int) < j ==> min i j < max i j"
  | 
| 
 | 
    66  | 
  by (tactic {* fast_arith_tac 1 *})
 | 
| 
 | 
    67  | 
  | 
| 
 | 
    68  | 
lemma "(0::int) <= abs i"
  | 
| 
 | 
    69  | 
  by (tactic {* fast_arith_tac 1 *})
 | 
| 
 | 
    70  | 
  | 
| 
 | 
    71  | 
lemma "(i::int) <= abs i"
  | 
| 
 | 
    72  | 
  by (tactic {* fast_arith_tac 1 *})
 | 
| 
 | 
    73  | 
  | 
| 
 | 
    74  | 
lemma "abs (abs (i::int)) = abs i"
  | 
| 
 | 
    75  | 
  by (tactic {* fast_arith_tac 1 *})
 | 
| 
 | 
    76  | 
  | 
| 
 | 
    77  | 
text {* Also testing subgoals with bound variables. *}
 | 
| 
 | 
    78  | 
  | 
| 
 | 
    79  | 
lemma "!!x. (x::nat) <= y ==> x - y = 0"
  | 
| 
 | 
    80  | 
  by (tactic {* fast_arith_tac 1 *})
 | 
| 
 | 
    81  | 
  | 
| 
 | 
    82  | 
lemma "!!x. (x::nat) - y = 0 ==> x <= y"
  | 
| 
 | 
    83  | 
  by (tactic {* fast_arith_tac 1 *})
 | 
| 
 | 
    84  | 
  | 
| 
 | 
    85  | 
lemma "!!x. ((x::nat) <= y) = (x - y = 0)"
  | 
| 
 | 
    86  | 
  by (tactic {* simple_arith_tac 1 *})
 | 
| 
 | 
    87  | 
  | 
| 
 | 
    88  | 
lemma "[| (x::nat) < y; d < 1 |] ==> x - y = d"
  | 
| 
 | 
    89  | 
  by (tactic {* fast_arith_tac 1 *})
 | 
| 
 | 
    90  | 
  | 
| 
 | 
    91  | 
lemma "[| (x::nat) < y; d < 1 |] ==> x - y - x = d - x"
  | 
| 
 | 
    92  | 
  by (tactic {* fast_arith_tac 1 *})
 | 
| 
 | 
    93  | 
  | 
| 
 | 
    94  | 
lemma "(x::int) < y ==> x - y < 0"
  | 
| 
 | 
    95  | 
  by (tactic {* fast_arith_tac 1 *})
 | 
| 
 | 
    96  | 
  | 
| 
 | 
    97  | 
lemma "nat (i + j) <= nat i + nat j"
  | 
| 
 | 
    98  | 
  by (tactic {* fast_arith_tac 1 *})
 | 
| 
 | 
    99  | 
  | 
| 
 | 
   100  | 
lemma "i < j ==> nat (i - j) = 0"
  | 
| 
 | 
   101  | 
  by (tactic {* fast_arith_tac 1 *})
 | 
| 
 | 
   102  | 
  | 
| 
 | 
   103  | 
lemma "(i::nat) mod 0 = i"
  | 
| 
23198
 | 
   104  | 
  (* FIXME: need to replace 0 by its numeral representation *)
  | 
| 
 | 
   105  | 
  apply (subst nat_numeral_0_eq_0 [symmetric])
  | 
| 
 | 
   106  | 
  by (tactic {* fast_arith_tac 1 *})
 | 
| 
 | 
   107  | 
  | 
| 
 | 
   108  | 
lemma "(i::nat) mod 1 = 0"
  | 
| 
 | 
   109  | 
  (* FIXME: need to replace 1 by its numeral representation *)
  | 
| 
 | 
   110  | 
  apply (subst nat_numeral_1_eq_1 [symmetric])
  | 
| 
 | 
   111  | 
  by (tactic {* fast_arith_tac 1 *})
 | 
| 
23193
 | 
   112  | 
  | 
| 
23198
 | 
   113  | 
lemma "(i::nat) mod 42 <= 41"
  | 
| 
 | 
   114  | 
  by (tactic {* fast_arith_tac 1 *})
 | 
| 
 | 
   115  | 
  | 
| 
 | 
   116  | 
lemma "(i::int) mod 0 = i"
  | 
| 
 | 
   117  | 
  (* FIXME: need to replace 0 by its numeral representation *)
  | 
| 
 | 
   118  | 
  apply (subst numeral_0_eq_0 [symmetric])
  | 
| 
 | 
   119  | 
  by (tactic {* fast_arith_tac 1 *})
 | 
| 
 | 
   120  | 
  | 
| 
 | 
   121  | 
lemma "(i::int) mod 1 = 0"
  | 
| 
 | 
   122  | 
  (* FIXME: need to replace 1 by its numeral representation *)
  | 
| 
 | 
   123  | 
  apply (subst numeral_1_eq_1 [symmetric])
  | 
| 
 | 
   124  | 
  (* FIXME: arith does not know about iszero *)
  | 
| 
 | 
   125  | 
  apply (tactic {* LA_Data_Ref.pre_tac 1 *})
 | 
| 
23193
 | 
   126  | 
oops
  | 
| 
 | 
   127  | 
  | 
| 
23198
 | 
   128  | 
lemma "(i::int) mod 42 <= 41"
  | 
| 
 | 
   129  | 
  (* FIXME: arith does not know about iszero *)
  | 
| 
 | 
   130  | 
  apply (tactic {* LA_Data_Ref.pre_tac 1 *})
 | 
| 
23193
 | 
   131  | 
oops
  | 
| 
 | 
   132  | 
  | 
| 
23218
 | 
   133  | 
  | 
| 
 | 
   134  | 
subsection {* Meta-Logic *}
 | 
| 
23193
 | 
   135  | 
  | 
| 
 | 
   136  | 
lemma "x < Suc y == x <= y"
  | 
| 
 | 
   137  | 
  by (tactic {* simple_arith_tac 1 *})
 | 
| 
 | 
   138  | 
  | 
| 
 | 
   139  | 
lemma "((x::nat) == z ==> x ~= y) ==> x ~= y | z ~= y"
  | 
| 
 | 
   140  | 
  by (tactic {* simple_arith_tac 1 *})
 | 
| 
 | 
   141  | 
  | 
| 
23218
 | 
   142  | 
  | 
| 
 | 
   143  | 
subsection {* Various Other Examples *}
 | 
| 
23193
 | 
   144  | 
  | 
| 
23198
 | 
   145  | 
lemma "(x < Suc y) = (x <= y)"
  | 
| 
 | 
   146  | 
  by (tactic {* simple_arith_tac 1 *})
 | 
| 
 | 
   147  | 
  | 
| 
23193
 | 
   148  | 
lemma "[| (x::nat) < y; y < z |] ==> x < z"
  | 
| 
 | 
   149  | 
  by (tactic {* fast_arith_tac 1 *})
 | 
| 
 | 
   150  | 
  | 
| 
 | 
   151  | 
lemma "(x::nat) < y & y < z ==> x < z"
  | 
| 
 | 
   152  | 
  by (tactic {* simple_arith_tac 1 *})
 | 
| 
 | 
   153  | 
  | 
| 
23208
 | 
   154  | 
text {* This example involves no arithmetic at all, but is solved by
 | 
| 
 | 
   155  | 
  preprocessing (i.e. NNF normalization) alone. *}
  | 
| 
 | 
   156  | 
  | 
| 
 | 
   157  | 
lemma "(P::bool) = Q ==> Q = P"
  | 
| 
 | 
   158  | 
  by (tactic {* simple_arith_tac 1 *})
 | 
| 
 | 
   159  | 
  | 
| 
 | 
   160  | 
lemma "[| P = (x = 0); (~P) = (y = 0) |] ==> min (x::nat) y = 0"
  | 
| 
 | 
   161  | 
  by (tactic {* simple_arith_tac 1 *})
 | 
| 
 | 
   162  | 
  | 
| 
 | 
   163  | 
lemma "[| P = (x = 0); (~P) = (y = 0) |] ==> max (x::nat) y = x + y"
  | 
| 
 | 
   164  | 
  by (tactic {* simple_arith_tac 1 *})
 | 
| 
 | 
   165  | 
  | 
| 
23193
 | 
   166  | 
lemma "[| (x::nat) ~= y; a + 2 = b; a < y; y < b; a < x; x < b |] ==> False"
  | 
| 
 | 
   167  | 
  by (tactic {* fast_arith_tac 1 *})
 | 
| 
 | 
   168  | 
  | 
| 
 | 
   169  | 
lemma "[| (x::nat) > y; y > z; z > x |] ==> False"
  | 
| 
 | 
   170  | 
  by (tactic {* fast_arith_tac 1 *})
 | 
| 
 | 
   171  | 
  | 
| 
 | 
   172  | 
lemma "(x::nat) - 5 > y ==> y < x"
  | 
| 
 | 
   173  | 
  by (tactic {* fast_arith_tac 1 *})
 | 
| 
 | 
   174  | 
  | 
| 
 | 
   175  | 
lemma "(x::nat) ~= 0 ==> 0 < x"
  | 
| 
 | 
   176  | 
  by (tactic {* fast_arith_tac 1 *})
 | 
| 
 | 
   177  | 
  | 
| 
 | 
   178  | 
lemma "[| (x::nat) ~= y; x <= y |] ==> x < y"
  | 
| 
 | 
   179  | 
  by (tactic {* fast_arith_tac 1 *})
 | 
| 
 | 
   180  | 
  | 
| 
23196
 | 
   181  | 
lemma "[| (x::nat) < y; P (x - y) |] ==> P 0"
  | 
| 
23193
 | 
   182  | 
  by (tactic {* simple_arith_tac 1 *})
 | 
| 
 | 
   183  | 
  | 
| 
 | 
   184  | 
lemma "(x - y) - (x::nat) = (x - x) - y"
  | 
| 
 | 
   185  | 
  by (tactic {* fast_arith_tac 1 *})
 | 
| 
 | 
   186  | 
  | 
| 
 | 
   187  | 
lemma "[| (a::nat) < b; c < d |] ==> (a - b) = (c - d)"
  | 
| 
 | 
   188  | 
  by (tactic {* fast_arith_tac 1 *})
 | 
| 
 | 
   189  | 
  | 
| 
 | 
   190  | 
lemma "((a::nat) - (b - (c - (d - e)))) = (a - (b - (c - (d - e))))"
  | 
| 
 | 
   191  | 
  by (tactic {* fast_arith_tac 1 *})
 | 
| 
 | 
   192  | 
  | 
| 
23198
 | 
   193  | 
lemma "(n < m & m < n') | (n < m & m = n') | (n < n' & n' < m) |
  | 
| 
 | 
   194  | 
  (n = n' & n' < m) | (n = m & m < n') |
  | 
| 
 | 
   195  | 
  (n' < m & m < n) | (n' < m & m = n) |
  | 
| 
 | 
   196  | 
  (n' < n & n < m) | (n' = n & n < m) | (n' = m & m < n) |
  | 
| 
 | 
   197  | 
  (m < n & n < n') | (m < n & n' = n) | (m < n' & n' < n) |
  | 
| 
 | 
   198  | 
  (m = n & n < n') | (m = n' & n' < n) |
  | 
| 
 | 
   199  | 
  (n' = m & m = (n::nat))"
  | 
| 
 | 
   200  | 
(* FIXME: this should work in principle, but is extremely slow because     *)
  | 
| 
 | 
   201  | 
(*        preprocessing negates the goal and tries to compute its negation *)
  | 
| 
 | 
   202  | 
(*        normal form, which creates lots of separate cases for this       *)
  | 
| 
 | 
   203  | 
(*        disjunction of conjunctions                                      *)
  | 
| 
 | 
   204  | 
(* by (tactic {* simple_arith_tac 1 *}) *)
 | 
| 
 | 
   205  | 
oops
  | 
| 
 | 
   206  | 
  | 
| 
 | 
   207  | 
lemma "2 * (x::nat) ~= 1"
  | 
| 
23208
 | 
   208  | 
(* FIXME: this is beyond the scope of the decision procedure at the moment, *)
  | 
| 
 | 
   209  | 
(*        because its negation is satisfiable in the rationals?             *)
  | 
| 
23198
 | 
   210  | 
(* by (tactic {* fast_arith_tac 1 *}) *)
 | 
| 
 | 
   211  | 
oops
  | 
| 
 | 
   212  | 
  | 
| 
 | 
   213  | 
text {* Constants. *}
 | 
| 
 | 
   214  | 
  | 
| 
 | 
   215  | 
lemma "(0::nat) < 1"
  | 
| 
 | 
   216  | 
  by (tactic {* fast_arith_tac 1 *})
 | 
| 
 | 
   217  | 
  | 
| 
 | 
   218  | 
lemma "(0::int) < 1"
  | 
| 
 | 
   219  | 
  by (tactic {* fast_arith_tac 1 *})
 | 
| 
 | 
   220  | 
  | 
| 
 | 
   221  | 
lemma "(47::nat) + 11 < 08 * 15"
  | 
| 
 | 
   222  | 
  by (tactic {* fast_arith_tac 1 *})
 | 
| 
 | 
   223  | 
  | 
| 
 | 
   224  | 
lemma "(47::int) + 11 < 08 * 15"
  | 
| 
 | 
   225  | 
  by (tactic {* fast_arith_tac 1 *})
 | 
| 
 | 
   226  | 
  | 
| 
23193
 | 
   227  | 
text {* Splitting of inequalities of different type. *}
 | 
| 
 | 
   228  | 
  | 
| 
 | 
   229  | 
lemma "[| (a::nat) ~= b; (i::int) ~= j; a < 2; b < 2 |] ==>
  | 
| 
 | 
   230  | 
  a + b <= nat (max (abs i) (abs j))"
  | 
| 
 | 
   231  | 
  by (tactic {* fast_arith_tac 1 *})
 | 
| 
 | 
   232  | 
  | 
| 
23198
 | 
   233  | 
text {* Again, but different order. *}
 | 
| 
 | 
   234  | 
  | 
| 
23193
 | 
   235  | 
lemma "[| (i::int) ~= j; (a::nat) ~= b; a < 2; b < 2 |] ==>
  | 
| 
 | 
   236  | 
  a + b <= nat (max (abs i) (abs j))"
  | 
| 
 | 
   237  | 
  by (tactic {* fast_arith_tac 1 *})
 | 
| 
 | 
   238  | 
  | 
| 
23196
 | 
   239  | 
(*
  | 
| 
23193
 | 
   240  | 
ML {* reset trace_arith; *}
 | 
| 
23196
 | 
   241  | 
*)
  | 
| 
23193
 | 
   242  | 
  | 
| 
 | 
   243  | 
end
  |