| 
23164
 | 
     1  | 
(*  Title:      HOL/IntArith.thy
  | 
| 
 | 
     2  | 
    ID:         $Id$
  | 
| 
 | 
     3  | 
    Authors:    Larry Paulson and Tobias Nipkow
  | 
| 
 | 
     4  | 
*)
  | 
| 
 | 
     5  | 
  | 
| 
 | 
     6  | 
header {* Integer arithmetic *}
 | 
| 
 | 
     7  | 
  | 
| 
 | 
     8  | 
theory IntArith
  | 
| 
 | 
     9  | 
imports Numeral Wellfounded_Relations
  | 
| 
23263
 | 
    10  | 
uses
  | 
| 
 | 
    11  | 
  "~~/src/Provers/Arith/assoc_fold.ML"
  | 
| 
 | 
    12  | 
  "~~/src/Provers/Arith/cancel_numerals.ML"
  | 
| 
 | 
    13  | 
  "~~/src/Provers/Arith/combine_numerals.ML"
  | 
| 
 | 
    14  | 
  ("int_arith1.ML")
 | 
| 
23164
 | 
    15  | 
begin
  | 
| 
 | 
    16  | 
  | 
| 
 | 
    17  | 
text{*Duplicate: can't understand why it's necessary*}
 | 
| 
 | 
    18  | 
declare numeral_0_eq_0 [simp]
  | 
| 
 | 
    19  | 
  | 
| 
 | 
    20  | 
  | 
| 
 | 
    21  | 
subsection{*Inequality Reasoning for the Arithmetic Simproc*}
 | 
| 
 | 
    22  | 
  | 
| 
 | 
    23  | 
lemma add_numeral_0: "Numeral0 + a = (a::'a::number_ring)"
  | 
| 
 | 
    24  | 
by simp 
  | 
| 
 | 
    25  | 
  | 
| 
 | 
    26  | 
lemma add_numeral_0_right: "a + Numeral0 = (a::'a::number_ring)"
  | 
| 
 | 
    27  | 
by simp
  | 
| 
 | 
    28  | 
  | 
| 
 | 
    29  | 
lemma mult_numeral_1: "Numeral1 * a = (a::'a::number_ring)"
  | 
| 
 | 
    30  | 
by simp 
  | 
| 
 | 
    31  | 
  | 
| 
 | 
    32  | 
lemma mult_numeral_1_right: "a * Numeral1 = (a::'a::number_ring)"
  | 
| 
 | 
    33  | 
by simp
  | 
| 
 | 
    34  | 
  | 
| 
 | 
    35  | 
lemma divide_numeral_1: "a / Numeral1 = (a::'a::{number_ring,field})"
 | 
| 
 | 
    36  | 
by simp
  | 
| 
 | 
    37  | 
  | 
| 
 | 
    38  | 
lemma inverse_numeral_1:
  | 
| 
 | 
    39  | 
  "inverse Numeral1 = (Numeral1::'a::{number_ring,field})"
 | 
| 
 | 
    40  | 
by simp
  | 
| 
 | 
    41  | 
  | 
| 
 | 
    42  | 
text{*Theorem lists for the cancellation simprocs. The use of binary numerals
 | 
| 
 | 
    43  | 
for 0 and 1 reduces the number of special cases.*}
  | 
| 
 | 
    44  | 
  | 
| 
 | 
    45  | 
lemmas add_0s = add_numeral_0 add_numeral_0_right
  | 
| 
 | 
    46  | 
lemmas mult_1s = mult_numeral_1 mult_numeral_1_right 
  | 
| 
 | 
    47  | 
                 mult_minus1 mult_minus1_right
  | 
| 
 | 
    48  | 
  | 
| 
 | 
    49  | 
  | 
| 
 | 
    50  | 
subsection{*Special Arithmetic Rules for Abstract 0 and 1*}
 | 
| 
 | 
    51  | 
  | 
| 
 | 
    52  | 
text{*Arithmetic computations are defined for binary literals, which leaves 0
 | 
| 
 | 
    53  | 
and 1 as special cases. Addition already has rules for 0, but not 1.
  | 
| 
 | 
    54  | 
Multiplication and unary minus already have rules for both 0 and 1.*}
  | 
| 
 | 
    55  | 
  | 
| 
 | 
    56  | 
  | 
