src/HOL/Word/Num_Lemmas.thy
author haftmann
Tue, 15 Jan 2008 16:19:23 +0100
changeset 25919 8b1c0d434824
parent 25875 536dfdc25e0a
child 25937 6e5b0f176dac
permissions -rw-r--r--
joined theories IntDef, Numeral, IntArith to theory Int
     1 (* 
     2   ID:      $Id$
     3   Author:  Jeremy Dawson, NICTA
     4 *) 
     5 
     6 header {* Useful Numerical Lemmas *}
     7 
     8 theory Num_Lemmas
     9 imports Main Parity
    10 begin
    11 
    12 lemma contentsI: "y = {x} ==> contents y = x" 
    13   unfolding contents_def by auto
    14 
    15 lemma prod_case_split: "prod_case = split"
    16   by (rule ext)+ auto
    17  
    18 lemmas split_split = prod.split [unfolded prod_case_split] 
    19 lemmas split_split_asm = prod.split_asm [unfolded prod_case_split]
    20 lemmas "split.splits" = split_split split_split_asm 
    21 
    22 lemmas funpow_0 = funpow.simps(1)
    23 lemmas funpow_Suc = funpow.simps(2)
    24 
    25 lemma nonemptyE: "S ~= {} ==> (!!x. x : S ==> R) ==> R"
    26   apply (erule contrapos_np)
    27   apply (rule equals0I)
    28   apply auto
    29   done
    30 
    31 lemma gt_or_eq_0: "0 < y \<or> 0 = (y::nat)" by auto
    32 
    33 constdefs
    34   mod_alt :: "'a => 'a => 'a :: Divides.div"
    35   "mod_alt n m == n mod m" 
    36 
    37   -- "alternative way of defining @{text bin_last}, @{text bin_rest}"
    38   bin_rl :: "int => int * bit" 
    39   "bin_rl w == SOME (r, l). w = r BIT l"
    40 
    41 declare iszero_0 [iff]
    42 
    43 lemmas xtr1 = xtrans(1)
    44 lemmas xtr2 = xtrans(2)
    45 lemmas xtr3 = xtrans(3)
    46 lemmas xtr4 = xtrans(4)
    47 lemmas xtr5 = xtrans(5)
    48 lemmas xtr6 = xtrans(6)
    49 lemmas xtr7 = xtrans(7)
    50 lemmas xtr8 = xtrans(8)
    51 
    52 lemma Min_ne_Pls [iff]:  
    53   "Int.Min ~= Int.Pls"
    54   unfolding Min_def Pls_def by auto
    55 
    56 lemmas Pls_ne_Min [iff] = Min_ne_Pls [symmetric]
    57 
    58 lemmas PlsMin_defs [intro!] = 
    59   Pls_def Min_def Pls_def [symmetric] Min_def [symmetric]
    60 
    61 lemmas PlsMin_simps [simp] = PlsMin_defs [THEN Eq_TrueI]
    62 
    63 lemma number_of_False_cong: 
    64   "False ==> number_of x = number_of y" 
    65   by auto
    66 
    67 lemmas nat_simps = diff_add_inverse2 diff_add_inverse
    68 lemmas nat_iffs = le_add1 le_add2
    69 
    70 lemma sum_imp_diff: "j = k + i ==> j - i = (k :: nat)"
    71   by (clarsimp simp add: nat_simps)
    72 
    73 lemma nobm1:
    74   "0 < (number_of w :: nat) ==> 
    75    number_of w - (1 :: nat) = number_of (Int.pred w)"
    76   apply (unfold nat_number_of_def One_nat_def nat_1 [symmetric] pred_def)
    77   apply (simp add: number_of_eq nat_diff_distrib [symmetric])
    78   done
    79 
    80 lemma of_int_power:
    81   "of_int (a ^ n) = (of_int a ^ n :: 'a :: {recpower, comm_ring_1})" 
    82   by (induct n) (auto simp add: power_Suc)
    83 
    84 lemma zless2: "0 < (2 :: int)" 
    85   by auto
    86 
    87 lemmas zless2p [simp] = zless2 [THEN zero_less_power]
    88 lemmas zle2p [simp] = zless2p [THEN order_less_imp_le]
    89 
    90 lemmas pos_mod_sign2 = zless2 [THEN pos_mod_sign [where b = "2::int"]]
    91 lemmas pos_mod_bound2 = zless2 [THEN pos_mod_bound [where b = "2::int"]]
    92 
    93 -- "the inverse(s) of @{text number_of}"
    94 lemma nmod2: "n mod (2::int) = 0 | n mod 2 = 1"
    95   using pos_mod_sign2 [of n] pos_mod_bound2 [of n]
    96   unfolding mod_alt_def [symmetric] by auto
    97 
    98 lemma emep1:
    99   "even n ==> even d ==> 0 <= d ==> (n + 1) mod (d :: int) = (n mod d) + 1"
   100   apply (simp add: add_commute)
   101   apply (safe dest!: even_equiv_def [THEN iffD1])
   102   apply (subst pos_zmod_mult_2)
   103    apply arith
   104   apply (simp add: zmod_zmult_zmult1)
   105  done
   106 
   107 lemmas eme1p = emep1 [simplified add_commute]
   108 
   109 lemma le_diff_eq': "(a \<le> c - b) = (b + a \<le> (c::int))"
   110   by (simp add: le_diff_eq add_commute)
   111 
   112 lemma less_diff_eq': "(a < c - b) = (b + a < (c::int))"
   113   by (simp add: less_diff_eq add_commute)
   114 
   115 lemma diff_le_eq': "(a - b \<le> c) = (a \<le> b + (c::int))"
   116   by (simp add: diff_le_eq add_commute)
   117 
   118 lemma diff_less_eq': "(a - b < c) = (a < b + (c::int))"
   119   by (simp add: diff_less_eq add_commute)
   120 
   121 
   122 lemmas m1mod2k = zless2p [THEN zmod_minus1]
   123 lemmas m1mod22k = mult_pos_pos [OF zless2 zless2p, THEN zmod_minus1]
   124 lemmas p1mod22k' = zless2p [THEN order_less_imp_le, THEN pos_zmod_mult_2]
   125 lemmas z1pmod2' = zero_le_one [THEN pos_zmod_mult_2, simplified]
   126 lemmas z1pdiv2' = zero_le_one [THEN pos_zdiv_mult_2, simplified]
   127 
   128 lemma p1mod22k:
   129   "(2 * b + 1) mod (2 * 2 ^ n) = 2 * (b mod 2 ^ n) + (1::int)"
   130   by (simp add: p1mod22k' add_commute)
   131 
   132 lemma z1pmod2:
   133   "(2 * b + 1) mod 2 = (1::int)"
   134   by (simp add: z1pmod2' add_commute)
   135   
   136 lemma z1pdiv2:
   137   "(2 * b + 1) div 2 = (b::int)"
   138   by (simp add: z1pdiv2' add_commute)
   139 
   140 lemmas zdiv_le_dividend = xtr3 [OF zdiv_1 [symmetric] zdiv_mono2,
   141   simplified int_one_le_iff_zero_less, simplified, standard]
   142   
   143 (** ways in which type Bin resembles a datatype **)
   144 
   145 lemma BIT_eq: "u BIT b = v BIT c ==> u = v & b = c"
   146   apply (unfold Bit_def)
   147   apply (simp (no_asm_use) split: bit.split_asm)
   148      apply simp_all
   149    apply (drule_tac f=even in arg_cong, clarsimp)+
   150   done
   151      
   152 lemmas BIT_eqE [elim!] = BIT_eq [THEN conjE, standard]
   153 
   154 lemma BIT_eq_iff [simp]: 
   155   "(u BIT b = v BIT c) = (u = v \<and> b = c)"
   156   by (rule iffI) auto
   157 
   158 lemmas BIT_eqI [intro!] = conjI [THEN BIT_eq_iff [THEN iffD2]]
   159 
   160 lemma less_Bits: 
   161   "(v BIT b < w BIT c) = (v < w | v <= w & b = bit.B0 & c = bit.B1)"
   162   unfolding Bit_def by (auto split: bit.split)
   163 
   164 lemma le_Bits: 
   165   "(v BIT b <= w BIT c) = (v < w | v <= w & (b ~= bit.B1 | c ~= bit.B0))" 
   166   unfolding Bit_def by (auto split: bit.split)
   167 
   168 lemma neB1E [elim!]:
   169   assumes ne: "y \<noteq> bit.B1"
   170   assumes y: "y = bit.B0 \<Longrightarrow> P"
   171   shows "P"
   172   apply (rule y)
   173   apply (cases y rule: bit.exhaust, simp)
   174   apply (simp add: ne)
   175   done
   176 
   177 lemma bin_ex_rl: "EX w b. w BIT b = bin"
   178   apply (unfold Bit_def)
   179   apply (cases "even bin")
   180    apply (clarsimp simp: even_equiv_def)
   181    apply (auto simp: odd_equiv_def split: bit.split)
   182   done
   183 
   184 lemma bin_exhaust:
   185   assumes Q: "\<And>x b. bin = x BIT b \<Longrightarrow> Q"
   186   shows "Q"
   187   apply (insert bin_ex_rl [of bin])  
   188   apply (erule exE)+
   189   apply (rule Q)
   190   apply force
   191   done
   192 
   193 lemma bin_rl_char: "(bin_rl w = (r, l)) = (r BIT l = w)"
   194   apply (unfold bin_rl_def)
   195   apply safe
   196    apply (cases w rule: bin_exhaust)
   197    apply auto
   198   done
   199 
   200 lemmas bin_rl_simps [THEN bin_rl_char [THEN iffD2], standard, simp] =
   201   Pls_0_eq Min_1_eq refl 
   202 
   203 lemma bin_abs_lem:
   204   "bin = (w BIT b) ==> ~ bin = Int.Min --> ~ bin = Int.Pls -->
   205     nat (abs w) < nat (abs bin)"
   206   apply (clarsimp simp add: bin_rl_char)
   207   apply (unfold Pls_def Min_def Bit_def)
   208   apply (cases b)
   209    apply (clarsimp, arith)
   210   apply (clarsimp, arith)
   211   done
   212 
   213 lemma bin_induct:
   214   assumes PPls: "P Int.Pls"
   215     and PMin: "P Int.Min"
   216     and PBit: "!!bin bit. P bin ==> P (bin BIT bit)"
   217   shows "P bin"
   218   apply (rule_tac P=P and a=bin and f1="nat o abs" 
   219                   in wf_measure [THEN wf_induct])
   220   apply (simp add: measure_def inv_image_def)
   221   apply (case_tac x rule: bin_exhaust)
   222   apply (frule bin_abs_lem)
   223   apply (auto simp add : PPls PMin PBit)
   224   done
   225 
   226 lemma no_no [simp]: "number_of (number_of i) = i"
   227   unfolding number_of_eq by simp
   228 
   229 lemma Bit_B0:
   230   "k BIT bit.B0 = k + k"
   231    by (unfold Bit_def) simp
   232 
   233 lemma Bit_B1:
   234   "k BIT bit.B1 = k + k + 1"
   235    by (unfold Bit_def) simp
   236   
   237 lemma Bit_B0_2t: "k BIT bit.B0 = 2 * k"
   238   by (rule trans, rule Bit_B0) simp
   239 
   240 lemma Bit_B1_2t: "k BIT bit.B1 = 2 * k + 1"
   241   by (rule trans, rule Bit_B1) simp
   242 
   243 lemma B_mod_2': 
   244   "X = 2 ==> (w BIT bit.B1) mod X = 1 & (w BIT bit.B0) mod X = 0"
   245   apply (simp (no_asm) only: Bit_B0 Bit_B1)
   246   apply (simp add: z1pmod2)
   247   done
   248     
   249 lemmas B1_mod_2 [simp] = B_mod_2' [OF refl, THEN conjunct1, standard]
   250 lemmas B0_mod_2 [simp] = B_mod_2' [OF refl, THEN conjunct2, standard]
   251 
   252 lemma axxbyy:
   253   "a + m + m = b + n + n ==> (a = 0 | a = 1) ==> (b = 0 | b = 1) ==>  
   254    a = b & m = (n :: int)"
   255   apply auto
   256    apply (drule_tac f="%n. n mod 2" in arg_cong)
   257    apply (clarsimp simp: z1pmod2)
   258   apply (drule_tac f="%n. n mod 2" in arg_cong)
   259   apply (clarsimp simp: z1pmod2)
   260   done
   261 
   262 lemma axxmod2:
   263   "(1 + x + x) mod 2 = (1 :: int) & (0 + x + x) mod 2 = (0 :: int)" 
   264   by simp (rule z1pmod2)
   265 
   266 lemma axxdiv2:
   267   "(1 + x + x) div 2 = (x :: int) & (0 + x + x) div 2 = (x :: int)" 
   268   by simp (rule z1pdiv2)
   269 
   270 lemmas iszero_minus = trans [THEN trans,
   271   OF iszero_def neg_equal_0_iff_equal iszero_def [symmetric], standard]
   272 
   273 lemmas zadd_diff_inverse = trans [OF diff_add_cancel [symmetric] add_commute,
   274   standard]
   275 
   276 lemmas add_diff_cancel2 = add_commute [THEN diff_eq_eq [THEN iffD2], standard]
   277 
   278 lemma zmod_uminus: "- ((a :: int) mod b) mod b = -a mod b"
   279   by (simp add : zmod_zminus1_eq_if)
   280 
   281 lemma zmod_zsub_distrib: "((a::int) - b) mod c = (a mod c - b mod c) mod c"
   282   apply (unfold diff_int_def)
   283   apply (rule trans [OF _ zmod_zadd1_eq [symmetric]])
   284   apply (simp add: zmod_uminus zmod_zadd1_eq [symmetric])
   285   done
   286 
   287 lemma zmod_zsub_right_eq: "((a::int) - b) mod c = (a - b mod c) mod c"
   288   apply (unfold diff_int_def)
   289   apply (rule trans [OF _ zmod_zadd_right_eq [symmetric]])
   290   apply (simp add : zmod_uminus zmod_zadd_right_eq [symmetric])
   291   done
   292 
   293 lemma zmod_zsub_left_eq: "((a::int) - b) mod c = (a mod c - b) mod c"
   294   by (rule zmod_zadd_left_eq [where b = "- b", simplified diff_int_def [symmetric]])
   295 
   296 lemma zmod_zsub_self [simp]: 
   297   "((b :: int) - a) mod a = b mod a"
   298   by (simp add: zmod_zsub_right_eq)
   299 
   300 lemma zmod_zmult1_eq_rev:
   301   "b * a mod c = b mod c * a mod (c::int)"
   302   apply (simp add: mult_commute)
   303   apply (subst zmod_zmult1_eq)
   304   apply simp
   305   done
   306 
   307 lemmas rdmods [symmetric] = zmod_uminus [symmetric]
   308   zmod_zsub_left_eq zmod_zsub_right_eq zmod_zadd_left_eq
   309   zmod_zadd_right_eq zmod_zmult1_eq zmod_zmult1_eq_rev
   310 
   311 lemma mod_plus_right:
   312   "((a + x) mod m = (b + x) mod m) = (a mod m = b mod (m :: nat))"
   313   apply (induct x)
   314    apply (simp_all add: mod_Suc)
   315   apply arith
   316   done
   317 
   318 lemma nat_minus_mod: "(n - n mod m) mod m = (0 :: nat)"
   319   by (induct n) (simp_all add : mod_Suc)
   320 
   321 lemmas nat_minus_mod_plus_right = trans [OF nat_minus_mod mod_0 [symmetric],
   322   THEN mod_plus_right [THEN iffD2], standard, simplified]
   323 
   324 lemmas push_mods' = zmod_zadd1_eq [standard]
   325   zmod_zmult_distrib [standard] zmod_zsub_distrib [standard]
   326   zmod_uminus [symmetric, standard]
   327 
   328 lemmas push_mods = push_mods' [THEN eq_reflection, standard]
   329 lemmas pull_mods = push_mods [symmetric] rdmods [THEN eq_reflection, standard]
   330 lemmas mod_simps = 
   331   zmod_zmult_self1 [THEN eq_reflection] zmod_zmult_self2 [THEN eq_reflection]
   332   mod_mod_trivial [THEN eq_reflection]
   333 
   334 lemma nat_mod_eq:
   335   "!!b. b < n ==> a mod n = b mod n ==> a mod n = (b :: nat)" 
   336   by (induct a) auto
   337 
   338 lemmas nat_mod_eq' = refl [THEN [2] nat_mod_eq]
   339 
   340 lemma nat_mod_lem: 
   341   "(0 :: nat) < n ==> b < n = (b mod n = b)"
   342   apply safe
   343    apply (erule nat_mod_eq')
   344   apply (erule subst)
   345   apply (erule mod_less_divisor)
   346   done
   347 
   348 lemma mod_nat_add: 
   349   "(x :: nat) < z ==> y < z ==> 
   350    (x + y) mod z = (if x + y < z then x + y else x + y - z)"
   351   apply (rule nat_mod_eq)
   352    apply auto
   353   apply (rule trans)
   354    apply (rule le_mod_geq)
   355    apply simp
   356   apply (rule nat_mod_eq')
   357   apply arith
   358   done
   359 
   360 lemma mod_nat_sub: 
   361   "(x :: nat) < z ==> (x - y) mod z = x - y"
   362   by (rule nat_mod_eq') arith
   363 
   364 lemma int_mod_lem: 
   365   "(0 :: int) < n ==> (0 <= b & b < n) = (b mod n = b)"
   366   apply safe
   367     apply (erule (1) mod_pos_pos_trivial)
   368    apply (erule_tac [!] subst)
   369    apply auto
   370   done
   371 
   372 lemma int_mod_eq:
   373   "(0 :: int) <= b ==> b < n ==> a mod n = b mod n ==> a mod n = b"
   374   by clarsimp (rule mod_pos_pos_trivial)
   375 
   376 lemmas int_mod_eq' = refl [THEN [3] int_mod_eq]
   377 
   378 lemma int_mod_le: "0 <= a ==> 0 < (n :: int) ==> a mod n <= a"
   379   apply (cases "a < n")
   380    apply (auto dest: mod_pos_pos_trivial pos_mod_bound [where a=a])
   381   done
   382 
   383 lemma int_mod_le': "0 <= b - n ==> 0 < (n :: int) ==> b mod n <= b - n"
   384   by (rule int_mod_le [where a = "b - n" and n = n, simplified])
   385 
   386 lemma int_mod_ge: "a < n ==> 0 < (n :: int) ==> a <= a mod n"
   387   apply (cases "0 <= a")
   388    apply (drule (1) mod_pos_pos_trivial)
   389    apply simp
   390   apply (rule order_trans [OF _ pos_mod_sign])
   391    apply simp
   392   apply assumption
   393   done
   394 
   395 lemma int_mod_ge': "b < 0 ==> 0 < (n :: int) ==> b + n <= b mod n"
   396   by (rule int_mod_ge [where a = "b + n" and n = n, simplified])
   397 
   398 lemma mod_add_if_z:
   399   "(x :: int) < z ==> y < z ==> 0 <= y ==> 0 <= x ==> 0 <= z ==> 
   400    (x + y) mod z = (if x + y < z then x + y else x + y - z)"
   401   by (auto intro: int_mod_eq)
   402 
   403 lemma mod_sub_if_z:
   404   "(x :: int) < z ==> y < z ==> 0 <= y ==> 0 <= x ==> 0 <= z ==> 
   405    (x - y) mod z = (if y <= x then x - y else x - y + z)"
   406   by (auto intro: int_mod_eq)
   407 
   408 lemmas zmde = zmod_zdiv_equality [THEN diff_eq_eq [THEN iffD2], symmetric]
   409 lemmas mcl = mult_cancel_left [THEN iffD1, THEN make_pos_rule]
   410 
   411 (* already have this for naturals, div_mult_self1/2, but not for ints *)
   412 lemma zdiv_mult_self: "m ~= (0 :: int) ==> (a + m * n) div m = a div m + n"
   413   apply (rule mcl)
   414    prefer 2
   415    apply (erule asm_rl)
   416   apply (simp add: zmde ring_distribs)
   417   apply (simp add: push_mods)
   418   done
   419 
   420 (** Rep_Integ **)
   421 lemma eqne: "equiv A r ==> X : A // r ==> X ~= {}"
   422   unfolding equiv_def refl_def quotient_def Image_def by auto
   423 
   424 lemmas Rep_Integ_ne = Integ.Rep_Integ 
   425   [THEN equiv_intrel [THEN eqne, simplified Integ_def [symmetric]], standard]
   426 
   427 lemmas riq = Integ.Rep_Integ [simplified Integ_def]
   428 lemmas intrel_refl = refl [THEN equiv_intrel_iff [THEN iffD1], standard]
   429 lemmas Rep_Integ_equiv = quotient_eq_iff
   430   [OF equiv_intrel riq riq, simplified Integ.Rep_Integ_inject, standard]
   431 lemmas Rep_Integ_same = 
   432   Rep_Integ_equiv [THEN intrel_refl [THEN rev_iffD2], standard]
   433 
   434 lemma RI_int: "(a, 0) : Rep_Integ (int a)"
   435   unfolding int_def by auto
   436 
   437 lemmas RI_intrel [simp] = UNIV_I [THEN quotientI,
   438   THEN Integ.Abs_Integ_inverse [simplified Integ_def], standard]
   439 
   440 lemma RI_minus: "(a, b) : Rep_Integ x ==> (b, a) : Rep_Integ (- x)"
   441   apply (rule_tac z=x in eq_Abs_Integ)
   442   apply (clarsimp simp: minus)
   443   done
   444 
   445 lemma RI_add: 
   446   "(a, b) : Rep_Integ x ==> (c, d) : Rep_Integ y ==> 
   447    (a + c, b + d) : Rep_Integ (x + y)"
   448   apply (rule_tac z=x in eq_Abs_Integ)
   449   apply (rule_tac z=y in eq_Abs_Integ) 
   450   apply (clarsimp simp: add)
   451   done
   452 
   453 lemma mem_same: "a : S ==> a = b ==> b : S"
   454   by fast
   455 
   456 (* two alternative proofs of this *)
   457 lemma RI_eq_diff': "(a, b) : Rep_Integ (int a - int b)"
   458   apply (unfold diff_def)
   459   apply (rule mem_same)
   460    apply (rule RI_minus RI_add RI_int)+
   461   apply simp
   462   done
   463 
   464 lemma RI_eq_diff: "((a, b) : Rep_Integ x) = (int a - int b = x)"
   465   apply safe
   466    apply (rule Rep_Integ_same)
   467     prefer 2
   468     apply (erule asm_rl)
   469    apply (rule RI_eq_diff')+
   470   done
   471 
   472 lemma mod_power_lem:
   473   "a > 1 ==> a ^ n mod a ^ m = (if m <= n then 0 else (a :: int) ^ n)"
   474   apply clarsimp
   475   apply safe
   476    apply (simp add: zdvd_iff_zmod_eq_0 [symmetric])
   477    apply (drule le_iff_add [THEN iffD1])
   478    apply (force simp: zpower_zadd_distrib)
   479   apply (rule mod_pos_pos_trivial)
   480    apply (simp)
   481   apply (rule power_strict_increasing)
   482    apply auto
   483   done
   484 
   485 lemma min_pm [simp]: "min a b + (a - b) = (a :: nat)"
   486   by arith
   487   
   488 lemmas min_pm1 [simp] = trans [OF add_commute min_pm]
   489 
   490 lemma rev_min_pm [simp]: "min b a + (a - b) = (a::nat)"
   491   by simp
   492 
   493 lemmas rev_min_pm1 [simp] = trans [OF add_commute rev_min_pm]
   494 
   495 lemma pl_pl_rels: 
   496   "a + b = c + d ==> 
   497    a >= c & b <= d | a <= c & b >= (d :: nat)"
   498   apply (cut_tac n=a and m=c in nat_le_linear)
   499   apply (safe dest!: le_iff_add [THEN iffD1])
   500          apply arith+
   501   done
   502 
   503 lemmas pl_pl_rels' = add_commute [THEN [2] trans, THEN pl_pl_rels]
   504 
   505 lemma minus_eq: "(m - k = m) = (k = 0 | m = (0 :: nat))"
   506   by arith
   507 
   508 lemma pl_pl_mm: "(a :: nat) + b = c + d ==> a - c = d - b"
   509   by arith
   510 
   511 lemmas pl_pl_mm' = add_commute [THEN [2] trans, THEN pl_pl_mm]
   512  
   513 lemma min_minus [simp] : "min m (m - k) = (m - k :: nat)"
   514   by arith
   515   
   516 lemmas min_minus' [simp] = trans [OF min_max.inf_commute min_minus]
   517 
   518 lemma nat_no_eq_iff: 
   519   "(number_of b :: int) >= 0 ==> (number_of c :: int) >= 0 ==> 
   520    (number_of b = (number_of c :: nat)) = (b = c)"
   521   apply (unfold nat_number_of_def)
   522   apply safe
   523   apply (drule (2) eq_nat_nat_iff [THEN iffD1])
   524   apply (simp add: number_of_eq)
   525   done
   526 
   527 lemmas dme = box_equals [OF div_mod_equality add_0_right add_0_right]
   528 lemmas dtle = xtr3 [OF dme [symmetric] le_add1]
   529 lemmas th2 = order_trans [OF order_refl [THEN [2] mult_le_mono] dtle]
   530 
   531 lemma td_gal: 
   532   "0 < c ==> (a >= b * c) = (a div c >= (b :: nat))"
   533   apply safe
   534    apply (erule (1) xtr4 [OF div_le_mono div_mult_self_is_m])
   535   apply (erule th2)
   536   done
   537   
   538 lemmas td_gal_lt = td_gal [simplified le_def, simplified]
   539 
   540 lemma div_mult_le: "(a :: nat) div b * b <= a"
   541   apply (cases b)
   542    prefer 2
   543    apply (rule order_refl [THEN th2])
   544   apply auto
   545   done
   546 
   547 lemmas sdl = split_div_lemma [THEN iffD1, symmetric]
   548 
   549 lemma given_quot: "f > (0 :: nat) ==> (f * l + (f - 1)) div f = l"
   550   by (rule sdl, assumption) (simp (no_asm))
   551 
   552 lemma given_quot_alt: "f > (0 :: nat) ==> (l * f + f - Suc 0) div f = l"
   553   apply (frule given_quot)
   554   apply (rule trans)
   555    prefer 2
   556    apply (erule asm_rl)
   557   apply (rule_tac f="%n. n div f" in arg_cong)
   558   apply (simp add : mult_ac)
   559   done
   560     
   561 lemma diff_mod_le: "(a::nat) < d ==> b dvd d ==> a - a mod b <= d - b"
   562   apply (unfold dvd_def)
   563   apply clarify
   564   apply (case_tac k)
   565    apply clarsimp
   566   apply clarify
   567   apply (cases "b > 0")
   568    apply (drule mult_commute [THEN xtr1])
   569    apply (frule (1) td_gal_lt [THEN iffD1])
   570    apply (clarsimp simp: le_simps)
   571    apply (rule mult_div_cancel [THEN [2] xtr4])
   572    apply (rule mult_mono)
   573       apply auto
   574   done
   575 
   576 lemma less_le_mult':
   577   "w * c < b * c ==> 0 \<le> c ==> (w + 1) * c \<le> b * (c::int)"
   578   apply (rule mult_right_mono)
   579    apply (rule zless_imp_add1_zle)
   580    apply (erule (1) mult_right_less_imp_less)
   581   apply assumption
   582   done
   583 
   584 lemmas less_le_mult = less_le_mult' [simplified left_distrib, simplified]
   585 
   586 lemmas less_le_mult_minus = iffD2 [OF le_diff_eq less_le_mult, 
   587   simplified left_diff_distrib, standard]
   588 
   589 lemma lrlem':
   590   assumes d: "(i::nat) \<le> j \<or> m < j'"
   591   assumes R1: "i * k \<le> j * k \<Longrightarrow> R"
   592   assumes R2: "Suc m * k' \<le> j' * k' \<Longrightarrow> R"
   593   shows "R" using d
   594   apply safe
   595    apply (rule R1, erule mult_le_mono1)
   596   apply (rule R2, erule Suc_le_eq [THEN iffD2 [THEN mult_le_mono1]])
   597   done
   598 
   599 lemma lrlem: "(0::nat) < sc ==>
   600     (sc - n + (n + lb * n) <= m * n) = (sc + lb * n <= m * n)"
   601   apply safe
   602    apply arith
   603   apply (case_tac "sc >= n")
   604    apply arith
   605   apply (insert linorder_le_less_linear [of m lb])
   606   apply (erule_tac k=n and k'=n in lrlem')
   607    apply arith
   608   apply simp
   609   done
   610 
   611 lemma gen_minus: "0 < n ==> f n = f (Suc (n - 1))"
   612   by auto
   613 
   614 lemma mpl_lem: "j <= (i :: nat) ==> k < j ==> i - j + k < i"
   615   apply (induct i, clarsimp)
   616   apply (cases j, clarsimp)
   617   apply arith
   618   done
   619 
   620 lemma nonneg_mod_div:
   621   "0 <= a ==> 0 <= b ==> 0 <= (a mod b :: int) & 0 <= a div b"
   622   apply (cases "b = 0", clarsimp)
   623   apply (auto intro: pos_imp_zdiv_nonneg_iff [THEN iffD2])
   624   done
   625 
   626 end