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