| 
 | 
    57  | 
lemma binop_eq: "[|f x y = g x y; x = x'; y = y'|] ==> f x' y' = g x' y'"
  | 
| 
 | 
    58  | 
by simp
  | 
| 
 | 
    59  | 
  | 
| 
 | 
    60  | 
  | 
| 
 | 
    61  | 
lemmas add_number_of_eq = number_of_add [symmetric]
  | 
| 
 | 
    62  | 
  | 
| 
 | 
    63  | 
text{*Allow 1 on either or both sides*}
 | 
| 
 | 
    64  | 
lemma one_add_one_is_two: "1 + 1 = (2::'a::number_ring)"
  | 
| 
 | 
    65  | 
by (simp del: numeral_1_eq_1 add: numeral_1_eq_1 [symmetric] add_number_of_eq)
  | 
| 
 | 
    66  | 
  | 
| 
 | 
    67  | 
lemmas add_special =
  | 
| 
 | 
    68  | 
    one_add_one_is_two
  | 
| 
 | 
    69  | 
    binop_eq [of "op +", OF add_number_of_eq numeral_1_eq_1 refl, standard]
  | 
| 
 | 
    70  | 
    binop_eq [of "op +", OF add_number_of_eq refl numeral_1_eq_1, standard]
  | 
| 
 | 
    71  | 
  | 
| 
 | 
    72  | 
text{*Allow 1 on either or both sides (1-1 already simplifies to 0)*}
 | 
| 
 | 
    73  | 
lemmas diff_special =
  | 
| 
 | 
    74  | 
    binop_eq [of "op -", OF diff_number_of_eq numeral_1_eq_1 refl, standard]
  | 
| 
 | 
    75  | 
    binop_eq [of "op -", OF diff_number_of_eq refl numeral_1_eq_1, standard]
  | 
| 
 | 
    76  | 
  | 
| 
 | 
    77  | 
text{*Allow 0 or 1 on either side with a binary numeral on the other*}
 | 
| 
 | 
    78  | 
lemmas eq_special =
  | 
| 
 | 
    79  | 
    binop_eq [of "op =", OF eq_number_of_eq numeral_0_eq_0 refl, standard]
  | 
| 
 | 
    80  | 
    binop_eq [of "op =", OF eq_number_of_eq numeral_1_eq_1 refl, standard]
  | 
| 
 | 
    81  | 
    binop_eq [of "op =", OF eq_number_of_eq refl numeral_0_eq_0, standard]
  | 
| 
 | 
    82  | 
    binop_eq [of "op =", OF eq_number_of_eq refl numeral_1_eq_1, standard]
  | 
| 
 | 
    83  | 
  | 
| 
 | 
    84  | 
text{*Allow 0 or 1 on either side with a binary numeral on the other*}
 | 
| 
 | 
    85  | 
lemmas less_special =
  | 
| 
 | 
    86  | 
  binop_eq [of "op <", OF less_number_of_eq_neg numeral_0_eq_0 refl, standard]
  | 
| 
 | 
    87  | 
  binop_eq [of "op <", OF less_number_of_eq_neg numeral_1_eq_1 refl, standard]
  | 
| 
 | 
    88  | 
  binop_eq [of "op <", OF less_number_of_eq_neg refl numeral_0_eq_0, standard]
  | 
| 
 | 
    89  | 
  binop_eq [of "op <", OF less_number_of_eq_neg refl numeral_1_eq_1, standard]
  | 
| 
 | 
    90  | 
  | 
| 
 | 
    91  | 
text{*Allow 0 or 1 on either side with a binary numeral on the other*}
 | 
| 
 | 
    92  | 
lemmas le_special =
  | 
| 
 | 
    93  | 
    binop_eq [of "op \<le>", OF le_number_of_eq numeral_0_eq_0 refl, standard]
  | 
| 
 | 
    94  | 
    binop_eq [of "op \<le>", OF le_number_of_eq numeral_1_eq_1 refl, standard]
  | 
| 
 | 
    95  | 
    binop_eq [of "op \<le>", OF le_number_of_eq refl numeral_0_eq_0, standard]
  | 
| 
 | 
    96  | 
    binop_eq [of "op \<le>", OF le_number_of_eq refl numeral_1_eq_1, standard]
  | 
| 
 | 
    97  | 
  | 
| 
 | 
    98  | 
