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