lemmas arith_special[simp] = 
  | 
| 
 | 
    99  | 
       add_special diff_special eq_special less_special le_special
  | 
| 
 | 
   100  | 
  | 
| 
 | 
   101  | 
  | 
| 
 | 
   102  | 
lemma min_max_01: "min (0::int) 1 = 0 & min (1::int) 0 = 0 &
  | 
| 
 | 
   103  | 
                   max (0::int) 1 = 1 & max (1::int) 0 = 1"
  | 
| 
 | 
   104  | 
by(simp add:min_def max_def)
  | 
| 
 | 
   105  | 
  | 
| 
 | 
   106  | 
lemmas min_max_special[simp] =
  | 
| 
 | 
   107  | 
 min_max_01
  | 
| 
 | 
   108  | 
 max_def[of "0::int" "number_of v", standard, simp]
  | 
| 
 | 
   109  | 
 min_def[of "0::int" "number_of v", standard, simp]
  | 
| 
 | 
   110  | 
 max_def[of "number_of u" "0::int", standard, simp]
  | 
| 
 | 
   111  | 
 min_def[of "number_of u" "0::int", standard, simp]
  | 
| 
 | 
   112  | 
 max_def[of "1::int" "number_of v", standard, simp]
  | 
| 
 | 
   113  | 
 min_def[of "1::int" "number_of v", standard, simp]
  | 
| 
 | 
   114  | 
 max_def[of "number_of u" "1::int", standard, simp]
  | 
| 
 | 
   115  | 
 min_def[of "number_of u" "1::int", standard, simp]
  | 
| 
 | 
   116  | 
  | 
| 
 | 
   117  | 
use "int_arith1.ML"
  | 
| 
 | 
   118  | 
setup int_arith_setup
  | 
| 
 | 
   119  | 
  | 
| 
 | 
   120  | 
  | 
| 
 | 
   121  | 
subsection{*Lemmas About Small Numerals*}
 | 
| 
 | 
   122  | 
  | 
| 
 | 
   123  | 
lemma of_int_m1 [simp]: "of_int -1 = (-1 :: 'a :: number_ring)"
  | 
| 
 | 
   124  | 
proof -
  | 
| 
 | 
   125  | 
  have "(of_int -1 :: 'a) = of_int (- 1)" by simp
  | 
| 
 | 
   126  | 
  also have "... = - of_int 1" by (simp only: of_int_minus)
  | 
| 
 | 
   127  | 
  also have "... = -1" by simp
  | 
| 
 | 
   128  | 
  finally show ?thesis .
  | 
| 
 | 
   129  | 
qed
  | 
| 
 | 
   130  | 
  | 
| 
 | 
   131  | 
lemma abs_minus_one [simp]: "abs (-1) = (1::'a::{ordered_idom,number_ring})"
 | 
| 
 | 
   132  | 
by (simp add: abs_if)
  | 
| 
 | 
   133  | 
  | 
| 
 | 
   134  | 
lemma abs_power_minus_one [simp]:
  | 
| 
 | 
   135  | 
     "abs(-1 ^ n) = (1::'a::{ordered_idom,number_ring,recpower})"
 | 
| 
 | 
   136  | 
by (simp add: power_abs)
  | 
| 
 | 
   137  | 
  | 
| 
 | 
   138  | 
lemma of_int_number_of_eq:
  | 
| 
 | 
   139  | 
     "of_int (number_of v) = (number_of v :: 'a :: number_ring)"
  | 
| 
 | 
   140  | 
by (simp add: number_of_eq) 
  | 
| 
 | 
   141  | 
  | 
| 
 | 
   142  | 
text{*Lemmas for specialist use, NOT as default simprules*}
 | 
| 
 | 
   143  | 
lemma mult_2: "2 * z = (z+z::'a::number_ring)"
  | 
| 
 | 
   144  | 
proof -
  | 
| 
 | 
   145  | 
  have "2*z = (1 + 1)*z" by simp
  | 
| 
 | 
   146  | 
  also have "... = z+z" by (simp add: left_distrib)
  | 
| 
 | 
   147  | 
  finally show ?thesis .
  | 
| 
 | 
   148  | 
qed
  | 
| 
 | 
   149  | 
  | 
| 
 | 
   150  | 
lemma mult_2_right: "z * 2 = (z+z::'a::number_ring)"
  | 
| 
 | 
   151  | 
by (subst mult_commute, rule mult_2)
  | 
| 
 | 
   152  | 
  | 
| 
 | 
   153  | 
  | 
| 
 | 
   154  | 
subsection{*More Inequality Reasoning*}
 | 
| 
 | 
   155  | 
  | 
| 
 | 
   156  | 
lemma zless_add1_eq: "(w < z + (1::int)) = (w<z | w=z)"
  | 
| 
 | 
   157  | 
by arith
  | 
| 
 | 
   158  | 
  | 
| 
 | 
   159  | 
lemma add1_zle_eq: "(w + (1::int) \<le> z) = (w<z)"
  | 
| 
 | 
   160  | 
by arith
  | 
| 
 | 
   161  | 
  | 
| 
 | 
   162  | 
lemma zle_diff1_eq [simp]: "(w \<le> z - (1::int)) = (w<z)"
  | 
| 
 | 
   163  | 
by arith
  | 
| 
 | 
   164  | 
  | 
| 
 | 
   165  | 
lemma zle_add1_eq_le [simp]: "(w < z + (1::int)) = (w\<le>z)"
  | 
| 
 | 
   166  | 
by arith
  | 
| 
 | 
   167  | 
  | 
| 
 | 
   168  | 
lemma int_one_le_iff_zero_less: "((1::int) \<le> z) = (0 < z)"
  | 
| 
 | 
   169  | 
by arith
  | 
| 
 | 
   170  | 
  | 
| 
 | 
   171  | 
  | 
| 
 | 
   172  | 
subsection{*The Functions @{term nat} and @{term int}*}
 | 
| 
 | 
   173  | 
  | 
| 
 | 
   174  | 
text{*Simplify the terms @{term "int 0"}, @{term "int(Suc 0)"} and
 | 
| 
 | 
   175  | 
  @{term "w + - z"}*}
 | 
| 
 | 
   176  | 
declare Zero_int_def [symmetric, simp]
  | 
| 
 | 
   177  | 
declare One_int_def [symmetric, simp]
  | 
| 
 | 
   178  | 
  | 
| 
 | 
   179  | 
lemmas diff_int_def_symmetric = diff_int_def [symmetric, simp]
  | 
| 
 | 
   180  | 
  | 
| 
 | 
   181  | 
lemma nat_0: "nat 0 = 0"
  | 
| 
 | 
   182  | 
by (simp add: nat_eq_iff)
  | 
| 
 | 
   183  | 
  | 
| 
 | 
   184  | 
lemma nat_1: "nat 1 = Suc 0"
  | 
| 
 | 
   185  | 
by (subst nat_eq_iff, simp)
  | 
| 
 | 
   186  | 
  | 
| 
 | 
   187  | 
lemma nat_2: "nat 2 = Suc (Suc 0)"
  | 
| 
 | 
   188  | 
by (subst nat_eq_iff, simp)
  | 
| 
 | 
   189  | 
  | 
| 
 | 
   190  | 
lemma one_less_nat_eq [simp]: "(Suc 0 < nat z) = (1 < z)"
  | 
| 
 | 
   191  | 
apply (insert zless_nat_conj [of 1 z])
  | 
| 
 | 
   192  | 
apply (auto simp add: nat_1)
  | 
| 
 | 
   193  | 
done
  | 
| 
 | 
   194  | 
  | 
| 
 | 
   195  | 
text{*This simplifies expressions of the form @{term "int n = z"} where
 | 
| 
 | 
   196  | 
      z is an integer literal.*}
  | 
| 
 | 
   197  | 
lemmas int_eq_iff_number_of [simp] = int_eq_iff [of _ "number_of v", standard]
  | 
| 
 | 
   198  | 
  | 
| 
 | 
   199  | 
  | 
| 
 | 
   200  | 
lemma split_nat [arith_split]:
  | 
| 
 | 
   201  | 
  "P(nat(i::int)) = ((\<forall>n. i = int n \<longrightarrow> P n) & (i < 0 \<longrightarrow> P 0))"
  | 
| 
 | 
   202  | 
  (is "?P = (?L & ?R)")
  | 
| 
 | 
   203  | 
proof (cases "i < 0")
  | 
| 
 | 
   204  | 
  case True thus ?thesis by simp
  | 
| 
 | 
   205  | 
next
  | 
| 
 | 
   206  | 
  case False
  | 
| 
 | 
   207  | 
  have "?P = ?L"
  | 
| 
 | 
   208  | 
  proof
  | 
| 
 | 
   209  | 
    assume ?P thus ?L using False by clarsimp
  | 
| 
 | 
   210  | 
  next
  | 
| 
 | 
   211  | 
    assume ?L thus ?P using False by simp
  | 
| 
 | 
   212  | 
  qed
  | 
| 
 | 
   213  | 
  with False show ?thesis by simp
  | 
| 
 | 
   214  | 
qed
  | 
| 
 | 
   215  | 
  | 
| 
 | 
   216  | 
  | 
| 
 | 
   217  | 
(*Analogous to zadd_int*)
  | 
| 
 | 
   218  | 
lemma zdiff_int: "n \<le> m ==> int m - int n = int (m-n)" 
  | 
| 
 | 
   219  | 
by (induct m n rule: diff_induct, simp_all)
  | 
| 
 | 
   220  | 
  | 
| 
 | 
   221  | 
lemma nat_mult_distrib: "(0::int) \<le> z ==> nat (z*z') = nat z * nat z'"
  | 
| 
 | 
   222  | 
apply (cases "0 \<le> z'")
  | 
| 
 | 
   223  | 
apply (rule inj_int [THEN injD])
  | 
| 
 | 
   224  | 
apply (simp add: int_mult zero_le_mult_iff)
  | 
| 
 | 
   225  | 
apply (simp add: mult_le_0_iff)
  | 
| 
 | 
   226  | 
done
  | 
| 
 | 
   227  | 
  | 
| 
 | 
   228  | 
lemma nat_mult_distrib_neg: "z \<le> (0::int) ==> nat(z*z') = nat(-z) * nat(-z')"
  | 
| 
 | 
   229  | 
apply (rule trans)
  | 
| 
 | 
   230  | 
apply (rule_tac [2] nat_mult_distrib, auto)
  | 
| 
 | 
   231  | 
done
  | 
| 
 | 
   232  | 
  | 
| 
 | 
   233  | 
lemma nat_abs_mult_distrib: "nat (abs (w * z)) = nat (abs w) * nat (abs z)"
  | 
| 
 | 
   234  | 
apply (cases "z=0 | w=0")
  | 
| 
 | 
   235  | 
apply (auto simp add: abs_if nat_mult_distrib [symmetric] 
  | 
| 
 | 
   236  | 
                      nat_mult_distrib_neg [symmetric] mult_less_0_iff)
  | 
| 
 | 
   237  | 
done
  | 
| 
 | 
   238  | 
  | 
| 
 | 
   239  | 
  | 
| 
 | 
   240  | 
subsection "Induction principles for int"
  | 
| 
 | 
   241  | 
  | 
| 
 | 
   242  | 
text{*Well-founded segments of the integers*}
 | 
| 
 | 
   243  | 
  | 
| 
 | 
   244  | 
definition
  | 
| 
 | 
   245  | 
  int_ge_less_than  ::  "int => (int * int) set"
  | 
| 
 | 
   246  | 
where
  | 
| 
 | 
   247  | 
  "int_ge_less_than d = {(z',z). d \<le> z' & z' < z}"
 | 
| 
 | 
   248  | 
  | 
| 
 | 
   249  | 
theorem wf_int_ge_less_than: "wf (int_ge_less_than d)"
  | 
| 
 | 
   250  | 
proof -
  | 
| 
 | 
   251  | 
  have "int_ge_less_than d \<subseteq> measure (%z. nat (z-d))"
  | 
| 
 | 
   252  | 
    by (auto simp add: int_ge_less_than_def)
  | 
| 
 | 
   253  | 
  thus ?thesis 
  | 
| 
 | 
   254  | 
    by (rule wf_subset [OF wf_measure]) 
  | 
| 
 | 
   255  | 
qed
  | 
| 
 | 
   256  | 
  | 
| 
 | 
   257  | 
text{*This variant looks odd, but is typical of the relations suggested
 | 
| 
 | 
   258  | 
by RankFinder.*}
  | 
| 
 | 
   259  | 
  | 
| 
 | 
   260  | 
definition
  | 
| 
 | 
   261  | 
  int_ge_less_than2 ::  "int => (int * int) set"
  | 
| 
 | 
   262  | 
where
  | 
| 
 | 
   263  | 
  "int_ge_less_than2 d = {(z',z). d \<le> z & z' < z}"
 | 
| 
 | 
   264  | 
  | 
| 
 | 
   265  | 
theorem wf_int_ge_less_than2: "wf (int_ge_less_than2 d)"
  | 
| 
 | 
   266  | 
proof -
  | 
| 
 | 
   267  | 
  have "int_ge_less_than2 d \<subseteq> measure (%z. nat (1+z-d))" 
  | 
| 
 | 
   268  | 
    by (auto simp add: int_ge_less_than2_def)
  | 
| 
 | 
   269  | 
  thus ?thesis 
  | 
| 
 | 
   270  | 
    by (rule wf_subset [OF wf_measure]) 
  | 
| 
 | 
   271  | 
qed
  | 
| 
 | 
   272  | 
  | 
| 
 | 
   273  | 
                     (* `set:int': dummy construction *)
  | 
| 
 | 
   274  | 
theorem int_ge_induct[case_names base step,induct set:int]:
  | 
| 
 | 
   275  | 
  assumes ge: "k \<le> (i::int)" and
  | 
| 
 | 
   276  | 
        base: "P(k)" and
  | 
| 
 | 
   277  | 
        step: "\<And>i. \<lbrakk>k \<le> i; P i\<rbrakk> \<Longrightarrow> P(i+1)"
  | 
| 
 | 
   278  | 
  shows "P i"
  | 
| 
 | 
   279  | 
proof -
  | 
| 
 | 
   280  | 
  { fix n have "\<And>i::int. n = nat(i-k) \<Longrightarrow> k \<le> i \<Longrightarrow> P i"
 | 
| 
 | 
   281  | 
    proof (induct n)
  | 
| 
 | 
   282  | 
      case 0
  | 
| 
 | 
   283  | 
      hence "i = k" by arith
  | 
| 
 | 
   284  | 
      thus "P i" using base by simp
  | 
| 
 | 
   285  | 
    next
  | 
| 
 | 
   286  | 
      case (Suc n)
  | 
| 
 | 
   287  | 
      hence "n = nat((i - 1) - k)" by arith
  | 
| 
 | 
   288  | 
      moreover
  | 
| 
 | 
   289  | 
      have ki1: "k \<le> i - 1" using Suc.prems by arith
  | 
| 
 | 
   290  | 
      ultimately
  | 
| 
 | 
   291  | 
      have "P(i - 1)" by(rule Suc.hyps)
  | 
| 
 | 
   292  | 
      from step[OF ki1 this] show ?case by simp
  | 
| 
 | 
   293  | 
    qed
  | 
| 
 | 
   294  | 
  }
  | 
| 
 | 
   295  | 
  with ge show ?thesis by fast
  | 
| 
 | 
   296  | 
qed
  | 
| 
 | 
   297  | 
  | 
| 
 | 
   298  | 
                     (* `set:int': dummy construction *)
  | 
| 
 | 
   299  | 
theorem int_gr_induct[case_names base step,induct set:int]:
  | 
| 
 | 
   300  | 
  assumes gr: "k < (i::int)" and
  | 
| 
 | 
   301  | 
        base: "P(k+1)" and
  | 
| 
 | 
   302  | 
        step: "\<And>i. \<lbrakk>k < i; P i\<rbrakk> \<Longrightarrow> P(i+1)"
  | 
| 
 | 
   303  | 
  shows "P i"
  | 
| 
 | 
   304  | 
apply(rule int_ge_induct[of "k + 1"])
  | 
| 
 | 
   305  | 
  using gr apply arith
  | 
| 
 | 
   306  | 
 apply(rule base)
  | 
| 
 | 
   307  | 
apply (rule step, simp+)
  | 
| 
 | 
   308  | 
done
  | 
| 
 | 
   309  | 
  | 
| 
 | 
   310  | 
theorem int_le_induct[consumes 1,case_names base step]:
  | 
| 
 | 
   311  | 
  assumes le: "i \<le> (k::int)" and
  | 
| 
 | 
   312  | 
        base: "P(k)" and
  | 
| 
 | 
   313  | 
        step: "\<And>i. \<lbrakk>i \<le> k; P i\<rbrakk> \<Longrightarrow> P(i - 1)"
  | 
| 
 | 
   314  | 
  shows "P i"
  | 
| 
 | 
   315  | 
proof -
  | 
| 
 | 
   316  | 
  { fix n have "\<And>i::int. n = nat(k-i) \<Longrightarrow> i \<le> k \<Longrightarrow> P i"
 | 
| 
 | 
   317  | 
    proof (induct n)
  | 
| 
 | 
   318  | 
      case 0
  | 
| 
 | 
   319  | 
      hence "i = k" by arith
  | 
| 
 | 
   320  | 
      thus "P i" using base by simp
  | 
| 
 | 
   321  | 
    next
  | 
| 
 | 
   322  | 
      case (Suc n)
  | 
| 
 | 
   323  | 
      hence "n = nat(k - (i+1))" by arith
  | 
| 
 | 
   324  | 
      moreover
  | 
| 
 | 
   325  | 
      have ki1: "i + 1 \<le> k" using Suc.prems by arith
  | 
| 
 | 
   326  | 
      ultimately
  | 
| 
 | 
   327  | 
      have "P(i+1)" by(rule Suc.hyps)
  | 
| 
 | 
   328  | 
      from step[OF ki1 this] show ?case by simp
  | 
| 
 | 
   329  | 
    qed
  | 
| 
 | 
   330  | 
  }
  | 
| 
 | 
   331  | 
  with le show ?thesis by fast
  | 
| 
 | 
   332  | 
qed
  | 
| 
 | 
   333  | 
  | 
| 
 | 
   334  | 
theorem int_less_induct [consumes 1,case_names base step]:
  | 
| 
 | 
   335  | 
  assumes less: "(i::int) < k" and
  | 
| 
 | 
   336  | 
        base: "P(k - 1)" and
  | 
| 
 | 
   337  | 
        step: "\<And>i. \<lbrakk>i < k; P i\<rbrakk> \<Longrightarrow> P(i - 1)"
  | 
| 
 | 
   338  | 
  shows "P i"
  | 
| 
 | 
   339  | 
apply(rule int_le_induct[of _ "k - 1"])
  | 
| 
 | 
   340  | 
  using less apply arith
  | 
| 
 | 
   341  | 
 apply(rule base)
  | 
| 
 | 
   342  | 
apply (rule step, simp+)
  | 
| 
 | 
   343  | 
done
  | 
| 
 | 
   344  | 
  | 
| 
 | 
   345  | 
subsection{*Intermediate value theorems*}
 | 
| 
 | 
   346  | 
  | 
| 
 | 
   347  | 
lemma int_val_lemma:
  | 
| 
 | 
   348  | 
     "(\<forall>i<n::nat. abs(f(i+1) - f i) \<le> 1) -->  
  | 
| 
 | 
   349  | 
      f 0 \<le> k --> k \<le> f n --> (\<exists>i \<le> n. f i = (k::int))"
  | 
| 
 | 
   350  | 
apply (induct_tac "n", simp)
  | 
| 
 | 
   351  | 
apply (intro strip)
  | 
| 
 | 
   352  | 
apply (erule impE, simp)
  | 
| 
 | 
   353  | 
apply (erule_tac x = n in allE, simp)
  | 
| 
 | 
   354  | 
apply (case_tac "k = f (n+1) ")
  | 
| 
 | 
   355  | 
 apply force
  | 
| 
 | 
   356  | 
apply (erule impE)
  | 
| 
 | 
   357  | 
 apply (simp add: abs_if split add: split_if_asm)
  | 
| 
 | 
   358  | 
apply (blast intro: le_SucI)
  | 
| 
 | 
   359  | 
done
  | 
| 
 | 
   360  | 
  | 
| 
 | 
   361  | 
lemmas nat0_intermed_int_val = int_val_lemma [rule_format (no_asm)]
  | 
| 
 | 
   362  | 
  | 
| 
 | 
   363  | 
lemma nat_intermed_int_val:
  | 
| 
 | 
   364  | 
     "[| \<forall>i. m \<le> i & i < n --> abs(f(i + 1::nat) - f i) \<le> 1; m < n;  
  | 
| 
 | 
   365  | 
         f m \<le> k; k \<le> f n |] ==> ? i. m \<le> i & i \<le> n & f i = (k::int)"
  | 
| 
 | 
   366  | 
apply (cut_tac n = "n-m" and f = "%i. f (i+m) " and k = k 
  | 
| 
 | 
   367  | 
       in int_val_lemma)
  | 
| 
 | 
   368  | 
apply simp
  | 
| 
 | 
   369  | 
apply (erule exE)
  | 
| 
 | 
   370  | 
apply (rule_tac x = "i+m" in exI, arith)
  | 
| 
 | 
   371  | 
done
  | 
| 
 | 
   372  | 
  | 
| 
 | 
   373  | 
  | 
| 
 | 
   374  | 
subsection{*Products and 1, by T. M. Rasmussen*}
 | 
| 
 | 
   375  | 
  | 
| 
 | 
   376  | 
lemma zabs_less_one_iff [simp]: "(\<bar>z\<bar> < 1) = (z = (0::int))"
  | 
| 
 | 
   377  | 
by arith
  | 
| 
 | 
   378  | 
  | 
| 
 | 
   379  | 
lemma abs_zmult_eq_1: "(\<bar>m * n\<bar> = 1) ==> \<bar>m\<bar> = (1::int)"
  | 
| 
 | 
   380  | 
apply (cases "\<bar>n\<bar>=1") 
  | 
| 
 | 
   381  | 
apply (simp add: abs_mult) 
  | 
| 
 | 
   382  | 
apply (rule ccontr) 
  | 
| 
 | 
   383  | 
apply (auto simp add: linorder_neq_iff abs_mult) 
  | 
| 
 | 
   384  | 
apply (subgoal_tac "2 \<le> \<bar>m\<bar> & 2 \<le> \<bar>n\<bar>")
  | 
| 
 | 
   385  | 
 prefer 2 apply arith 
  | 
| 
 | 
   386  | 
apply (subgoal_tac "2*2 \<le> \<bar>m\<bar> * \<bar>n\<bar>", simp) 
  | 
| 
 | 
   387  | 
apply (rule mult_mono, auto) 
  | 
| 
 | 
   388  | 
done
  | 
| 
 | 
   389  | 
  | 
| 
 | 
   390  | 
lemma pos_zmult_eq_1_iff_lemma: "(m * n = 1) ==> m = (1::int) | m = -1"
  | 
| 
 | 
   391  | 
by (insert abs_zmult_eq_1 [of m n], arith)
  | 
| 
 | 
   392  | 
  | 
| 
 | 
   393  | 
lemma pos_zmult_eq_1_iff: "0 < (m::int) ==> (m * n = 1) = (m = 1 & n = 1)"
  | 
| 
 | 
   394  | 
apply (auto dest: pos_zmult_eq_1_iff_lemma) 
  | 
| 
 | 
   395  | 
apply (simp add: mult_commute [of m]) 
  | 
| 
 | 
   396  | 
apply (frule pos_zmult_eq_1_iff_lemma, auto) 
  | 
| 
 | 
   397  | 
done
  | 
| 
 | 
   398  | 
  | 
| 
 | 
   399  | 
lemma zmult_eq_1_iff: "(m*n = (1::int)) = ((m = 1 & n = 1) | (m = -1 & n = -1))"
  | 
| 
 | 
   400  | 
apply (rule iffI) 
  | 
| 
 | 
   401  | 
 apply (frule pos_zmult_eq_1_iff_lemma)
  | 
| 
 | 
   402  | 
 apply (simp add: mult_commute [of m]) 
  | 
| 
 | 
   403  | 
 apply (frule pos_zmult_eq_1_iff_lemma, auto) 
  | 
| 
 | 
   404  | 
done
  | 
| 
 | 
   405  | 
  | 
| 
 | 
   406  | 
  | 
| 
 | 
   407  | 
subsection {* Legacy ML bindings *}
 | 
| 
 | 
   408  | 
  | 
| 
 | 
   409  | 
ML {*
 | 
| 
 | 
   410  | 
val of_int_number_of_eq = @{thm of_int_number_of_eq};
 | 
| 
 | 
   411  | 
val nat_0 = @{thm nat_0};
 | 
| 
 | 
   412  | 
val nat_1 = @{thm nat_1};
 | 
| 
 | 
   413  | 
*}
  | 
| 
 | 
   414  | 
  | 
| 
 | 
   415  | 
end
  |