src/HOL/Divides.thy
author huffman
Mon, 20 Feb 2012 14:23:46 +0100
changeset 46551 866bce5442a3
parent 46026 83caa4f4bd56
child 46552 5d33a3269029
permissions -rw-r--r--
simplify projections on simultaneous computations of div and mod; tuned structure (from Florian Haftmann)
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
3366
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
     1
(*  Title:      HOL/Divides.thy
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
     2
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
6865
5577ffe4c2f1 now div and mod are overloaded; dvd is polymorphic
paulson
parents: 3366
diff changeset
     3
    Copyright   1999  University of Cambridge
18154
0c05abaf6244 add header
huffman
parents: 17609
diff changeset
     4
*)
3366
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
     5
27651
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27540
diff changeset
     6
header {* The division operators div and mod *}
3366
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
     7
15131
c69542757a4d New theory header syntax.
nipkow
parents: 14640
diff changeset
     8
theory Divides
33318
ddd97d9dfbfb moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 33296
diff changeset
     9
imports Nat_Numeral Nat_Transfer
33340
a165b97f3658 moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 33318
diff changeset
    10
uses "~~/src/Provers/Arith/cancel_div_mod.ML"
15131
c69542757a4d New theory header syntax.
nipkow
parents: 14640
diff changeset
    11
begin
3366
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
    12
25942
a52309ac4a4d added class semiring_div
haftmann
parents: 25571
diff changeset
    13
subsection {* Syntactic division operations *}
a52309ac4a4d added class semiring_div
haftmann
parents: 25571
diff changeset
    14
27651
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27540
diff changeset
    15
class div = dvd +
27540
dc38e79f5a1c separate class dvd for divisibility predicate
haftmann
parents: 26748
diff changeset
    16
  fixes div :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixl "div" 70)
27651
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27540
diff changeset
    17
    and mod :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixl "mod" 70)
27540
dc38e79f5a1c separate class dvd for divisibility predicate
haftmann
parents: 26748
diff changeset
    18
dc38e79f5a1c separate class dvd for divisibility predicate
haftmann
parents: 26748
diff changeset
    19
27651
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27540
diff changeset
    20
subsection {* Abstract division in commutative semirings. *}
25942
a52309ac4a4d added class semiring_div
haftmann
parents: 25571
diff changeset
    21
30930
11010e5f18f0 tightended specification of class semiring_div
haftmann
parents: 30923
diff changeset
    22
class semiring_div = comm_semiring_1_cancel + no_zero_divisors + div +
25942
a52309ac4a4d added class semiring_div
haftmann
parents: 25571
diff changeset
    23
  assumes mod_div_equality: "a div b * b + a mod b = a"
27651
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27540
diff changeset
    24
    and div_by_0 [simp]: "a div 0 = 0"
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27540
diff changeset
    25
    and div_0 [simp]: "0 div a = 0"
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27540
diff changeset
    26
    and div_mult_self1 [simp]: "b \<noteq> 0 \<Longrightarrow> (a + c * b) div b = c + a div b"
30930
11010e5f18f0 tightended specification of class semiring_div
haftmann
parents: 30923
diff changeset
    27
    and div_mult_mult1 [simp]: "c \<noteq> 0 \<Longrightarrow> (c * a) div (c * b) = a div b"
25942
a52309ac4a4d added class semiring_div
haftmann
parents: 25571
diff changeset
    28
begin
a52309ac4a4d added class semiring_div
haftmann
parents: 25571
diff changeset
    29
26100
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
    30
text {* @{const div} and @{const mod} *}
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
    31
26062
16f334d7156a more abstract lemmas
haftmann
parents: 25947
diff changeset
    32
lemma mod_div_equality2: "b * (a div b) + a mod b = a"
16f334d7156a more abstract lemmas
haftmann
parents: 25947
diff changeset
    33
  unfolding mult_commute [of b]
16f334d7156a more abstract lemmas
haftmann
parents: 25947
diff changeset
    34
  by (rule mod_div_equality)
16f334d7156a more abstract lemmas
haftmann
parents: 25947
diff changeset
    35
29403
fe17df4e4ab3 generalize some div/mod lemmas; remove type-specific proofs
huffman
parents: 29252
diff changeset
    36
lemma mod_div_equality': "a mod b + a div b * b = a"
fe17df4e4ab3 generalize some div/mod lemmas; remove type-specific proofs
huffman
parents: 29252
diff changeset
    37
  using mod_div_equality [of a b]
fe17df4e4ab3 generalize some div/mod lemmas; remove type-specific proofs
huffman
parents: 29252
diff changeset
    38
  by (simp only: add_ac)
fe17df4e4ab3 generalize some div/mod lemmas; remove type-specific proofs
huffman
parents: 29252
diff changeset
    39
26062
16f334d7156a more abstract lemmas
haftmann
parents: 25947
diff changeset
    40
lemma div_mod_equality: "((a div b) * b + a mod b) + c = a + c"
30934
ed5377c2b0a3 tuned setups of CancelDivMod
haftmann
parents: 30930
diff changeset
    41
  by (simp add: mod_div_equality)
26062
16f334d7156a more abstract lemmas
haftmann
parents: 25947
diff changeset
    42
16f334d7156a more abstract lemmas
haftmann
parents: 25947
diff changeset
    43
lemma div_mod_equality2: "(b * (a div b) + a mod b) + c = a + c"
30934
ed5377c2b0a3 tuned setups of CancelDivMod
haftmann
parents: 30930
diff changeset
    44
  by (simp add: mod_div_equality2)
26062
16f334d7156a more abstract lemmas
haftmann
parents: 25947
diff changeset
    45
27651
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27540
diff changeset
    46
lemma mod_by_0 [simp]: "a mod 0 = a"
30934
ed5377c2b0a3 tuned setups of CancelDivMod
haftmann
parents: 30930
diff changeset
    47
  using mod_div_equality [of a zero] by simp
27651
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27540
diff changeset
    48
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27540
diff changeset
    49
lemma mod_0 [simp]: "0 mod a = 0"
30934
ed5377c2b0a3 tuned setups of CancelDivMod
haftmann
parents: 30930
diff changeset
    50
  using mod_div_equality [of zero a] div_0 by simp
27651
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27540
diff changeset
    51
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27540
diff changeset
    52
lemma div_mult_self2 [simp]:
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27540
diff changeset
    53
  assumes "b \<noteq> 0"
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27540
diff changeset
    54
  shows "(a + b * c) div b = c + a div b"
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27540
diff changeset
    55
  using assms div_mult_self1 [of b a c] by (simp add: mult_commute)
26100
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
    56
27651
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27540
diff changeset
    57
lemma mod_mult_self1 [simp]: "(a + c * b) mod b = a mod b"
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27540
diff changeset
    58
proof (cases "b = 0")
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27540
diff changeset
    59
  case True then show ?thesis by simp
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27540
diff changeset
    60
next
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27540
diff changeset
    61
  case False
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27540
diff changeset
    62
  have "a + c * b = (a + c * b) div b * b + (a + c * b) mod b"
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27540
diff changeset
    63
    by (simp add: mod_div_equality)
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27540
diff changeset
    64
  also from False div_mult_self1 [of b a c] have
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27540
diff changeset
    65
    "\<dots> = (c + a div b) * b + (a + c * b) mod b"
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29509
diff changeset
    66
      by (simp add: algebra_simps)
27651
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27540
diff changeset
    67
  finally have "a = a div b * b + (a + c * b) mod b"
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27540
diff changeset
    68
    by (simp add: add_commute [of a] add_assoc left_distrib)
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27540
diff changeset
    69
  then have "a div b * b + (a + c * b) mod b = a div b * b + a mod b"
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27540
diff changeset
    70
    by (simp add: mod_div_equality)
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27540
diff changeset
    71
  then show ?thesis by simp
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27540
diff changeset
    72
qed
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27540
diff changeset
    73
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27540
diff changeset
    74
lemma mod_mult_self2 [simp]: "(a + b * c) mod b = a mod b"
30934
ed5377c2b0a3 tuned setups of CancelDivMod
haftmann
parents: 30930
diff changeset
    75
  by (simp add: mult_commute [of b])
27651
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27540
diff changeset
    76
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27540
diff changeset
    77
lemma div_mult_self1_is_id [simp]: "b \<noteq> 0 \<Longrightarrow> b * a div b = a"
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27540
diff changeset
    78
  using div_mult_self2 [of b 0 a] by simp
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27540
diff changeset
    79
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27540
diff changeset
    80
lemma div_mult_self2_is_id [simp]: "b \<noteq> 0 \<Longrightarrow> a * b div b = a"
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27540
diff changeset
    81
  using div_mult_self1 [of b 0 a] by simp
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27540
diff changeset
    82
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27540
diff changeset
    83
lemma mod_mult_self1_is_0 [simp]: "b * a mod b = 0"
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27540
diff changeset
    84
  using mod_mult_self2 [of 0 b a] by simp
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27540
diff changeset
    85
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27540
diff changeset
    86
lemma mod_mult_self2_is_0 [simp]: "a * b mod b = 0"
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27540
diff changeset
    87
  using mod_mult_self1 [of 0 a b] by simp
26062
16f334d7156a more abstract lemmas
haftmann
parents: 25947
diff changeset
    88
27651
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27540
diff changeset
    89
lemma div_by_1 [simp]: "a div 1 = a"
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27540
diff changeset
    90
  using div_mult_self2_is_id [of 1 a] zero_neq_one by simp
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27540
diff changeset
    91
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27540
diff changeset
    92
lemma mod_by_1 [simp]: "a mod 1 = 0"
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27540
diff changeset
    93
proof -
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27540
diff changeset
    94
  from mod_div_equality [of a one] div_by_1 have "a + a mod 1 = a" by simp
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27540
diff changeset
    95
  then have "a + a mod 1 = a + 0" by simp
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27540
diff changeset
    96
  then show ?thesis by (rule add_left_imp_eq)
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27540
diff changeset
    97
qed
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27540
diff changeset
    98
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27540
diff changeset
    99
lemma mod_self [simp]: "a mod a = 0"
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27540
diff changeset
   100
  using mod_mult_self2_is_0 [of 1] by simp
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27540
diff changeset
   101
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27540
diff changeset
   102
lemma div_self [simp]: "a \<noteq> 0 \<Longrightarrow> a div a = 1"
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27540
diff changeset
   103
  using div_mult_self2_is_id [of _ 1] by simp
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27540
diff changeset
   104
27676
55676111ed69 (re-)added simp rules for (_ + _) div/mod _
haftmann
parents: 27651
diff changeset
   105
lemma div_add_self1 [simp]:
27651
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27540
diff changeset
   106
  assumes "b \<noteq> 0"
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27540
diff changeset
   107
  shows "(b + a) div b = a div b + 1"
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27540
diff changeset
   108
  using assms div_mult_self1 [of b a 1] by (simp add: add_commute)
26062
16f334d7156a more abstract lemmas
haftmann
parents: 25947
diff changeset
   109
27676
55676111ed69 (re-)added simp rules for (_ + _) div/mod _
haftmann
parents: 27651
diff changeset
   110
lemma div_add_self2 [simp]:
27651
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27540
diff changeset
   111
  assumes "b \<noteq> 0"
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27540
diff changeset
   112
  shows "(a + b) div b = a div b + 1"
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27540
diff changeset
   113
  using assms div_add_self1 [of b a] by (simp add: add_commute)
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27540
diff changeset
   114
27676
55676111ed69 (re-)added simp rules for (_ + _) div/mod _
haftmann
parents: 27651
diff changeset
   115
lemma mod_add_self1 [simp]:
27651
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27540
diff changeset
   116
  "(b + a) mod b = a mod b"
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27540
diff changeset
   117
  using mod_mult_self1 [of a 1 b] by (simp add: add_commute)
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27540
diff changeset
   118
27676
55676111ed69 (re-)added simp rules for (_ + _) div/mod _
haftmann
parents: 27651
diff changeset
   119
lemma mod_add_self2 [simp]:
27651
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27540
diff changeset
   120
  "(a + b) mod b = a mod b"
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27540
diff changeset
   121
  using mod_mult_self1 [of a 1 b] by simp
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27540
diff changeset
   122
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27540
diff changeset
   123
lemma mod_div_decomp:
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27540
diff changeset
   124
  fixes a b
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27540
diff changeset
   125
  obtains q r where "q = a div b" and "r = a mod b"
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27540
diff changeset
   126
    and "a = q * b + r"
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27540
diff changeset
   127
proof -
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27540
diff changeset
   128
  from mod_div_equality have "a = a div b * b + a mod b" by simp
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27540
diff changeset
   129
  moreover have "a div b = a div b" ..
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27540
diff changeset
   130
  moreover have "a mod b = a mod b" ..
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27540
diff changeset
   131
  note that ultimately show thesis by blast
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27540
diff changeset
   132
qed
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27540
diff changeset
   133
45231
d85a2fdc586c replacing code_inline by code_unfold, removing obsolete code_unfold, code_inline del now that the ancient code generator is removed
bulwahn
parents: 44890
diff changeset
   134
lemma dvd_eq_mod_eq_0 [code]: "a dvd b \<longleftrightarrow> b mod a = 0"
25942
a52309ac4a4d added class semiring_div
haftmann
parents: 25571
diff changeset
   135
proof
a52309ac4a4d added class semiring_div
haftmann
parents: 25571
diff changeset
   136
  assume "b mod a = 0"
a52309ac4a4d added class semiring_div
haftmann
parents: 25571
diff changeset
   137
  with mod_div_equality [of b a] have "b div a * a = b" by simp
a52309ac4a4d added class semiring_div
haftmann
parents: 25571
diff changeset
   138
  then have "b = a * (b div a)" unfolding mult_commute ..
a52309ac4a4d added class semiring_div
haftmann
parents: 25571
diff changeset
   139
  then have "\<exists>c. b = a * c" ..
a52309ac4a4d added class semiring_div
haftmann
parents: 25571
diff changeset
   140
  then show "a dvd b" unfolding dvd_def .
a52309ac4a4d added class semiring_div
haftmann
parents: 25571
diff changeset
   141
next
a52309ac4a4d added class semiring_div
haftmann
parents: 25571
diff changeset
   142
  assume "a dvd b"
a52309ac4a4d added class semiring_div
haftmann
parents: 25571
diff changeset
   143
  then have "\<exists>c. b = a * c" unfolding dvd_def .
a52309ac4a4d added class semiring_div
haftmann
parents: 25571
diff changeset
   144
  then obtain c where "b = a * c" ..
a52309ac4a4d added class semiring_div
haftmann
parents: 25571
diff changeset
   145
  then have "b mod a = a * c mod a" by simp
a52309ac4a4d added class semiring_div
haftmann
parents: 25571
diff changeset
   146
  then have "b mod a = c * a mod a" by (simp add: mult_commute)
27651
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27540
diff changeset
   147
  then show "b mod a = 0" by simp
25942
a52309ac4a4d added class semiring_div
haftmann
parents: 25571
diff changeset
   148
qed
a52309ac4a4d added class semiring_div
haftmann
parents: 25571
diff changeset
   149
29403
fe17df4e4ab3 generalize some div/mod lemmas; remove type-specific proofs
huffman
parents: 29252
diff changeset
   150
lemma mod_div_trivial [simp]: "a mod b div b = 0"
fe17df4e4ab3 generalize some div/mod lemmas; remove type-specific proofs
huffman
parents: 29252
diff changeset
   151
proof (cases "b = 0")
fe17df4e4ab3 generalize some div/mod lemmas; remove type-specific proofs
huffman
parents: 29252
diff changeset
   152
  assume "b = 0"
fe17df4e4ab3 generalize some div/mod lemmas; remove type-specific proofs
huffman
parents: 29252
diff changeset
   153
  thus ?thesis by simp
fe17df4e4ab3 generalize some div/mod lemmas; remove type-specific proofs
huffman
parents: 29252
diff changeset
   154
next
fe17df4e4ab3 generalize some div/mod lemmas; remove type-specific proofs
huffman
parents: 29252
diff changeset
   155
  assume "b \<noteq> 0"
fe17df4e4ab3 generalize some div/mod lemmas; remove type-specific proofs
huffman
parents: 29252
diff changeset
   156
  hence "a div b + a mod b div b = (a mod b + a div b * b) div b"
fe17df4e4ab3 generalize some div/mod lemmas; remove type-specific proofs
huffman
parents: 29252
diff changeset
   157
    by (rule div_mult_self1 [symmetric])
fe17df4e4ab3 generalize some div/mod lemmas; remove type-specific proofs
huffman
parents: 29252
diff changeset
   158
  also have "\<dots> = a div b"
fe17df4e4ab3 generalize some div/mod lemmas; remove type-specific proofs
huffman
parents: 29252
diff changeset
   159
    by (simp only: mod_div_equality')
fe17df4e4ab3 generalize some div/mod lemmas; remove type-specific proofs
huffman
parents: 29252
diff changeset
   160
  also have "\<dots> = a div b + 0"
fe17df4e4ab3 generalize some div/mod lemmas; remove type-specific proofs
huffman
parents: 29252
diff changeset
   161
    by simp
fe17df4e4ab3 generalize some div/mod lemmas; remove type-specific proofs
huffman
parents: 29252
diff changeset
   162
  finally show ?thesis
fe17df4e4ab3 generalize some div/mod lemmas; remove type-specific proofs
huffman
parents: 29252
diff changeset
   163
    by (rule add_left_imp_eq)
fe17df4e4ab3 generalize some div/mod lemmas; remove type-specific proofs
huffman
parents: 29252
diff changeset
   164
qed
fe17df4e4ab3 generalize some div/mod lemmas; remove type-specific proofs
huffman
parents: 29252
diff changeset
   165
fe17df4e4ab3 generalize some div/mod lemmas; remove type-specific proofs
huffman
parents: 29252
diff changeset
   166
lemma mod_mod_trivial [simp]: "a mod b mod b = a mod b"
fe17df4e4ab3 generalize some div/mod lemmas; remove type-specific proofs
huffman
parents: 29252
diff changeset
   167
proof -
fe17df4e4ab3 generalize some div/mod lemmas; remove type-specific proofs
huffman
parents: 29252
diff changeset
   168
  have "a mod b mod b = (a mod b + a div b * b) mod b"
fe17df4e4ab3 generalize some div/mod lemmas; remove type-specific proofs
huffman
parents: 29252
diff changeset
   169
    by (simp only: mod_mult_self1)
fe17df4e4ab3 generalize some div/mod lemmas; remove type-specific proofs
huffman
parents: 29252
diff changeset
   170
  also have "\<dots> = a mod b"
fe17df4e4ab3 generalize some div/mod lemmas; remove type-specific proofs
huffman
parents: 29252
diff changeset
   171
    by (simp only: mod_div_equality')
fe17df4e4ab3 generalize some div/mod lemmas; remove type-specific proofs
huffman
parents: 29252
diff changeset
   172
  finally show ?thesis .
fe17df4e4ab3 generalize some div/mod lemmas; remove type-specific proofs
huffman
parents: 29252
diff changeset
   173
qed
fe17df4e4ab3 generalize some div/mod lemmas; remove type-specific proofs
huffman
parents: 29252
diff changeset
   174
29925
17d1e32ef867 dvd and setprod lemmas
nipkow
parents: 29667
diff changeset
   175
lemma dvd_imp_mod_0: "a dvd b \<Longrightarrow> b mod a = 0"
29948
cdf12a1cb963 Cleaned up IntDiv and removed subsumed lemmas.
nipkow
parents: 29925
diff changeset
   176
by (rule dvd_eq_mod_eq_0[THEN iffD1])
29925
17d1e32ef867 dvd and setprod lemmas
nipkow
parents: 29667
diff changeset
   177
17d1e32ef867 dvd and setprod lemmas
nipkow
parents: 29667
diff changeset
   178
lemma dvd_div_mult_self: "a dvd b \<Longrightarrow> (b div a) * a = b"
17d1e32ef867 dvd and setprod lemmas
nipkow
parents: 29667
diff changeset
   179
by (subst (2) mod_div_equality [of b a, symmetric]) (simp add:dvd_imp_mod_0)
17d1e32ef867 dvd and setprod lemmas
nipkow
parents: 29667
diff changeset
   180
33274
b6ff7db522b5 moved lemmas for dvd on nat to theories Nat and Power
haftmann
parents: 32010
diff changeset
   181
lemma dvd_mult_div_cancel: "a dvd b \<Longrightarrow> a * (b div a) = b"
b6ff7db522b5 moved lemmas for dvd on nat to theories Nat and Power
haftmann
parents: 32010
diff changeset
   182
by (drule dvd_div_mult_self) (simp add: mult_commute)
b6ff7db522b5 moved lemmas for dvd on nat to theories Nat and Power
haftmann
parents: 32010
diff changeset
   183
30052
410fefc247aa added dvd_div_mult
nipkow
parents: 30042
diff changeset
   184
lemma dvd_div_mult: "a dvd b \<Longrightarrow> (b div a) * c = b * c div a"
410fefc247aa added dvd_div_mult
nipkow
parents: 30042
diff changeset
   185
apply (cases "a = 0")
410fefc247aa added dvd_div_mult
nipkow
parents: 30042
diff changeset
   186
 apply simp
410fefc247aa added dvd_div_mult
nipkow
parents: 30042
diff changeset
   187
apply (auto simp: dvd_def mult_assoc)
410fefc247aa added dvd_div_mult
nipkow
parents: 30042
diff changeset
   188
done
410fefc247aa added dvd_div_mult
nipkow
parents: 30042
diff changeset
   189
29925
17d1e32ef867 dvd and setprod lemmas
nipkow
parents: 29667
diff changeset
   190
lemma div_dvd_div[simp]:
17d1e32ef867 dvd and setprod lemmas
nipkow
parents: 29667
diff changeset
   191
  "a dvd b \<Longrightarrow> a dvd c \<Longrightarrow> (b div a dvd c div a) = (b dvd c)"
17d1e32ef867 dvd and setprod lemmas
nipkow
parents: 29667
diff changeset
   192
apply (cases "a = 0")
17d1e32ef867 dvd and setprod lemmas
nipkow
parents: 29667
diff changeset
   193
 apply simp
17d1e32ef867 dvd and setprod lemmas
nipkow
parents: 29667
diff changeset
   194
apply (unfold dvd_def)
17d1e32ef867 dvd and setprod lemmas
nipkow
parents: 29667
diff changeset
   195
apply auto
17d1e32ef867 dvd and setprod lemmas
nipkow
parents: 29667
diff changeset
   196
 apply(blast intro:mult_assoc[symmetric])
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 44766
diff changeset
   197
apply(fastforce simp add: mult_assoc)
29925
17d1e32ef867 dvd and setprod lemmas
nipkow
parents: 29667
diff changeset
   198
done
17d1e32ef867 dvd and setprod lemmas
nipkow
parents: 29667
diff changeset
   199
30078
beee83623cc9 move lemma dvd_mod_imp_dvd into class semiring_div
huffman
parents: 30052
diff changeset
   200
lemma dvd_mod_imp_dvd: "[| k dvd m mod n;  k dvd n |] ==> k dvd m"
beee83623cc9 move lemma dvd_mod_imp_dvd into class semiring_div
huffman
parents: 30052
diff changeset
   201
  apply (subgoal_tac "k dvd (m div n) *n + m mod n")
beee83623cc9 move lemma dvd_mod_imp_dvd into class semiring_div
huffman
parents: 30052
diff changeset
   202
   apply (simp add: mod_div_equality)
beee83623cc9 move lemma dvd_mod_imp_dvd into class semiring_div
huffman
parents: 30052
diff changeset
   203
  apply (simp only: dvd_add dvd_mult)
beee83623cc9 move lemma dvd_mod_imp_dvd into class semiring_div
huffman
parents: 30052
diff changeset
   204
  done
beee83623cc9 move lemma dvd_mod_imp_dvd into class semiring_div
huffman
parents: 30052
diff changeset
   205
29403
fe17df4e4ab3 generalize some div/mod lemmas; remove type-specific proofs
huffman
parents: 29252
diff changeset
   206
text {* Addition respects modular equivalence. *}
fe17df4e4ab3 generalize some div/mod lemmas; remove type-specific proofs
huffman
parents: 29252
diff changeset
   207
fe17df4e4ab3 generalize some div/mod lemmas; remove type-specific proofs
huffman
parents: 29252
diff changeset
   208
lemma mod_add_left_eq: "(a + b) mod c = (a mod c + b) mod c"
fe17df4e4ab3 generalize some div/mod lemmas; remove type-specific proofs
huffman
parents: 29252
diff changeset
   209
proof -
fe17df4e4ab3 generalize some div/mod lemmas; remove type-specific proofs
huffman
parents: 29252
diff changeset
   210
  have "(a + b) mod c = (a div c * c + a mod c + b) mod c"
fe17df4e4ab3 generalize some div/mod lemmas; remove type-specific proofs
huffman
parents: 29252
diff changeset
   211
    by (simp only: mod_div_equality)
fe17df4e4ab3 generalize some div/mod lemmas; remove type-specific proofs
huffman
parents: 29252
diff changeset
   212
  also have "\<dots> = (a mod c + b + a div c * c) mod c"
fe17df4e4ab3 generalize some div/mod lemmas; remove type-specific proofs
huffman
parents: 29252
diff changeset
   213
    by (simp only: add_ac)
fe17df4e4ab3 generalize some div/mod lemmas; remove type-specific proofs
huffman
parents: 29252
diff changeset
   214
  also have "\<dots> = (a mod c + b) mod c"
fe17df4e4ab3 generalize some div/mod lemmas; remove type-specific proofs
huffman
parents: 29252
diff changeset
   215
    by (rule mod_mult_self1)
fe17df4e4ab3 generalize some div/mod lemmas; remove type-specific proofs
huffman
parents: 29252
diff changeset
   216
  finally show ?thesis .
fe17df4e4ab3 generalize some div/mod lemmas; remove type-specific proofs
huffman
parents: 29252
diff changeset
   217
qed
fe17df4e4ab3 generalize some div/mod lemmas; remove type-specific proofs
huffman
parents: 29252
diff changeset
   218
fe17df4e4ab3 generalize some div/mod lemmas; remove type-specific proofs
huffman
parents: 29252
diff changeset
   219
lemma mod_add_right_eq: "(a + b) mod c = (a + b mod c) mod c"
fe17df4e4ab3 generalize some div/mod lemmas; remove type-specific proofs
huffman
parents: 29252
diff changeset
   220
proof -
fe17df4e4ab3 generalize some div/mod lemmas; remove type-specific proofs
huffman
parents: 29252
diff changeset
   221
  have "(a + b) mod c = (a + (b div c * c + b mod c)) mod c"
fe17df4e4ab3 generalize some div/mod lemmas; remove type-specific proofs
huffman
parents: 29252
diff changeset
   222
    by (simp only: mod_div_equality)
fe17df4e4ab3 generalize some div/mod lemmas; remove type-specific proofs
huffman
parents: 29252
diff changeset
   223
  also have "\<dots> = (a + b mod c + b div c * c) mod c"
fe17df4e4ab3 generalize some div/mod lemmas; remove type-specific proofs
huffman
parents: 29252
diff changeset
   224
    by (simp only: add_ac)
fe17df4e4ab3 generalize some div/mod lemmas; remove type-specific proofs
huffman
parents: 29252
diff changeset
   225
  also have "\<dots> = (a + b mod c) mod c"
fe17df4e4ab3 generalize some div/mod lemmas; remove type-specific proofs
huffman
parents: 29252
diff changeset
   226
    by (rule mod_mult_self1)
fe17df4e4ab3 generalize some div/mod lemmas; remove type-specific proofs
huffman
parents: 29252
diff changeset
   227
  finally show ?thesis .
fe17df4e4ab3 generalize some div/mod lemmas; remove type-specific proofs
huffman
parents: 29252
diff changeset
   228
qed
fe17df4e4ab3 generalize some div/mod lemmas; remove type-specific proofs
huffman
parents: 29252
diff changeset
   229
fe17df4e4ab3 generalize some div/mod lemmas; remove type-specific proofs
huffman
parents: 29252
diff changeset
   230
lemma mod_add_eq: "(a + b) mod c = (a mod c + b mod c) mod c"
fe17df4e4ab3 generalize some div/mod lemmas; remove type-specific proofs
huffman
parents: 29252
diff changeset
   231
by (rule trans [OF mod_add_left_eq mod_add_right_eq])
fe17df4e4ab3 generalize some div/mod lemmas; remove type-specific proofs
huffman
parents: 29252
diff changeset
   232
fe17df4e4ab3 generalize some div/mod lemmas; remove type-specific proofs
huffman
parents: 29252
diff changeset
   233
lemma mod_add_cong:
fe17df4e4ab3 generalize some div/mod lemmas; remove type-specific proofs
huffman
parents: 29252
diff changeset
   234
  assumes "a mod c = a' mod c"
fe17df4e4ab3 generalize some div/mod lemmas; remove type-specific proofs
huffman
parents: 29252
diff changeset
   235
  assumes "b mod c = b' mod c"
fe17df4e4ab3 generalize some div/mod lemmas; remove type-specific proofs
huffman
parents: 29252
diff changeset
   236
  shows "(a + b) mod c = (a' + b') mod c"
fe17df4e4ab3 generalize some div/mod lemmas; remove type-specific proofs
huffman
parents: 29252
diff changeset
   237
proof -
fe17df4e4ab3 generalize some div/mod lemmas; remove type-specific proofs
huffman
parents: 29252
diff changeset
   238
  have "(a mod c + b mod c) mod c = (a' mod c + b' mod c) mod c"
fe17df4e4ab3 generalize some div/mod lemmas; remove type-specific proofs
huffman
parents: 29252
diff changeset
   239
    unfolding assms ..
fe17df4e4ab3 generalize some div/mod lemmas; remove type-specific proofs
huffman
parents: 29252
diff changeset
   240
  thus ?thesis
fe17df4e4ab3 generalize some div/mod lemmas; remove type-specific proofs
huffman
parents: 29252
diff changeset
   241
    by (simp only: mod_add_eq [symmetric])
fe17df4e4ab3 generalize some div/mod lemmas; remove type-specific proofs
huffman
parents: 29252
diff changeset
   242
qed
fe17df4e4ab3 generalize some div/mod lemmas; remove type-specific proofs
huffman
parents: 29252
diff changeset
   243
30923
2697a1d1d34a more coherent developement in Divides.thy and IntDiv.thy
haftmann
parents: 30840
diff changeset
   244
lemma div_add [simp]: "z dvd x \<Longrightarrow> z dvd y
30837
3d4832d9f7e4 added strong_setprod_cong[cong] (in analogy with setsum)
nipkow
parents: 30729
diff changeset
   245
  \<Longrightarrow> (x + y) div z = x div z + y div z"
30923
2697a1d1d34a more coherent developement in Divides.thy and IntDiv.thy
haftmann
parents: 30840
diff changeset
   246
by (cases "z = 0", simp, unfold dvd_def, auto simp add: algebra_simps)
30837
3d4832d9f7e4 added strong_setprod_cong[cong] (in analogy with setsum)
nipkow
parents: 30729
diff changeset
   247
29403
fe17df4e4ab3 generalize some div/mod lemmas; remove type-specific proofs
huffman
parents: 29252
diff changeset
   248
text {* Multiplication respects modular equivalence. *}
fe17df4e4ab3 generalize some div/mod lemmas; remove type-specific proofs
huffman
parents: 29252
diff changeset
   249
fe17df4e4ab3 generalize some div/mod lemmas; remove type-specific proofs
huffman
parents: 29252
diff changeset
   250
lemma mod_mult_left_eq: "(a * b) mod c = ((a mod c) * b) mod c"
fe17df4e4ab3 generalize some div/mod lemmas; remove type-specific proofs
huffman
parents: 29252
diff changeset
   251
proof -
fe17df4e4ab3 generalize some div/mod lemmas; remove type-specific proofs
huffman
parents: 29252
diff changeset
   252
  have "(a * b) mod c = ((a div c * c + a mod c) * b) mod c"
fe17df4e4ab3 generalize some div/mod lemmas; remove type-specific proofs
huffman
parents: 29252
diff changeset
   253
    by (simp only: mod_div_equality)
fe17df4e4ab3 generalize some div/mod lemmas; remove type-specific proofs
huffman
parents: 29252
diff changeset
   254
  also have "\<dots> = (a mod c * b + a div c * b * c) mod c"
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29509
diff changeset
   255
    by (simp only: algebra_simps)
29403
fe17df4e4ab3 generalize some div/mod lemmas; remove type-specific proofs
huffman
parents: 29252
diff changeset
   256
  also have "\<dots> = (a mod c * b) mod c"
fe17df4e4ab3 generalize some div/mod lemmas; remove type-specific proofs
huffman
parents: 29252
diff changeset
   257
    by (rule mod_mult_self1)
fe17df4e4ab3 generalize some div/mod lemmas; remove type-specific proofs
huffman
parents: 29252
diff changeset
   258
  finally show ?thesis .
fe17df4e4ab3 generalize some div/mod lemmas; remove type-specific proofs
huffman
parents: 29252
diff changeset
   259
qed
fe17df4e4ab3 generalize some div/mod lemmas; remove type-specific proofs
huffman
parents: 29252
diff changeset
   260
fe17df4e4ab3 generalize some div/mod lemmas; remove type-specific proofs
huffman
parents: 29252
diff changeset
   261
lemma mod_mult_right_eq: "(a * b) mod c = (a * (b mod c)) mod c"
fe17df4e4ab3 generalize some div/mod lemmas; remove type-specific proofs
huffman
parents: 29252
diff changeset
   262
proof -
fe17df4e4ab3 generalize some div/mod lemmas; remove type-specific proofs
huffman
parents: 29252
diff changeset
   263
  have "(a * b) mod c = (a * (b div c * c + b mod c)) mod c"
fe17df4e4ab3 generalize some div/mod lemmas; remove type-specific proofs
huffman
parents: 29252
diff changeset
   264
    by (simp only: mod_div_equality)
fe17df4e4ab3 generalize some div/mod lemmas; remove type-specific proofs
huffman
parents: 29252
diff changeset
   265
  also have "\<dots> = (a * (b mod c) + a * (b div c) * c) mod c"
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29509
diff changeset
   266
    by (simp only: algebra_simps)
29403
fe17df4e4ab3 generalize some div/mod lemmas; remove type-specific proofs
huffman
parents: 29252
diff changeset
   267
  also have "\<dots> = (a * (b mod c)) mod c"
fe17df4e4ab3 generalize some div/mod lemmas; remove type-specific proofs
huffman
parents: 29252
diff changeset
   268
    by (rule mod_mult_self1)
fe17df4e4ab3 generalize some div/mod lemmas; remove type-specific proofs
huffman
parents: 29252
diff changeset
   269
  finally show ?thesis .
fe17df4e4ab3 generalize some div/mod lemmas; remove type-specific proofs
huffman
parents: 29252
diff changeset
   270
qed
fe17df4e4ab3 generalize some div/mod lemmas; remove type-specific proofs
huffman
parents: 29252
diff changeset
   271
fe17df4e4ab3 generalize some div/mod lemmas; remove type-specific proofs
huffman
parents: 29252
diff changeset
   272
lemma mod_mult_eq: "(a * b) mod c = ((a mod c) * (b mod c)) mod c"
fe17df4e4ab3 generalize some div/mod lemmas; remove type-specific proofs
huffman
parents: 29252
diff changeset
   273
by (rule trans [OF mod_mult_left_eq mod_mult_right_eq])
fe17df4e4ab3 generalize some div/mod lemmas; remove type-specific proofs
huffman
parents: 29252
diff changeset
   274
fe17df4e4ab3 generalize some div/mod lemmas; remove type-specific proofs
huffman
parents: 29252
diff changeset
   275
lemma mod_mult_cong:
fe17df4e4ab3 generalize some div/mod lemmas; remove type-specific proofs
huffman
parents: 29252
diff changeset
   276
  assumes "a mod c = a' mod c"
fe17df4e4ab3 generalize some div/mod lemmas; remove type-specific proofs
huffman
parents: 29252
diff changeset
   277
  assumes "b mod c = b' mod c"
fe17df4e4ab3 generalize some div/mod lemmas; remove type-specific proofs
huffman
parents: 29252
diff changeset
   278
  shows "(a * b) mod c = (a' * b') mod c"
fe17df4e4ab3 generalize some div/mod lemmas; remove type-specific proofs
huffman
parents: 29252
diff changeset
   279
proof -
fe17df4e4ab3 generalize some div/mod lemmas; remove type-specific proofs
huffman
parents: 29252
diff changeset
   280
  have "(a mod c * (b mod c)) mod c = (a' mod c * (b' mod c)) mod c"
fe17df4e4ab3 generalize some div/mod lemmas; remove type-specific proofs
huffman
parents: 29252
diff changeset
   281
    unfolding assms ..
fe17df4e4ab3 generalize some div/mod lemmas; remove type-specific proofs
huffman
parents: 29252
diff changeset
   282
  thus ?thesis
fe17df4e4ab3 generalize some div/mod lemmas; remove type-specific proofs
huffman
parents: 29252
diff changeset
   283
    by (simp only: mod_mult_eq [symmetric])
fe17df4e4ab3 generalize some div/mod lemmas; remove type-specific proofs
huffman
parents: 29252
diff changeset
   284
qed
fe17df4e4ab3 generalize some div/mod lemmas; remove type-specific proofs
huffman
parents: 29252
diff changeset
   285
29404
ee15ccdeaa72 generalize zmod_zmod_cancel -> mod_mod_cancel
huffman
parents: 29403
diff changeset
   286
lemma mod_mod_cancel:
ee15ccdeaa72 generalize zmod_zmod_cancel -> mod_mod_cancel
huffman
parents: 29403
diff changeset
   287
  assumes "c dvd b"
ee15ccdeaa72 generalize zmod_zmod_cancel -> mod_mod_cancel
huffman
parents: 29403
diff changeset
   288
  shows "a mod b mod c = a mod c"
ee15ccdeaa72 generalize zmod_zmod_cancel -> mod_mod_cancel
huffman
parents: 29403
diff changeset
   289
proof -
ee15ccdeaa72 generalize zmod_zmod_cancel -> mod_mod_cancel
huffman
parents: 29403
diff changeset
   290
  from `c dvd b` obtain k where "b = c * k"
ee15ccdeaa72 generalize zmod_zmod_cancel -> mod_mod_cancel
huffman
parents: 29403
diff changeset
   291
    by (rule dvdE)
ee15ccdeaa72 generalize zmod_zmod_cancel -> mod_mod_cancel
huffman
parents: 29403
diff changeset
   292
  have "a mod b mod c = a mod (c * k) mod c"
ee15ccdeaa72 generalize zmod_zmod_cancel -> mod_mod_cancel
huffman
parents: 29403
diff changeset
   293
    by (simp only: `b = c * k`)
ee15ccdeaa72 generalize zmod_zmod_cancel -> mod_mod_cancel
huffman
parents: 29403
diff changeset
   294
  also have "\<dots> = (a mod (c * k) + a div (c * k) * k * c) mod c"
ee15ccdeaa72 generalize zmod_zmod_cancel -> mod_mod_cancel
huffman
parents: 29403
diff changeset
   295
    by (simp only: mod_mult_self1)
ee15ccdeaa72 generalize zmod_zmod_cancel -> mod_mod_cancel
huffman
parents: 29403
diff changeset
   296
  also have "\<dots> = (a div (c * k) * (c * k) + a mod (c * k)) mod c"
ee15ccdeaa72 generalize zmod_zmod_cancel -> mod_mod_cancel
huffman
parents: 29403
diff changeset
   297
    by (simp only: add_ac mult_ac)
ee15ccdeaa72 generalize zmod_zmod_cancel -> mod_mod_cancel
huffman
parents: 29403
diff changeset
   298
  also have "\<dots> = a mod c"
ee15ccdeaa72 generalize zmod_zmod_cancel -> mod_mod_cancel
huffman
parents: 29403
diff changeset
   299
    by (simp only: mod_div_equality)
ee15ccdeaa72 generalize zmod_zmod_cancel -> mod_mod_cancel
huffman
parents: 29403
diff changeset
   300
  finally show ?thesis .
ee15ccdeaa72 generalize zmod_zmod_cancel -> mod_mod_cancel
huffman
parents: 29403
diff changeset
   301
qed
ee15ccdeaa72 generalize zmod_zmod_cancel -> mod_mod_cancel
huffman
parents: 29403
diff changeset
   302
30930
11010e5f18f0 tightended specification of class semiring_div
haftmann
parents: 30923
diff changeset
   303
lemma div_mult_div_if_dvd:
11010e5f18f0 tightended specification of class semiring_div
haftmann
parents: 30923
diff changeset
   304
  "y dvd x \<Longrightarrow> z dvd w \<Longrightarrow> (x div y) * (w div z) = (x * w) div (y * z)"
11010e5f18f0 tightended specification of class semiring_div
haftmann
parents: 30923
diff changeset
   305
  apply (cases "y = 0", simp)
11010e5f18f0 tightended specification of class semiring_div
haftmann
parents: 30923
diff changeset
   306
  apply (cases "z = 0", simp)
11010e5f18f0 tightended specification of class semiring_div
haftmann
parents: 30923
diff changeset
   307
  apply (auto elim!: dvdE simp add: algebra_simps)
30476
0a41b0662264 added div lemmas
nipkow
parents: 30242
diff changeset
   308
  apply (subst mult_assoc [symmetric])
0a41b0662264 added div lemmas
nipkow
parents: 30242
diff changeset
   309
  apply (simp add: no_zero_divisors)
30930
11010e5f18f0 tightended specification of class semiring_div
haftmann
parents: 30923
diff changeset
   310
  done
11010e5f18f0 tightended specification of class semiring_div
haftmann
parents: 30923
diff changeset
   311
35367
45a193f0ed0c lemma div_mult_swap, dvd_div_eq_mult, dvd_div_div_eq_mult
haftmann
parents: 35216
diff changeset
   312
lemma div_mult_swap:
45a193f0ed0c lemma div_mult_swap, dvd_div_eq_mult, dvd_div_div_eq_mult
haftmann
parents: 35216
diff changeset
   313
  assumes "c dvd b"
45a193f0ed0c lemma div_mult_swap, dvd_div_eq_mult, dvd_div_div_eq_mult
haftmann
parents: 35216
diff changeset
   314
  shows "a * (b div c) = (a * b) div c"
45a193f0ed0c lemma div_mult_swap, dvd_div_eq_mult, dvd_div_div_eq_mult
haftmann
parents: 35216
diff changeset
   315
proof -
45a193f0ed0c lemma div_mult_swap, dvd_div_eq_mult, dvd_div_div_eq_mult
haftmann
parents: 35216
diff changeset
   316
  from assms have "b div c * (a div 1) = b * a div (c * 1)"
45a193f0ed0c lemma div_mult_swap, dvd_div_eq_mult, dvd_div_div_eq_mult
haftmann
parents: 35216
diff changeset
   317
    by (simp only: div_mult_div_if_dvd one_dvd)
45a193f0ed0c lemma div_mult_swap, dvd_div_eq_mult, dvd_div_div_eq_mult
haftmann
parents: 35216
diff changeset
   318
  then show ?thesis by (simp add: mult_commute)
45a193f0ed0c lemma div_mult_swap, dvd_div_eq_mult, dvd_div_div_eq_mult
haftmann
parents: 35216
diff changeset
   319
qed
45a193f0ed0c lemma div_mult_swap, dvd_div_eq_mult, dvd_div_div_eq_mult
haftmann
parents: 35216
diff changeset
   320
   
30930
11010e5f18f0 tightended specification of class semiring_div
haftmann
parents: 30923
diff changeset
   321
lemma div_mult_mult2 [simp]:
11010e5f18f0 tightended specification of class semiring_div
haftmann
parents: 30923
diff changeset
   322
  "c \<noteq> 0 \<Longrightarrow> (a * c) div (b * c) = a div b"
11010e5f18f0 tightended specification of class semiring_div
haftmann
parents: 30923
diff changeset
   323
  by (drule div_mult_mult1) (simp add: mult_commute)
11010e5f18f0 tightended specification of class semiring_div
haftmann
parents: 30923
diff changeset
   324
11010e5f18f0 tightended specification of class semiring_div
haftmann
parents: 30923
diff changeset
   325
lemma div_mult_mult1_if [simp]:
11010e5f18f0 tightended specification of class semiring_div
haftmann
parents: 30923
diff changeset
   326
  "(c * a) div (c * b) = (if c = 0 then 0 else a div b)"
11010e5f18f0 tightended specification of class semiring_div
haftmann
parents: 30923
diff changeset
   327
  by simp_all
30476
0a41b0662264 added div lemmas
nipkow
parents: 30242
diff changeset
   328
30930
11010e5f18f0 tightended specification of class semiring_div
haftmann
parents: 30923
diff changeset
   329
lemma mod_mult_mult1:
11010e5f18f0 tightended specification of class semiring_div
haftmann
parents: 30923
diff changeset
   330
  "(c * a) mod (c * b) = c * (a mod b)"
11010e5f18f0 tightended specification of class semiring_div
haftmann
parents: 30923
diff changeset
   331
proof (cases "c = 0")
11010e5f18f0 tightended specification of class semiring_div
haftmann
parents: 30923
diff changeset
   332
  case True then show ?thesis by simp
11010e5f18f0 tightended specification of class semiring_div
haftmann
parents: 30923
diff changeset
   333
next
11010e5f18f0 tightended specification of class semiring_div
haftmann
parents: 30923
diff changeset
   334
  case False
11010e5f18f0 tightended specification of class semiring_div
haftmann
parents: 30923
diff changeset
   335
  from mod_div_equality
11010e5f18f0 tightended specification of class semiring_div
haftmann
parents: 30923
diff changeset
   336
  have "((c * a) div (c * b)) * (c * b) + (c * a) mod (c * b) = c * a" .
11010e5f18f0 tightended specification of class semiring_div
haftmann
parents: 30923
diff changeset
   337
  with False have "c * ((a div b) * b + a mod b) + (c * a) mod (c * b)
11010e5f18f0 tightended specification of class semiring_div
haftmann
parents: 30923
diff changeset
   338
    = c * a + c * (a mod b)" by (simp add: algebra_simps)
11010e5f18f0 tightended specification of class semiring_div
haftmann
parents: 30923
diff changeset
   339
  with mod_div_equality show ?thesis by simp 
11010e5f18f0 tightended specification of class semiring_div
haftmann
parents: 30923
diff changeset
   340
qed
11010e5f18f0 tightended specification of class semiring_div
haftmann
parents: 30923
diff changeset
   341
  
11010e5f18f0 tightended specification of class semiring_div
haftmann
parents: 30923
diff changeset
   342
lemma mod_mult_mult2:
11010e5f18f0 tightended specification of class semiring_div
haftmann
parents: 30923
diff changeset
   343
  "(a * c) mod (b * c) = (a mod b) * c"
11010e5f18f0 tightended specification of class semiring_div
haftmann
parents: 30923
diff changeset
   344
  using mod_mult_mult1 [of c a b] by (simp add: mult_commute)
11010e5f18f0 tightended specification of class semiring_div
haftmann
parents: 30923
diff changeset
   345
31662
57f7ef0dba8e generalize lemmas dvd_mod and dvd_mod_iff to class semiring_div
huffman
parents: 31661
diff changeset
   346
lemma dvd_mod: "k dvd m \<Longrightarrow> k dvd n \<Longrightarrow> k dvd (m mod n)"
57f7ef0dba8e generalize lemmas dvd_mod and dvd_mod_iff to class semiring_div
huffman
parents: 31661
diff changeset
   347
  unfolding dvd_def by (auto simp add: mod_mult_mult1)
57f7ef0dba8e generalize lemmas dvd_mod and dvd_mod_iff to class semiring_div
huffman
parents: 31661
diff changeset
   348
57f7ef0dba8e generalize lemmas dvd_mod and dvd_mod_iff to class semiring_div
huffman
parents: 31661
diff changeset
   349
lemma dvd_mod_iff: "k dvd n \<Longrightarrow> k dvd (m mod n) \<longleftrightarrow> k dvd m"
57f7ef0dba8e generalize lemmas dvd_mod and dvd_mod_iff to class semiring_div
huffman
parents: 31661
diff changeset
   350
by (blast intro: dvd_mod_imp_dvd dvd_mod)
57f7ef0dba8e generalize lemmas dvd_mod and dvd_mod_iff to class semiring_div
huffman
parents: 31661
diff changeset
   351
31009
41fd307cab30 dropped reference to class recpower and lemma duplicate
haftmann
parents: 30934
diff changeset
   352
lemma div_power:
31661
1e252b8b2334 move lemma div_power into semiring_div context; class ring_div inherits from idom
huffman
parents: 31009
diff changeset
   353
  "y dvd x \<Longrightarrow> (x div y) ^ n = x ^ n div y ^ n"
30476
0a41b0662264 added div lemmas
nipkow
parents: 30242
diff changeset
   354
apply (induct n)
0a41b0662264 added div lemmas
nipkow
parents: 30242
diff changeset
   355
 apply simp
0a41b0662264 added div lemmas
nipkow
parents: 30242
diff changeset
   356
apply(simp add: div_mult_div_if_dvd dvd_power_same)
0a41b0662264 added div lemmas
nipkow
parents: 30242
diff changeset
   357
done
0a41b0662264 added div lemmas
nipkow
parents: 30242
diff changeset
   358
35367
45a193f0ed0c lemma div_mult_swap, dvd_div_eq_mult, dvd_div_div_eq_mult
haftmann
parents: 35216
diff changeset
   359
lemma dvd_div_eq_mult:
45a193f0ed0c lemma div_mult_swap, dvd_div_eq_mult, dvd_div_div_eq_mult
haftmann
parents: 35216
diff changeset
   360
  assumes "a \<noteq> 0" and "a dvd b"  
45a193f0ed0c lemma div_mult_swap, dvd_div_eq_mult, dvd_div_div_eq_mult
haftmann
parents: 35216
diff changeset
   361
  shows "b div a = c \<longleftrightarrow> b = c * a"
45a193f0ed0c lemma div_mult_swap, dvd_div_eq_mult, dvd_div_div_eq_mult
haftmann
parents: 35216
diff changeset
   362
proof
45a193f0ed0c lemma div_mult_swap, dvd_div_eq_mult, dvd_div_div_eq_mult
haftmann
parents: 35216
diff changeset
   363
  assume "b = c * a"
45a193f0ed0c lemma div_mult_swap, dvd_div_eq_mult, dvd_div_div_eq_mult
haftmann
parents: 35216
diff changeset
   364
  then show "b div a = c" by (simp add: assms)
45a193f0ed0c lemma div_mult_swap, dvd_div_eq_mult, dvd_div_div_eq_mult
haftmann
parents: 35216
diff changeset
   365
next
45a193f0ed0c lemma div_mult_swap, dvd_div_eq_mult, dvd_div_div_eq_mult
haftmann
parents: 35216
diff changeset
   366
  assume "b div a = c"
45a193f0ed0c lemma div_mult_swap, dvd_div_eq_mult, dvd_div_div_eq_mult
haftmann
parents: 35216
diff changeset
   367
  then have "b div a * a = c * a" by simp
45a193f0ed0c lemma div_mult_swap, dvd_div_eq_mult, dvd_div_div_eq_mult
haftmann
parents: 35216
diff changeset
   368
  moreover from `a dvd b` have "b div a * a = b" by (simp add: dvd_div_mult_self)
45a193f0ed0c lemma div_mult_swap, dvd_div_eq_mult, dvd_div_div_eq_mult
haftmann
parents: 35216
diff changeset
   369
  ultimately show "b = c * a" by simp
45a193f0ed0c lemma div_mult_swap, dvd_div_eq_mult, dvd_div_div_eq_mult
haftmann
parents: 35216
diff changeset
   370
qed
45a193f0ed0c lemma div_mult_swap, dvd_div_eq_mult, dvd_div_div_eq_mult
haftmann
parents: 35216
diff changeset
   371
   
45a193f0ed0c lemma div_mult_swap, dvd_div_eq_mult, dvd_div_div_eq_mult
haftmann
parents: 35216
diff changeset
   372
lemma dvd_div_div_eq_mult:
45a193f0ed0c lemma div_mult_swap, dvd_div_eq_mult, dvd_div_div_eq_mult
haftmann
parents: 35216
diff changeset
   373
  assumes "a \<noteq> 0" "c \<noteq> 0" and "a dvd b" "c dvd d"
45a193f0ed0c lemma div_mult_swap, dvd_div_eq_mult, dvd_div_div_eq_mult
haftmann
parents: 35216
diff changeset
   374
  shows "b div a = d div c \<longleftrightarrow> b * c = a * d"
45a193f0ed0c lemma div_mult_swap, dvd_div_eq_mult, dvd_div_div_eq_mult
haftmann
parents: 35216
diff changeset
   375
  using assms by (auto simp add: mult_commute [of _ a] dvd_div_mult_self dvd_div_eq_mult div_mult_swap intro: sym)
45a193f0ed0c lemma div_mult_swap, dvd_div_eq_mult, dvd_div_div_eq_mult
haftmann
parents: 35216
diff changeset
   376
31661
1e252b8b2334 move lemma div_power into semiring_div context; class ring_div inherits from idom
huffman
parents: 31009
diff changeset
   377
end
1e252b8b2334 move lemma div_power into semiring_div context; class ring_div inherits from idom
huffman
parents: 31009
diff changeset
   378
35673
178caf872f95 weakend class ring_div; tuned
haftmann
parents: 35644
diff changeset
   379
class ring_div = semiring_div + comm_ring_1
29405
98ab21b14f09 add class ring_div; generalize mod/diff/minus proofs for class ring_div
huffman
parents: 29404
diff changeset
   380
begin
98ab21b14f09 add class ring_div; generalize mod/diff/minus proofs for class ring_div
huffman
parents: 29404
diff changeset
   381
36634
f9b43d197d16 a ring_div is a ring_1_no_zero_divisors
haftmann
parents: 35815
diff changeset
   382
subclass ring_1_no_zero_divisors ..
f9b43d197d16 a ring_div is a ring_1_no_zero_divisors
haftmann
parents: 35815
diff changeset
   383
29405
98ab21b14f09 add class ring_div; generalize mod/diff/minus proofs for class ring_div
huffman
parents: 29404
diff changeset
   384
text {* Negation respects modular equivalence. *}
98ab21b14f09 add class ring_div; generalize mod/diff/minus proofs for class ring_div
huffman
parents: 29404
diff changeset
   385
98ab21b14f09 add class ring_div; generalize mod/diff/minus proofs for class ring_div
huffman
parents: 29404
diff changeset
   386
lemma mod_minus_eq: "(- a) mod b = (- (a mod b)) mod b"
98ab21b14f09 add class ring_div; generalize mod/diff/minus proofs for class ring_div
huffman
parents: 29404
diff changeset
   387
proof -
98ab21b14f09 add class ring_div; generalize mod/diff/minus proofs for class ring_div
huffman
parents: 29404
diff changeset
   388
  have "(- a) mod b = (- (a div b * b + a mod b)) mod b"
98ab21b14f09 add class ring_div; generalize mod/diff/minus proofs for class ring_div
huffman
parents: 29404
diff changeset
   389
    by (simp only: mod_div_equality)
98ab21b14f09 add class ring_div; generalize mod/diff/minus proofs for class ring_div
huffman
parents: 29404
diff changeset
   390
  also have "\<dots> = (- (a mod b) + - (a div b) * b) mod b"
98ab21b14f09 add class ring_div; generalize mod/diff/minus proofs for class ring_div
huffman
parents: 29404
diff changeset
   391
    by (simp only: minus_add_distrib minus_mult_left add_ac)
98ab21b14f09 add class ring_div; generalize mod/diff/minus proofs for class ring_div
huffman
parents: 29404
diff changeset
   392
  also have "\<dots> = (- (a mod b)) mod b"
98ab21b14f09 add class ring_div; generalize mod/diff/minus proofs for class ring_div
huffman
parents: 29404
diff changeset
   393
    by (rule mod_mult_self1)
98ab21b14f09 add class ring_div; generalize mod/diff/minus proofs for class ring_div
huffman
parents: 29404
diff changeset
   394
  finally show ?thesis .
98ab21b14f09 add class ring_div; generalize mod/diff/minus proofs for class ring_div
huffman
parents: 29404
diff changeset
   395
qed
98ab21b14f09 add class ring_div; generalize mod/diff/minus proofs for class ring_div
huffman
parents: 29404
diff changeset
   396
98ab21b14f09 add class ring_div; generalize mod/diff/minus proofs for class ring_div
huffman
parents: 29404
diff changeset
   397
lemma mod_minus_cong:
98ab21b14f09 add class ring_div; generalize mod/diff/minus proofs for class ring_div
huffman
parents: 29404
diff changeset
   398
  assumes "a mod b = a' mod b"
98ab21b14f09 add class ring_div; generalize mod/diff/minus proofs for class ring_div
huffman
parents: 29404
diff changeset
   399
  shows "(- a) mod b = (- a') mod b"
98ab21b14f09 add class ring_div; generalize mod/diff/minus proofs for class ring_div
huffman
parents: 29404
diff changeset
   400
proof -
98ab21b14f09 add class ring_div; generalize mod/diff/minus proofs for class ring_div
huffman
parents: 29404
diff changeset
   401
  have "(- (a mod b)) mod b = (- (a' mod b)) mod b"
98ab21b14f09 add class ring_div; generalize mod/diff/minus proofs for class ring_div
huffman
parents: 29404
diff changeset
   402
    unfolding assms ..
98ab21b14f09 add class ring_div; generalize mod/diff/minus proofs for class ring_div
huffman
parents: 29404
diff changeset
   403
  thus ?thesis
98ab21b14f09 add class ring_div; generalize mod/diff/minus proofs for class ring_div
huffman
parents: 29404
diff changeset
   404
    by (simp only: mod_minus_eq [symmetric])
98ab21b14f09 add class ring_div; generalize mod/diff/minus proofs for class ring_div
huffman
parents: 29404
diff changeset
   405
qed
98ab21b14f09 add class ring_div; generalize mod/diff/minus proofs for class ring_div
huffman
parents: 29404
diff changeset
   406
98ab21b14f09 add class ring_div; generalize mod/diff/minus proofs for class ring_div
huffman
parents: 29404
diff changeset
   407
text {* Subtraction respects modular equivalence. *}
98ab21b14f09 add class ring_div; generalize mod/diff/minus proofs for class ring_div
huffman
parents: 29404
diff changeset
   408
98ab21b14f09 add class ring_div; generalize mod/diff/minus proofs for class ring_div
huffman
parents: 29404
diff changeset
   409
lemma mod_diff_left_eq: "(a - b) mod c = (a mod c - b) mod c"
98ab21b14f09 add class ring_div; generalize mod/diff/minus proofs for class ring_div
huffman
parents: 29404
diff changeset
   410
  unfolding diff_minus
98ab21b14f09 add class ring_div; generalize mod/diff/minus proofs for class ring_div
huffman
parents: 29404
diff changeset
   411
  by (intro mod_add_cong mod_minus_cong) simp_all
98ab21b14f09 add class ring_div; generalize mod/diff/minus proofs for class ring_div
huffman
parents: 29404
diff changeset
   412
98ab21b14f09 add class ring_div; generalize mod/diff/minus proofs for class ring_div
huffman
parents: 29404
diff changeset
   413
lemma mod_diff_right_eq: "(a - b) mod c = (a - b mod c) mod c"
98ab21b14f09 add class ring_div; generalize mod/diff/minus proofs for class ring_div
huffman
parents: 29404
diff changeset
   414
  unfolding diff_minus
98ab21b14f09 add class ring_div; generalize mod/diff/minus proofs for class ring_div
huffman
parents: 29404
diff changeset
   415
  by (intro mod_add_cong mod_minus_cong) simp_all
98ab21b14f09 add class ring_div; generalize mod/diff/minus proofs for class ring_div
huffman
parents: 29404
diff changeset
   416
98ab21b14f09 add class ring_div; generalize mod/diff/minus proofs for class ring_div
huffman
parents: 29404
diff changeset
   417
lemma mod_diff_eq: "(a - b) mod c = (a mod c - b mod c) mod c"
98ab21b14f09 add class ring_div; generalize mod/diff/minus proofs for class ring_div
huffman
parents: 29404
diff changeset
   418
  unfolding diff_minus
98ab21b14f09 add class ring_div; generalize mod/diff/minus proofs for class ring_div
huffman
parents: 29404
diff changeset
   419
  by (intro mod_add_cong mod_minus_cong) simp_all
98ab21b14f09 add class ring_div; generalize mod/diff/minus proofs for class ring_div
huffman
parents: 29404
diff changeset
   420
98ab21b14f09 add class ring_div; generalize mod/diff/minus proofs for class ring_div
huffman
parents: 29404
diff changeset
   421
lemma mod_diff_cong:
98ab21b14f09 add class ring_div; generalize mod/diff/minus proofs for class ring_div
huffman
parents: 29404
diff changeset
   422
  assumes "a mod c = a' mod c"
98ab21b14f09 add class ring_div; generalize mod/diff/minus proofs for class ring_div
huffman
parents: 29404
diff changeset
   423
  assumes "b mod c = b' mod c"
98ab21b14f09 add class ring_div; generalize mod/diff/minus proofs for class ring_div
huffman
parents: 29404
diff changeset
   424
  shows "(a - b) mod c = (a' - b') mod c"
98ab21b14f09 add class ring_div; generalize mod/diff/minus proofs for class ring_div
huffman
parents: 29404
diff changeset
   425
  unfolding diff_minus using assms
98ab21b14f09 add class ring_div; generalize mod/diff/minus proofs for class ring_div
huffman
parents: 29404
diff changeset
   426
  by (intro mod_add_cong mod_minus_cong)
98ab21b14f09 add class ring_div; generalize mod/diff/minus proofs for class ring_div
huffman
parents: 29404
diff changeset
   427
30180
6d29a873141f added lemmas by Jeremy Avigad
nipkow
parents: 30079
diff changeset
   428
lemma dvd_neg_div: "y dvd x \<Longrightarrow> -x div y = - (x div y)"
6d29a873141f added lemmas by Jeremy Avigad
nipkow
parents: 30079
diff changeset
   429
apply (case_tac "y = 0") apply simp
6d29a873141f added lemmas by Jeremy Avigad
nipkow
parents: 30079
diff changeset
   430
apply (auto simp add: dvd_def)
6d29a873141f added lemmas by Jeremy Avigad
nipkow
parents: 30079
diff changeset
   431
apply (subgoal_tac "-(y * k) = y * - k")
6d29a873141f added lemmas by Jeremy Avigad
nipkow
parents: 30079
diff changeset
   432
 apply (erule ssubst)
6d29a873141f added lemmas by Jeremy Avigad
nipkow
parents: 30079
diff changeset
   433
 apply (erule div_mult_self1_is_id)
6d29a873141f added lemmas by Jeremy Avigad
nipkow
parents: 30079
diff changeset
   434
apply simp
6d29a873141f added lemmas by Jeremy Avigad
nipkow
parents: 30079
diff changeset
   435
done
6d29a873141f added lemmas by Jeremy Avigad
nipkow
parents: 30079
diff changeset
   436
6d29a873141f added lemmas by Jeremy Avigad
nipkow
parents: 30079
diff changeset
   437
lemma dvd_div_neg: "y dvd x \<Longrightarrow> x div -y = - (x div y)"
6d29a873141f added lemmas by Jeremy Avigad
nipkow
parents: 30079
diff changeset
   438
apply (case_tac "y = 0") apply simp
6d29a873141f added lemmas by Jeremy Avigad
nipkow
parents: 30079
diff changeset
   439
apply (auto simp add: dvd_def)
6d29a873141f added lemmas by Jeremy Avigad
nipkow
parents: 30079
diff changeset
   440
apply (subgoal_tac "y * k = -y * -k")
6d29a873141f added lemmas by Jeremy Avigad
nipkow
parents: 30079
diff changeset
   441
 apply (erule ssubst)
6d29a873141f added lemmas by Jeremy Avigad
nipkow
parents: 30079
diff changeset
   442
 apply (rule div_mult_self1_is_id)
6d29a873141f added lemmas by Jeremy Avigad
nipkow
parents: 30079
diff changeset
   443
 apply simp
6d29a873141f added lemmas by Jeremy Avigad
nipkow
parents: 30079
diff changeset
   444
apply simp
6d29a873141f added lemmas by Jeremy Avigad
nipkow
parents: 30079
diff changeset
   445
done
6d29a873141f added lemmas by Jeremy Avigad
nipkow
parents: 30079
diff changeset
   446
29405
98ab21b14f09 add class ring_div; generalize mod/diff/minus proofs for class ring_div
huffman
parents: 29404
diff changeset
   447
end
98ab21b14f09 add class ring_div; generalize mod/diff/minus proofs for class ring_div
huffman
parents: 29404
diff changeset
   448
25942
a52309ac4a4d added class semiring_div
haftmann
parents: 25571
diff changeset
   449
26100
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   450
subsection {* Division on @{typ nat} *}
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   451
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   452
text {*
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   453
  We define @{const div} and @{const mod} on @{typ nat} by means
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   454
  of a characteristic relation with two input arguments
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   455
  @{term "m\<Colon>nat"}, @{term "n\<Colon>nat"} and two output arguments
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   456
  @{term "q\<Colon>nat"}(uotient) and @{term "r\<Colon>nat"}(emainder).
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   457
*}
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   458
33340
a165b97f3658 moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 33318
diff changeset
   459
definition divmod_nat_rel :: "nat \<Rightarrow> nat \<Rightarrow> nat \<times> nat \<Rightarrow> bool" where
a165b97f3658 moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 33318
diff changeset
   460
  "divmod_nat_rel m n qr \<longleftrightarrow>
30923
2697a1d1d34a more coherent developement in Divides.thy and IntDiv.thy
haftmann
parents: 30840
diff changeset
   461
    m = fst qr * n + snd qr \<and>
2697a1d1d34a more coherent developement in Divides.thy and IntDiv.thy
haftmann
parents: 30840
diff changeset
   462
      (if n = 0 then fst qr = 0 else if n > 0 then 0 \<le> snd qr \<and> snd qr < n else n < snd qr \<and> snd qr \<le> 0)"
26100
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   463
33340
a165b97f3658 moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 33318
diff changeset
   464
text {* @{const divmod_nat_rel} is total: *}
26100
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   465
33340
a165b97f3658 moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 33318
diff changeset
   466
lemma divmod_nat_rel_ex:
a165b97f3658 moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 33318
diff changeset
   467
  obtains q r where "divmod_nat_rel m n (q, r)"
26100
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   468
proof (cases "n = 0")
30923
2697a1d1d34a more coherent developement in Divides.thy and IntDiv.thy
haftmann
parents: 30840
diff changeset
   469
  case True  with that show thesis
33340
a165b97f3658 moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 33318
diff changeset
   470
    by (auto simp add: divmod_nat_rel_def)
26100
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   471
next
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   472
  case False
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   473
  have "\<exists>q r. m = q * n + r \<and> r < n"
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   474
  proof (induct m)
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   475
    case 0 with `n \<noteq> 0`
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   476
    have "(0\<Colon>nat) = 0 * n + 0 \<and> 0 < n" by simp
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   477
    then show ?case by blast
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   478
  next
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   479
    case (Suc m) then obtain q' r'
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   480
      where m: "m = q' * n + r'" and n: "r' < n" by auto
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   481
    then show ?case proof (cases "Suc r' < n")
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   482
      case True
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   483
      from m n have "Suc m = q' * n + Suc r'" by simp
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   484
      with True show ?thesis by blast
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   485
    next
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   486
      case False then have "n \<le> Suc r'" by auto
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   487
      moreover from n have "Suc r' \<le> n" by auto
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   488
      ultimately have "n = Suc r'" by auto
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   489
      with m have "Suc m = Suc q' * n + 0" by simp
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   490
      with `n \<noteq> 0` show ?thesis by blast
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   491
    qed
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   492
  qed
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   493
  with that show thesis
33340
a165b97f3658 moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 33318
diff changeset
   494
    using `n \<noteq> 0` by (auto simp add: divmod_nat_rel_def)
26100
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   495
qed
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   496
33340
a165b97f3658 moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 33318
diff changeset
   497
text {* @{const divmod_nat_rel} is injective: *}
26100
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   498
33340
a165b97f3658 moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 33318
diff changeset
   499
lemma divmod_nat_rel_unique:
a165b97f3658 moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 33318
diff changeset
   500
  assumes "divmod_nat_rel m n qr"
a165b97f3658 moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 33318
diff changeset
   501
    and "divmod_nat_rel m n qr'"
30923
2697a1d1d34a more coherent developement in Divides.thy and IntDiv.thy
haftmann
parents: 30840
diff changeset
   502
  shows "qr = qr'"
26100
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   503
proof (cases "n = 0")
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   504
  case True with assms show ?thesis
30923
2697a1d1d34a more coherent developement in Divides.thy and IntDiv.thy
haftmann
parents: 30840
diff changeset
   505
    by (cases qr, cases qr')
33340
a165b97f3658 moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 33318
diff changeset
   506
      (simp add: divmod_nat_rel_def)
26100
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   507
next
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   508
  case False
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   509
  have aux: "\<And>q r q' r'. q' * n + r' = q * n + r \<Longrightarrow> r < n \<Longrightarrow> q' \<le> (q\<Colon>nat)"
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   510
  apply (rule leI)
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   511
  apply (subst less_iff_Suc_add)
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   512
  apply (auto simp add: add_mult_distrib)
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   513
  done
30923
2697a1d1d34a more coherent developement in Divides.thy and IntDiv.thy
haftmann
parents: 30840
diff changeset
   514
  from `n \<noteq> 0` assms have "fst qr = fst qr'"
33340
a165b97f3658 moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 33318
diff changeset
   515
    by (auto simp add: divmod_nat_rel_def intro: order_antisym dest: aux sym)
30923
2697a1d1d34a more coherent developement in Divides.thy and IntDiv.thy
haftmann
parents: 30840
diff changeset
   516
  moreover from this assms have "snd qr = snd qr'"
33340
a165b97f3658 moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 33318
diff changeset
   517
    by (simp add: divmod_nat_rel_def)
30923
2697a1d1d34a more coherent developement in Divides.thy and IntDiv.thy
haftmann
parents: 30840
diff changeset
   518
  ultimately show ?thesis by (cases qr, cases qr') simp
26100
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   519
qed
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   520
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   521
text {*
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   522
  We instantiate divisibility on the natural numbers by
33340
a165b97f3658 moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 33318
diff changeset
   523
  means of @{const divmod_nat_rel}:
26100
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   524
*}
25942
a52309ac4a4d added class semiring_div
haftmann
parents: 25571
diff changeset
   525
33340
a165b97f3658 moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 33318
diff changeset
   526
definition divmod_nat :: "nat \<Rightarrow> nat \<Rightarrow> nat \<times> nat" where
37767
a2b7a20d6ea3 dropped superfluous [code del]s
haftmann
parents: 36634
diff changeset
   527
  "divmod_nat m n = (THE qr. divmod_nat_rel m n qr)"
30923
2697a1d1d34a more coherent developement in Divides.thy and IntDiv.thy
haftmann
parents: 30840
diff changeset
   528
33340
a165b97f3658 moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 33318
diff changeset
   529
lemma divmod_nat_rel_divmod_nat:
a165b97f3658 moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 33318
diff changeset
   530
  "divmod_nat_rel m n (divmod_nat m n)"
30923
2697a1d1d34a more coherent developement in Divides.thy and IntDiv.thy
haftmann
parents: 30840
diff changeset
   531
proof -
33340
a165b97f3658 moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 33318
diff changeset
   532
  from divmod_nat_rel_ex
a165b97f3658 moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 33318
diff changeset
   533
    obtain qr where rel: "divmod_nat_rel m n qr" .
30923
2697a1d1d34a more coherent developement in Divides.thy and IntDiv.thy
haftmann
parents: 30840
diff changeset
   534
  then show ?thesis
33340
a165b97f3658 moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 33318
diff changeset
   535
  by (auto simp add: divmod_nat_def intro: theI elim: divmod_nat_rel_unique)
30923
2697a1d1d34a more coherent developement in Divides.thy and IntDiv.thy
haftmann
parents: 30840
diff changeset
   536
qed
2697a1d1d34a more coherent developement in Divides.thy and IntDiv.thy
haftmann
parents: 30840
diff changeset
   537
33340
a165b97f3658 moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 33318
diff changeset
   538
lemma divmod_nat_eq:
a165b97f3658 moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 33318
diff changeset
   539
  assumes "divmod_nat_rel m n qr" 
a165b97f3658 moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 33318
diff changeset
   540
  shows "divmod_nat m n = qr"
a165b97f3658 moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 33318
diff changeset
   541
  using assms by (auto intro: divmod_nat_rel_unique divmod_nat_rel_divmod_nat)
26100
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   542
46551
866bce5442a3 simplify projections on simultaneous computations of div and mod; tuned structure (from Florian Haftmann)
huffman
parents: 46026
diff changeset
   543
instantiation nat :: semiring_div
866bce5442a3 simplify projections on simultaneous computations of div and mod; tuned structure (from Florian Haftmann)
huffman
parents: 46026
diff changeset
   544
begin
866bce5442a3 simplify projections on simultaneous computations of div and mod; tuned structure (from Florian Haftmann)
huffman
parents: 46026
diff changeset
   545
26100
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   546
definition div_nat where
33340
a165b97f3658 moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 33318
diff changeset
   547
  "m div n = fst (divmod_nat m n)"
26100
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   548
46551
866bce5442a3 simplify projections on simultaneous computations of div and mod; tuned structure (from Florian Haftmann)
huffman
parents: 46026
diff changeset
   549
lemma fst_divmod_nat [simp]:
866bce5442a3 simplify projections on simultaneous computations of div and mod; tuned structure (from Florian Haftmann)
huffman
parents: 46026
diff changeset
   550
  "fst (divmod_nat m n) = m div n"
866bce5442a3 simplify projections on simultaneous computations of div and mod; tuned structure (from Florian Haftmann)
huffman
parents: 46026
diff changeset
   551
  by (simp add: div_nat_def)
866bce5442a3 simplify projections on simultaneous computations of div and mod; tuned structure (from Florian Haftmann)
huffman
parents: 46026
diff changeset
   552
26100
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   553
definition mod_nat where
33340
a165b97f3658 moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 33318
diff changeset
   554
  "m mod n = snd (divmod_nat m n)"
25571
c9e39eafc7a0 instantiation target rather than legacy instance
haftmann
parents: 25162
diff changeset
   555
46551
866bce5442a3 simplify projections on simultaneous computations of div and mod; tuned structure (from Florian Haftmann)
huffman
parents: 46026
diff changeset
   556
lemma snd_divmod_nat [simp]:
866bce5442a3 simplify projections on simultaneous computations of div and mod; tuned structure (from Florian Haftmann)
huffman
parents: 46026
diff changeset
   557
  "snd (divmod_nat m n) = m mod n"
866bce5442a3 simplify projections on simultaneous computations of div and mod; tuned structure (from Florian Haftmann)
huffman
parents: 46026
diff changeset
   558
  by (simp add: mod_nat_def)
866bce5442a3 simplify projections on simultaneous computations of div and mod; tuned structure (from Florian Haftmann)
huffman
parents: 46026
diff changeset
   559
33340
a165b97f3658 moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 33318
diff changeset
   560
lemma divmod_nat_div_mod:
a165b97f3658 moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 33318
diff changeset
   561
  "divmod_nat m n = (m div n, m mod n)"
46551
866bce5442a3 simplify projections on simultaneous computations of div and mod; tuned structure (from Florian Haftmann)
huffman
parents: 46026
diff changeset
   562
  by (simp add: prod_eq_iff)
26100
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   563
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   564
lemma div_eq:
33340
a165b97f3658 moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 33318
diff changeset
   565
  assumes "divmod_nat_rel m n (q, r)" 
26100
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   566
  shows "m div n = q"
46551
866bce5442a3 simplify projections on simultaneous computations of div and mod; tuned structure (from Florian Haftmann)
huffman
parents: 46026
diff changeset
   567
  using assms by (auto dest!: divmod_nat_eq simp add: prod_eq_iff)
26100
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   568
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   569
lemma mod_eq:
33340
a165b97f3658 moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 33318
diff changeset
   570
  assumes "divmod_nat_rel m n (q, r)" 
26100
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   571
  shows "m mod n = r"
46551
866bce5442a3 simplify projections on simultaneous computations of div and mod; tuned structure (from Florian Haftmann)
huffman
parents: 46026
diff changeset
   572
  using assms by (auto dest!: divmod_nat_eq simp add: prod_eq_iff)
25571
c9e39eafc7a0 instantiation target rather than legacy instance
haftmann
parents: 25162
diff changeset
   573
33340
a165b97f3658 moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 33318
diff changeset
   574
lemma divmod_nat_rel: "divmod_nat_rel m n (m div n, m mod n)"
46551
866bce5442a3 simplify projections on simultaneous computations of div and mod; tuned structure (from Florian Haftmann)
huffman
parents: 46026
diff changeset
   575
  using divmod_nat_rel_divmod_nat by (simp add: divmod_nat_div_mod)
14267
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14208
diff changeset
   576
33340
a165b97f3658 moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 33318
diff changeset
   577
lemma divmod_nat_zero:
a165b97f3658 moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 33318
diff changeset
   578
  "divmod_nat m 0 = (0, m)"
26100
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   579
proof -
33340
a165b97f3658 moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 33318
diff changeset
   580
  from divmod_nat_rel [of m 0] show ?thesis
a165b97f3658 moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 33318
diff changeset
   581
    unfolding divmod_nat_div_mod divmod_nat_rel_def by simp
26100
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   582
qed
25942
a52309ac4a4d added class semiring_div
haftmann
parents: 25571
diff changeset
   583
33340
a165b97f3658 moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 33318
diff changeset
   584
lemma divmod_nat_base:
26100
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   585
  assumes "m < n"
33340
a165b97f3658 moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 33318
diff changeset
   586
  shows "divmod_nat m n = (0, m)"
26100
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   587
proof -
33340
a165b97f3658 moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 33318
diff changeset
   588
  from divmod_nat_rel [of m n] show ?thesis
a165b97f3658 moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 33318
diff changeset
   589
    unfolding divmod_nat_div_mod divmod_nat_rel_def
26100
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   590
    using assms by (cases "m div n = 0")
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   591
      (auto simp add: gr0_conv_Suc [of "m div n"])
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   592
qed
25942
a52309ac4a4d added class semiring_div
haftmann
parents: 25571
diff changeset
   593
33340
a165b97f3658 moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 33318
diff changeset
   594
lemma divmod_nat_step:
26100
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   595
  assumes "0 < n" and "n \<le> m"
33340
a165b97f3658 moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 33318
diff changeset
   596
  shows "divmod_nat m n = (Suc ((m - n) div n), (m - n) mod n)"
26100
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   597
proof -
33340
a165b97f3658 moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 33318
diff changeset
   598
  from divmod_nat_rel have divmod_nat_m_n: "divmod_nat_rel m n (m div n, m mod n)" .
26100
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   599
  with assms have m_div_n: "m div n \<ge> 1"
33340
a165b97f3658 moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 33318
diff changeset
   600
    by (cases "m div n") (auto simp add: divmod_nat_rel_def)
35815
10e723e54076 tuned proofs (to avoid linarith error message caused by bootstrapping of HOL)
boehmes
parents: 35673
diff changeset
   601
  have "divmod_nat_rel (m - n) n (m div n - Suc 0, m mod n)"
10e723e54076 tuned proofs (to avoid linarith error message caused by bootstrapping of HOL)
boehmes
parents: 35673
diff changeset
   602
  proof -
10e723e54076 tuned proofs (to avoid linarith error message caused by bootstrapping of HOL)
boehmes
parents: 35673
diff changeset
   603
    from assms have
10e723e54076 tuned proofs (to avoid linarith error message caused by bootstrapping of HOL)
boehmes
parents: 35673
diff changeset
   604
      "n \<noteq> 0"
10e723e54076 tuned proofs (to avoid linarith error message caused by bootstrapping of HOL)
boehmes
parents: 35673
diff changeset
   605
      "\<And>k. m = Suc k * n + m mod n ==> m - n = (Suc k - Suc 0) * n + m mod n"
10e723e54076 tuned proofs (to avoid linarith error message caused by bootstrapping of HOL)
boehmes
parents: 35673
diff changeset
   606
      by simp_all
10e723e54076 tuned proofs (to avoid linarith error message caused by bootstrapping of HOL)
boehmes
parents: 35673
diff changeset
   607
    then show ?thesis using assms divmod_nat_m_n 
10e723e54076 tuned proofs (to avoid linarith error message caused by bootstrapping of HOL)
boehmes
parents: 35673
diff changeset
   608
      by (cases "m div n")
10e723e54076 tuned proofs (to avoid linarith error message caused by bootstrapping of HOL)
boehmes
parents: 35673
diff changeset
   609
         (simp_all only: divmod_nat_rel_def fst_conv snd_conv, simp_all)
10e723e54076 tuned proofs (to avoid linarith error message caused by bootstrapping of HOL)
boehmes
parents: 35673
diff changeset
   610
  qed
33340
a165b97f3658 moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 33318
diff changeset
   611
  with divmod_nat_eq have "divmod_nat (m - n) n = (m div n - Suc 0, m mod n)" by simp
a165b97f3658 moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 33318
diff changeset
   612
  moreover from divmod_nat_div_mod have "divmod_nat (m - n) n = ((m - n) div n, (m - n) mod n)" .
26100
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   613
  ultimately have "m div n = Suc ((m - n) div n)"
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   614
    and "m mod n = (m - n) mod n" using m_div_n by simp_all
33340
a165b97f3658 moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 33318
diff changeset
   615
  then show ?thesis using divmod_nat_div_mod by simp
26100
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   616
qed
25942
a52309ac4a4d added class semiring_div
haftmann
parents: 25571
diff changeset
   617
26300
03def556e26e removed duplicate lemmas;
wenzelm
parents: 26100
diff changeset
   618
text {* The ''recursion'' equations for @{const div} and @{const mod} *}
26100
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   619
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   620
lemma div_less [simp]:
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   621
  fixes m n :: nat
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   622
  assumes "m < n"
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   623
  shows "m div n = 0"
46551
866bce5442a3 simplify projections on simultaneous computations of div and mod; tuned structure (from Florian Haftmann)
huffman
parents: 46026
diff changeset
   624
  using assms divmod_nat_base by (simp add: prod_eq_iff)
25942
a52309ac4a4d added class semiring_div
haftmann
parents: 25571
diff changeset
   625
26100
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   626
lemma le_div_geq:
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   627
  fixes m n :: nat
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   628
  assumes "0 < n" and "n \<le> m"
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   629
  shows "m div n = Suc ((m - n) div n)"
46551
866bce5442a3 simplify projections on simultaneous computations of div and mod; tuned structure (from Florian Haftmann)
huffman
parents: 46026
diff changeset
   630
  using assms divmod_nat_step by (simp add: prod_eq_iff)
14267
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14208
diff changeset
   631
26100
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   632
lemma mod_less [simp]:
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   633
  fixes m n :: nat
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   634
  assumes "m < n"
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   635
  shows "m mod n = m"
46551
866bce5442a3 simplify projections on simultaneous computations of div and mod; tuned structure (from Florian Haftmann)
huffman
parents: 46026
diff changeset
   636
  using assms divmod_nat_base by (simp add: prod_eq_iff)
26100
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   637
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   638
lemma le_mod_geq:
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   639
  fixes m n :: nat
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   640
  assumes "n \<le> m"
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   641
  shows "m mod n = (m - n) mod n"
46551
866bce5442a3 simplify projections on simultaneous computations of div and mod; tuned structure (from Florian Haftmann)
huffman
parents: 46026
diff changeset
   642
  using assms divmod_nat_step by (cases "n = 0") (simp_all add: prod_eq_iff)
14267
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14208
diff changeset
   643
30930
11010e5f18f0 tightended specification of class semiring_div
haftmann
parents: 30923
diff changeset
   644
instance proof -
11010e5f18f0 tightended specification of class semiring_div
haftmann
parents: 30923
diff changeset
   645
  have [simp]: "\<And>n::nat. n div 0 = 0"
33340
a165b97f3658 moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 33318
diff changeset
   646
    by (simp add: div_nat_def divmod_nat_zero)
30930
11010e5f18f0 tightended specification of class semiring_div
haftmann
parents: 30923
diff changeset
   647
  have [simp]: "\<And>n::nat. 0 div n = 0"
11010e5f18f0 tightended specification of class semiring_div
haftmann
parents: 30923
diff changeset
   648
  proof -
11010e5f18f0 tightended specification of class semiring_div
haftmann
parents: 30923
diff changeset
   649
    fix n :: nat
11010e5f18f0 tightended specification of class semiring_div
haftmann
parents: 30923
diff changeset
   650
    show "0 div n = 0"
11010e5f18f0 tightended specification of class semiring_div
haftmann
parents: 30923
diff changeset
   651
      by (cases "n = 0") simp_all
11010e5f18f0 tightended specification of class semiring_div
haftmann
parents: 30923
diff changeset
   652
  qed
11010e5f18f0 tightended specification of class semiring_div
haftmann
parents: 30923
diff changeset
   653
  show "OFCLASS(nat, semiring_div_class)" proof
11010e5f18f0 tightended specification of class semiring_div
haftmann
parents: 30923
diff changeset
   654
    fix m n :: nat
11010e5f18f0 tightended specification of class semiring_div
haftmann
parents: 30923
diff changeset
   655
    show "m div n * n + m mod n = m"
33340
a165b97f3658 moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 33318
diff changeset
   656
      using divmod_nat_rel [of m n] by (simp add: divmod_nat_rel_def)
30930
11010e5f18f0 tightended specification of class semiring_div
haftmann
parents: 30923
diff changeset
   657
  next
11010e5f18f0 tightended specification of class semiring_div
haftmann
parents: 30923
diff changeset
   658
    fix m n q :: nat
11010e5f18f0 tightended specification of class semiring_div
haftmann
parents: 30923
diff changeset
   659
    assume "n \<noteq> 0"
11010e5f18f0 tightended specification of class semiring_div
haftmann
parents: 30923
diff changeset
   660
    then show "(q + m * n) div n = m + q div n"
11010e5f18f0 tightended specification of class semiring_div
haftmann
parents: 30923
diff changeset
   661
      by (induct m) (simp_all add: le_div_geq)
11010e5f18f0 tightended specification of class semiring_div
haftmann
parents: 30923
diff changeset
   662
  next
11010e5f18f0 tightended specification of class semiring_div
haftmann
parents: 30923
diff changeset
   663
    fix m n q :: nat
11010e5f18f0 tightended specification of class semiring_div
haftmann
parents: 30923
diff changeset
   664
    assume "m \<noteq> 0"
11010e5f18f0 tightended specification of class semiring_div
haftmann
parents: 30923
diff changeset
   665
    then show "(m * n) div (m * q) = n div q"
11010e5f18f0 tightended specification of class semiring_div
haftmann
parents: 30923
diff changeset
   666
    proof (cases "n \<noteq> 0 \<and> q \<noteq> 0")
11010e5f18f0 tightended specification of class semiring_div
haftmann
parents: 30923
diff changeset
   667
      case False then show ?thesis by auto
11010e5f18f0 tightended specification of class semiring_div
haftmann
parents: 30923
diff changeset
   668
    next
11010e5f18f0 tightended specification of class semiring_div
haftmann
parents: 30923
diff changeset
   669
      case True with `m \<noteq> 0`
11010e5f18f0 tightended specification of class semiring_div
haftmann
parents: 30923
diff changeset
   670
        have "m > 0" and "n > 0" and "q > 0" by auto
33340
a165b97f3658 moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 33318
diff changeset
   671
      then have "\<And>a b. divmod_nat_rel n q (a, b) \<Longrightarrow> divmod_nat_rel (m * n) (m * q) (a, m * b)"
a165b97f3658 moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 33318
diff changeset
   672
        by (auto simp add: divmod_nat_rel_def) (simp_all add: algebra_simps)
a165b97f3658 moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 33318
diff changeset
   673
      moreover from divmod_nat_rel have "divmod_nat_rel n q (n div q, n mod q)" .
a165b97f3658 moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 33318
diff changeset
   674
      ultimately have "divmod_nat_rel (m * n) (m * q) (n div q, m * (n mod q))" .
30930
11010e5f18f0 tightended specification of class semiring_div
haftmann
parents: 30923
diff changeset
   675
      then show ?thesis by (simp add: div_eq)
11010e5f18f0 tightended specification of class semiring_div
haftmann
parents: 30923
diff changeset
   676
    qed
11010e5f18f0 tightended specification of class semiring_div
haftmann
parents: 30923
diff changeset
   677
  qed simp_all
25942
a52309ac4a4d added class semiring_div
haftmann
parents: 25571
diff changeset
   678
qed
26100
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   679
25942
a52309ac4a4d added class semiring_div
haftmann
parents: 25571
diff changeset
   680
end
14267
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14208
diff changeset
   681
33361
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
   682
lemma divmod_nat_if [code]: "divmod_nat m n = (if n = 0 \<or> m < n then (0, m) else
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
   683
  let (q, r) = divmod_nat (m - n) n in (Suc q, r))"
46551
866bce5442a3 simplify projections on simultaneous computations of div and mod; tuned structure (from Florian Haftmann)
huffman
parents: 46026
diff changeset
   684
  by (simp add: prod_eq_iff prod_case_beta not_less le_div_geq le_mod_geq)
33361
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
   685
26100
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   686
text {* Simproc for cancelling @{const div} and @{const mod} *}
25942
a52309ac4a4d added class semiring_div
haftmann
parents: 25571
diff changeset
   687
30934
ed5377c2b0a3 tuned setups of CancelDivMod
haftmann
parents: 30930
diff changeset
   688
ML {*
43594
ef1ddc59b825 modernized some simproc setup;
wenzelm
parents: 41792
diff changeset
   689
structure Cancel_Div_Mod_Nat = Cancel_Div_Mod
41550
efa734d9b221 eliminated global prems;
wenzelm
parents: 39489
diff changeset
   690
(
30934
ed5377c2b0a3 tuned setups of CancelDivMod
haftmann
parents: 30930
diff changeset
   691
  val div_name = @{const_name div};
ed5377c2b0a3 tuned setups of CancelDivMod
haftmann
parents: 30930
diff changeset
   692
  val mod_name = @{const_name mod};
ed5377c2b0a3 tuned setups of CancelDivMod
haftmann
parents: 30930
diff changeset
   693
  val mk_binop = HOLogic.mk_binop;
ed5377c2b0a3 tuned setups of CancelDivMod
haftmann
parents: 30930
diff changeset
   694
  val mk_sum = Nat_Arith.mk_sum;
ed5377c2b0a3 tuned setups of CancelDivMod
haftmann
parents: 30930
diff changeset
   695
  val dest_sum = Nat_Arith.dest_sum;
25942
a52309ac4a4d added class semiring_div
haftmann
parents: 25571
diff changeset
   696
30934
ed5377c2b0a3 tuned setups of CancelDivMod
haftmann
parents: 30930
diff changeset
   697
  val div_mod_eqs = map mk_meta_eq [@{thm div_mod_equality}, @{thm div_mod_equality2}];
14267
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14208
diff changeset
   698
30934
ed5377c2b0a3 tuned setups of CancelDivMod
haftmann
parents: 30930
diff changeset
   699
  val prove_eq_sums = Arith_Data.prove_conv2 all_tac (Arith_Data.simp_all_tac
35050
9f841f20dca6 renamed OrderedGroup to Groups; split theory Ring_and_Field into Rings Fields
haftmann
parents: 34982
diff changeset
   700
    (@{thm add_0_left} :: @{thm add_0_right} :: @{thms add_ac}))
41550
efa734d9b221 eliminated global prems;
wenzelm
parents: 39489
diff changeset
   701
)
25942
a52309ac4a4d added class semiring_div
haftmann
parents: 25571
diff changeset
   702
*}
a52309ac4a4d added class semiring_div
haftmann
parents: 25571
diff changeset
   703
43594
ef1ddc59b825 modernized some simproc setup;
wenzelm
parents: 41792
diff changeset
   704
simproc_setup cancel_div_mod_nat ("(m::nat) + n") = {* K Cancel_Div_Mod_Nat.proc *}
ef1ddc59b825 modernized some simproc setup;
wenzelm
parents: 41792
diff changeset
   705
26100
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   706
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   707
subsubsection {* Quotient *}
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   708
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   709
lemma div_geq: "0 < n \<Longrightarrow>  \<not> m < n \<Longrightarrow> m div n = Suc ((m - n) div n)"
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29509
diff changeset
   710
by (simp add: le_div_geq linorder_not_less)
26100
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   711
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   712
lemma div_if: "0 < n \<Longrightarrow> m div n = (if m < n then 0 else Suc ((m - n) div n))"
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29509
diff changeset
   713
by (simp add: div_geq)
26100
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   714
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   715
lemma div_mult_self_is_m [simp]: "0<n ==> (m*n) div n = (m::nat)"
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29509
diff changeset
   716
by simp
26100
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   717
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   718
lemma div_mult_self1_is_m [simp]: "0<n ==> (n*m) div n = (m::nat)"
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29509
diff changeset
   719
by simp
26100
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   720
25942
a52309ac4a4d added class semiring_div
haftmann
parents: 25571
diff changeset
   721
a52309ac4a4d added class semiring_div
haftmann
parents: 25571
diff changeset
   722
subsubsection {* Remainder *}
a52309ac4a4d added class semiring_div
haftmann
parents: 25571
diff changeset
   723
26100
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   724
lemma mod_less_divisor [simp]:
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   725
  fixes m n :: nat
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   726
  assumes "n > 0"
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   727
  shows "m mod n < (n::nat)"
33340
a165b97f3658 moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 33318
diff changeset
   728
  using assms divmod_nat_rel [of m n] unfolding divmod_nat_rel_def by auto
14267
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14208
diff changeset
   729
26100
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   730
lemma mod_less_eq_dividend [simp]:
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   731
  fixes m n :: nat
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   732
  shows "m mod n \<le> m"
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   733
proof (rule add_leD2)
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   734
  from mod_div_equality have "m div n * n + m mod n = m" .
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   735
  then show "m div n * n + m mod n \<le> m" by auto
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   736
qed
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   737
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   738
lemma mod_geq: "\<not> m < (n\<Colon>nat) \<Longrightarrow> m mod n = (m - n) mod n"
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29509
diff changeset
   739
by (simp add: le_mod_geq linorder_not_less)
14267
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14208
diff changeset
   740
26100
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   741
lemma mod_if: "m mod (n\<Colon>nat) = (if m < n then m else (m - n) mod n)"
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29509
diff changeset
   742
by (simp add: le_mod_geq)
26100
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   743
14267
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14208
diff changeset
   744
lemma mod_1 [simp]: "m mod Suc 0 = 0"
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29509
diff changeset
   745
by (induct m) (simp_all add: mod_geq)
14267
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14208
diff changeset
   746
26100
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   747
lemma mod_mult_distrib: "(m mod n) * (k\<Colon>nat) = (m * k) mod (n * k)"
22718
936f7580937d tuned proofs;
wenzelm
parents: 22473
diff changeset
   748
  apply (cases "n = 0", simp)
936f7580937d tuned proofs;
wenzelm
parents: 22473
diff changeset
   749
  apply (cases "k = 0", simp)
936f7580937d tuned proofs;
wenzelm
parents: 22473
diff changeset
   750
  apply (induct m rule: nat_less_induct)
936f7580937d tuned proofs;
wenzelm
parents: 22473
diff changeset
   751
  apply (subst mod_if, simp)
936f7580937d tuned proofs;
wenzelm
parents: 22473
diff changeset
   752
  apply (simp add: mod_geq diff_mult_distrib)
936f7580937d tuned proofs;
wenzelm
parents: 22473
diff changeset
   753
  done
14267
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14208
diff changeset
   754
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14208
diff changeset
   755
lemma mod_mult_distrib2: "(k::nat) * (m mod n) = (k*m) mod (k*n)"
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29509
diff changeset
   756
by (simp add: mult_commute [of k] mod_mult_distrib)
14267
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14208
diff changeset
   757
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14208
diff changeset
   758
(* a simple rearrangement of mod_div_equality: *)
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14208
diff changeset
   759
lemma mult_div_cancel: "(n::nat) * (m div n) = m - (m mod n)"
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29509
diff changeset
   760
by (cut_tac a = m and b = n in mod_div_equality2, arith)
14267
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14208
diff changeset
   761
15439
71c0f98e31f1 made diff_less a simp rule
nipkow
parents: 15251
diff changeset
   762
lemma mod_le_divisor[simp]: "0 < n \<Longrightarrow> m mod n \<le> (n::nat)"
22718
936f7580937d tuned proofs;
wenzelm
parents: 22473
diff changeset
   763
  apply (drule mod_less_divisor [where m = m])
936f7580937d tuned proofs;
wenzelm
parents: 22473
diff changeset
   764
  apply simp
936f7580937d tuned proofs;
wenzelm
parents: 22473
diff changeset
   765
  done
14267
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14208
diff changeset
   766
26100
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   767
subsubsection {* Quotient and Remainder *}
14267
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14208
diff changeset
   768
33340
a165b97f3658 moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 33318
diff changeset
   769
lemma divmod_nat_rel_mult1_eq:
a165b97f3658 moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 33318
diff changeset
   770
  "divmod_nat_rel b c (q, r) \<Longrightarrow> c > 0
a165b97f3658 moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 33318
diff changeset
   771
   \<Longrightarrow> divmod_nat_rel (a * b) c (a * q + a * r div c, a * r mod c)"
a165b97f3658 moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 33318
diff changeset
   772
by (auto simp add: split_ifs divmod_nat_rel_def algebra_simps)
14267
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14208
diff changeset
   773
30923
2697a1d1d34a more coherent developement in Divides.thy and IntDiv.thy
haftmann
parents: 30840
diff changeset
   774
lemma div_mult1_eq:
2697a1d1d34a more coherent developement in Divides.thy and IntDiv.thy
haftmann
parents: 30840
diff changeset
   775
  "(a * b) div c = a * (b div c) + a * (b mod c) div (c::nat)"
25134
3d4953e88449 Eliminated most of the neq0_conv occurrences. As a result, many
nipkow
parents: 25112
diff changeset
   776
apply (cases "c = 0", simp)
33340
a165b97f3658 moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 33318
diff changeset
   777
apply (blast intro: divmod_nat_rel [THEN divmod_nat_rel_mult1_eq, THEN div_eq])
25134
3d4953e88449 Eliminated most of the neq0_conv occurrences. As a result, many
nipkow
parents: 25112
diff changeset
   778
done
14267
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14208
diff changeset
   779
33340
a165b97f3658 moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 33318
diff changeset
   780
lemma divmod_nat_rel_add1_eq:
a165b97f3658 moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 33318
diff changeset
   781
  "divmod_nat_rel a c (aq, ar) \<Longrightarrow> divmod_nat_rel b c (bq, br) \<Longrightarrow>  c > 0
a165b97f3658 moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 33318
diff changeset
   782
   \<Longrightarrow> divmod_nat_rel (a + b) c (aq + bq + (ar + br) div c, (ar + br) mod c)"
a165b97f3658 moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 33318
diff changeset
   783
by (auto simp add: split_ifs divmod_nat_rel_def algebra_simps)
14267
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14208
diff changeset
   784
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14208
diff changeset
   785
(*NOT suitable for rewriting: the RHS has an instance of the LHS*)
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14208
diff changeset
   786
lemma div_add1_eq:
25134
3d4953e88449 Eliminated most of the neq0_conv occurrences. As a result, many
nipkow
parents: 25112
diff changeset
   787
  "(a+b) div (c::nat) = a div c + b div c + ((a mod c + b mod c) div c)"
3d4953e88449 Eliminated most of the neq0_conv occurrences. As a result, many
nipkow
parents: 25112
diff changeset
   788
apply (cases "c = 0", simp)
33340
a165b97f3658 moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 33318
diff changeset
   789
apply (blast intro: divmod_nat_rel_add1_eq [THEN div_eq] divmod_nat_rel)
25134
3d4953e88449 Eliminated most of the neq0_conv occurrences. As a result, many
nipkow
parents: 25112
diff changeset
   790
done
14267
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14208
diff changeset
   791
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14208
diff changeset
   792
lemma mod_lemma: "[| (0::nat) < c; r < b |] ==> b * (q mod c) + r < b * c"
22718
936f7580937d tuned proofs;
wenzelm
parents: 22473
diff changeset
   793
  apply (cut_tac m = q and n = c in mod_less_divisor)
936f7580937d tuned proofs;
wenzelm
parents: 22473
diff changeset
   794
  apply (drule_tac [2] m = "q mod c" in less_imp_Suc_add, auto)
936f7580937d tuned proofs;
wenzelm
parents: 22473
diff changeset
   795
  apply (erule_tac P = "%x. ?lhs < ?rhs x" in ssubst)
936f7580937d tuned proofs;
wenzelm
parents: 22473
diff changeset
   796
  apply (simp add: add_mult_distrib2)
936f7580937d tuned proofs;
wenzelm
parents: 22473
diff changeset
   797
  done
10559
d3fd54fc659b many new div and mod properties (borrowed from Integ/IntDiv)
paulson
parents: 10214
diff changeset
   798
33340
a165b97f3658 moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 33318
diff changeset
   799
lemma divmod_nat_rel_mult2_eq:
a165b97f3658 moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 33318
diff changeset
   800
  "divmod_nat_rel a b (q, r) \<Longrightarrow> 0 < b \<Longrightarrow> 0 < c
a165b97f3658 moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 33318
diff changeset
   801
   \<Longrightarrow> divmod_nat_rel a (b * c) (q div c, b *(q mod c) + r)"
a165b97f3658 moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 33318
diff changeset
   802
by (auto simp add: mult_ac divmod_nat_rel_def add_mult_distrib2 [symmetric] mod_lemma)
14267
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14208
diff changeset
   803
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14208
diff changeset
   804
lemma div_mult2_eq: "a div (b*c) = (a div b) div (c::nat)"
22718
936f7580937d tuned proofs;
wenzelm
parents: 22473
diff changeset
   805
  apply (cases "b = 0", simp)
936f7580937d tuned proofs;
wenzelm
parents: 22473
diff changeset
   806
  apply (cases "c = 0", simp)
33340
a165b97f3658 moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 33318
diff changeset
   807
  apply (force simp add: divmod_nat_rel [THEN divmod_nat_rel_mult2_eq, THEN div_eq])
22718
936f7580937d tuned proofs;
wenzelm
parents: 22473
diff changeset
   808
  done
14267
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14208
diff changeset
   809
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14208
diff changeset
   810
lemma mod_mult2_eq: "a mod (b*c) = b*(a div b mod c) + a mod (b::nat)"
22718
936f7580937d tuned proofs;
wenzelm
parents: 22473
diff changeset
   811
  apply (cases "b = 0", simp)
936f7580937d tuned proofs;
wenzelm
parents: 22473
diff changeset
   812
  apply (cases "c = 0", simp)
33340
a165b97f3658 moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 33318
diff changeset
   813
  apply (auto simp add: mult_commute divmod_nat_rel [THEN divmod_nat_rel_mult2_eq, THEN mod_eq])
22718
936f7580937d tuned proofs;
wenzelm
parents: 22473
diff changeset
   814
  done
14267
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14208
diff changeset
   815
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14208
diff changeset
   816
46551
866bce5442a3 simplify projections on simultaneous computations of div and mod; tuned structure (from Florian Haftmann)
huffman
parents: 46026
diff changeset
   817
subsubsection {* Further Facts about Quotient and Remainder *}
14267
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14208
diff changeset
   818
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14208
diff changeset
   819
lemma div_1 [simp]: "m div Suc 0 = m"
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29509
diff changeset
   820
by (induct m) (simp_all add: div_geq)
14267
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14208
diff changeset
   821
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14208
diff changeset
   822
(* Monotonicity of div in first argument *)
30923
2697a1d1d34a more coherent developement in Divides.thy and IntDiv.thy
haftmann
parents: 30840
diff changeset
   823
lemma div_le_mono [rule_format (no_asm)]:
22718
936f7580937d tuned proofs;
wenzelm
parents: 22473
diff changeset
   824
    "\<forall>m::nat. m \<le> n --> (m div k) \<le> (n div k)"
14267
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14208
diff changeset
   825
apply (case_tac "k=0", simp)
15251
bb6f072c8d10 converted some induct_tac to induct
paulson
parents: 15140
diff changeset
   826
apply (induct "n" rule: nat_less_induct, clarify)
14267
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14208
diff changeset
   827
apply (case_tac "n<k")
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14208
diff changeset
   828
(* 1  case n<k *)
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14208
diff changeset
   829
apply simp
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14208
diff changeset
   830
(* 2  case n >= k *)
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14208
diff changeset
   831
apply (case_tac "m<k")
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14208
diff changeset
   832
(* 2.1  case m<k *)
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14208
diff changeset
   833
apply simp
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14208
diff changeset
   834
(* 2.2  case m>=k *)
15439
71c0f98e31f1 made diff_less a simp rule
nipkow
parents: 15251
diff changeset
   835
apply (simp add: div_geq diff_le_mono)
14267
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14208
diff changeset
   836
done
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14208
diff changeset
   837
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14208
diff changeset
   838
(* Antimonotonicity of div in second argument *)
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14208
diff changeset
   839
lemma div_le_mono2: "!!m::nat. [| 0<m; m\<le>n |] ==> (k div n) \<le> (k div m)"
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14208
diff changeset
   840
apply (subgoal_tac "0<n")
22718
936f7580937d tuned proofs;
wenzelm
parents: 22473
diff changeset
   841
 prefer 2 apply simp
15251
bb6f072c8d10 converted some induct_tac to induct
paulson
parents: 15140
diff changeset
   842
apply (induct_tac k rule: nat_less_induct)
14267
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14208
diff changeset
   843
apply (rename_tac "k")
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14208
diff changeset
   844
apply (case_tac "k<n", simp)
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14208
diff changeset
   845
apply (subgoal_tac "~ (k<m) ")
22718
936f7580937d tuned proofs;
wenzelm
parents: 22473
diff changeset
   846
 prefer 2 apply simp
14267
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14208
diff changeset
   847
apply (simp add: div_geq)
15251
bb6f072c8d10 converted some induct_tac to induct
paulson
parents: 15140
diff changeset
   848
apply (subgoal_tac "(k-n) div n \<le> (k-m) div n")
14267
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14208
diff changeset
   849
 prefer 2
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14208
diff changeset
   850
 apply (blast intro: div_le_mono diff_le_mono2)
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14208
diff changeset
   851
apply (rule le_trans, simp)
15439
71c0f98e31f1 made diff_less a simp rule
nipkow
parents: 15251
diff changeset
   852
apply (simp)
14267
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14208
diff changeset
   853
done
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14208
diff changeset
   854
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14208
diff changeset
   855
lemma div_le_dividend [simp]: "m div n \<le> (m::nat)"
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14208
diff changeset
   856
apply (case_tac "n=0", simp)
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14208
diff changeset
   857
apply (subgoal_tac "m div n \<le> m div 1", simp)
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14208
diff changeset
   858
apply (rule div_le_mono2)
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14208
diff changeset
   859
apply (simp_all (no_asm_simp))
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14208
diff changeset
   860
done
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14208
diff changeset
   861
22718
936f7580937d tuned proofs;
wenzelm
parents: 22473
diff changeset
   862
(* Similar for "less than" *)
17085
5b57f995a179 more simprules now have names
paulson
parents: 17084
diff changeset
   863
lemma div_less_dividend [rule_format]:
14267
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14208
diff changeset
   864
     "!!n::nat. 1<n ==> 0 < m --> m div n < m"
15251
bb6f072c8d10 converted some induct_tac to induct
paulson
parents: 15140
diff changeset
   865
apply (induct_tac m rule: nat_less_induct)
14267
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14208
diff changeset
   866
apply (rename_tac "m")
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14208
diff changeset
   867
apply (case_tac "m<n", simp)
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14208
diff changeset
   868
apply (subgoal_tac "0<n")
22718
936f7580937d tuned proofs;
wenzelm
parents: 22473
diff changeset
   869
 prefer 2 apply simp
14267
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14208
diff changeset
   870
apply (simp add: div_geq)
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14208
diff changeset
   871
apply (case_tac "n<m")
15251
bb6f072c8d10 converted some induct_tac to induct
paulson
parents: 15140
diff changeset
   872
 apply (subgoal_tac "(m-n) div n < (m-n) ")
14267
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14208
diff changeset
   873
  apply (rule impI less_trans_Suc)+
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14208
diff changeset
   874
apply assumption
15439
71c0f98e31f1 made diff_less a simp rule
nipkow
parents: 15251
diff changeset
   875
  apply (simp_all)
14267
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14208
diff changeset
   876
done
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14208
diff changeset
   877
17085
5b57f995a179 more simprules now have names
paulson
parents: 17084
diff changeset
   878
declare div_less_dividend [simp]
5b57f995a179 more simprules now have names
paulson
parents: 17084
diff changeset
   879
14267
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14208
diff changeset
   880
text{*A fact for the mutilated chess board*}
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14208
diff changeset
   881
lemma mod_Suc: "Suc(m) mod n = (if Suc(m mod n) = n then 0 else Suc(m mod n))"
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14208
diff changeset
   882
apply (case_tac "n=0", simp)
15251
bb6f072c8d10 converted some induct_tac to induct
paulson
parents: 15140
diff changeset
   883
apply (induct "m" rule: nat_less_induct)
14267
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14208
diff changeset
   884
apply (case_tac "Suc (na) <n")
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14208
diff changeset
   885
(* case Suc(na) < n *)
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14208
diff changeset
   886
apply (frule lessI [THEN less_trans], simp add: less_not_refl3)
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14208
diff changeset
   887
(* case n \<le> Suc(na) *)
16796
140f1e0ea846 generlization of some "nat" theorems
paulson
parents: 16733
diff changeset
   888
apply (simp add: linorder_not_less le_Suc_eq mod_geq)
15439
71c0f98e31f1 made diff_less a simp rule
nipkow
parents: 15251
diff changeset
   889
apply (auto simp add: Suc_diff_le le_mod_geq)
14267
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14208
diff changeset
   890
done
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14208
diff changeset
   891
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14208
diff changeset
   892
lemma mod_eq_0_iff: "(m mod d = 0) = (\<exists>q::nat. m = d*q)"
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29509
diff changeset
   893
by (auto simp add: dvd_eq_mod_eq_0 [symmetric] dvd_def)
17084
fb0a80aef0be classical rules must have names for ATP integration
paulson
parents: 16796
diff changeset
   894
22718
936f7580937d tuned proofs;
wenzelm
parents: 22473
diff changeset
   895
lemmas mod_eq_0D [dest!] = mod_eq_0_iff [THEN iffD1]
14267
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14208
diff changeset
   896
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14208
diff changeset
   897
(*Loses information, namely we also have r<d provided d is nonzero*)
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14208
diff changeset
   898
lemma mod_eqD: "(m mod d = r) ==> \<exists>q::nat. m = r + q*d"
27651
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27540
diff changeset
   899
  apply (cut_tac a = m in mod_div_equality)
22718
936f7580937d tuned proofs;
wenzelm
parents: 22473
diff changeset
   900
  apply (simp only: add_ac)
936f7580937d tuned proofs;
wenzelm
parents: 22473
diff changeset
   901
  apply (blast intro: sym)
936f7580937d tuned proofs;
wenzelm
parents: 22473
diff changeset
   902
  done
14267
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14208
diff changeset
   903
13152
2a54f99b44b3 Divides.ML -> Divides_lemmas.ML
nipkow
parents: 12338
diff changeset
   904
lemma split_div:
13189
81ed5c6de890 Now arith can deal with div/mod arbitrary nat numerals.
nipkow
parents: 13152
diff changeset
   905
 "P(n div k :: nat) =
81ed5c6de890 Now arith can deal with div/mod arbitrary nat numerals.
nipkow
parents: 13152
diff changeset
   906
 ((k = 0 \<longrightarrow> P 0) \<and> (k \<noteq> 0 \<longrightarrow> (!i. !j<k. n = k*i + j \<longrightarrow> P i)))"
81ed5c6de890 Now arith can deal with div/mod arbitrary nat numerals.
nipkow
parents: 13152
diff changeset
   907
 (is "?P = ?Q" is "_ = (_ \<and> (_ \<longrightarrow> ?R))")
81ed5c6de890 Now arith can deal with div/mod arbitrary nat numerals.
nipkow
parents: 13152
diff changeset
   908
proof
81ed5c6de890 Now arith can deal with div/mod arbitrary nat numerals.
nipkow
parents: 13152
diff changeset
   909
  assume P: ?P
81ed5c6de890 Now arith can deal with div/mod arbitrary nat numerals.
nipkow
parents: 13152
diff changeset
   910
  show ?Q
81ed5c6de890 Now arith can deal with div/mod arbitrary nat numerals.
nipkow
parents: 13152
diff changeset
   911
  proof (cases)
81ed5c6de890 Now arith can deal with div/mod arbitrary nat numerals.
nipkow
parents: 13152
diff changeset
   912
    assume "k = 0"
27651
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27540
diff changeset
   913
    with P show ?Q by simp
13189
81ed5c6de890 Now arith can deal with div/mod arbitrary nat numerals.
nipkow
parents: 13152
diff changeset
   914
  next
81ed5c6de890 Now arith can deal with div/mod arbitrary nat numerals.
nipkow
parents: 13152
diff changeset
   915
    assume not0: "k \<noteq> 0"
81ed5c6de890 Now arith can deal with div/mod arbitrary nat numerals.
nipkow
parents: 13152
diff changeset
   916
    thus ?Q
81ed5c6de890 Now arith can deal with div/mod arbitrary nat numerals.
nipkow
parents: 13152
diff changeset
   917
    proof (simp, intro allI impI)
81ed5c6de890 Now arith can deal with div/mod arbitrary nat numerals.
nipkow
parents: 13152
diff changeset
   918
      fix i j
81ed5c6de890 Now arith can deal with div/mod arbitrary nat numerals.
nipkow
parents: 13152
diff changeset
   919
      assume n: "n = k*i + j" and j: "j < k"
81ed5c6de890 Now arith can deal with div/mod arbitrary nat numerals.
nipkow
parents: 13152
diff changeset
   920
      show "P i"
81ed5c6de890 Now arith can deal with div/mod arbitrary nat numerals.
nipkow
parents: 13152
diff changeset
   921
      proof (cases)
22718
936f7580937d tuned proofs;
wenzelm
parents: 22473
diff changeset
   922
        assume "i = 0"
936f7580937d tuned proofs;
wenzelm
parents: 22473
diff changeset
   923
        with n j P show "P i" by simp
13189
81ed5c6de890 Now arith can deal with div/mod arbitrary nat numerals.
nipkow
parents: 13152
diff changeset
   924
      next
22718
936f7580937d tuned proofs;
wenzelm
parents: 22473
diff changeset
   925
        assume "i \<noteq> 0"
936f7580937d tuned proofs;
wenzelm
parents: 22473
diff changeset
   926
        with not0 n j P show "P i" by(simp add:add_ac)
13189
81ed5c6de890 Now arith can deal with div/mod arbitrary nat numerals.
nipkow
parents: 13152
diff changeset
   927
      qed
81ed5c6de890 Now arith can deal with div/mod arbitrary nat numerals.
nipkow
parents: 13152
diff changeset
   928
    qed
81ed5c6de890 Now arith can deal with div/mod arbitrary nat numerals.
nipkow
parents: 13152
diff changeset
   929
  qed
81ed5c6de890 Now arith can deal with div/mod arbitrary nat numerals.
nipkow
parents: 13152
diff changeset
   930
next
81ed5c6de890 Now arith can deal with div/mod arbitrary nat numerals.
nipkow
parents: 13152
diff changeset
   931
  assume Q: ?Q
81ed5c6de890 Now arith can deal with div/mod arbitrary nat numerals.
nipkow
parents: 13152
diff changeset
   932
  show ?P
81ed5c6de890 Now arith can deal with div/mod arbitrary nat numerals.
nipkow
parents: 13152
diff changeset
   933
  proof (cases)
81ed5c6de890 Now arith can deal with div/mod arbitrary nat numerals.
nipkow
parents: 13152
diff changeset
   934
    assume "k = 0"
27651
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27540
diff changeset
   935
    with Q show ?P by simp
13189
81ed5c6de890 Now arith can deal with div/mod arbitrary nat numerals.
nipkow
parents: 13152
diff changeset
   936
  next
81ed5c6de890 Now arith can deal with div/mod arbitrary nat numerals.
nipkow
parents: 13152
diff changeset
   937
    assume not0: "k \<noteq> 0"
81ed5c6de890 Now arith can deal with div/mod arbitrary nat numerals.
nipkow
parents: 13152
diff changeset
   938
    with Q have R: ?R by simp
81ed5c6de890 Now arith can deal with div/mod arbitrary nat numerals.
nipkow
parents: 13152
diff changeset
   939
    from not0 R[THEN spec,of "n div k",THEN spec, of "n mod k"]
13517
42efec18f5b2 Added div+mod cancelling simproc
nipkow
parents: 13189
diff changeset
   940
    show ?P by simp
13189
81ed5c6de890 Now arith can deal with div/mod arbitrary nat numerals.
nipkow
parents: 13152
diff changeset
   941
  qed
81ed5c6de890 Now arith can deal with div/mod arbitrary nat numerals.
nipkow
parents: 13152
diff changeset
   942
qed
81ed5c6de890 Now arith can deal with div/mod arbitrary nat numerals.
nipkow
parents: 13152
diff changeset
   943
13882
2266550ab316 New theorems split_div' and mod_div_equality'.
berghofe
parents: 13517
diff changeset
   944
lemma split_div_lemma:
26100
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   945
  assumes "0 < n"
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   946
  shows "n * q \<le> m \<and> m < n * Suc q \<longleftrightarrow> q = ((m\<Colon>nat) div n)" (is "?lhs \<longleftrightarrow> ?rhs")
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   947
proof
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   948
  assume ?rhs
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   949
  with mult_div_cancel have nq: "n * q = m - (m mod n)" by simp
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   950
  then have A: "n * q \<le> m" by simp
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   951
  have "n - (m mod n) > 0" using mod_less_divisor assms by auto
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   952
  then have "m < m + (n - (m mod n))" by simp
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   953
  then have "m < n + (m - (m mod n))" by simp
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   954
  with nq have "m < n + n * q" by simp
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   955
  then have B: "m < n * Suc q" by simp
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   956
  from A B show ?lhs ..
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   957
next
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   958
  assume P: ?lhs
33340
a165b97f3658 moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 33318
diff changeset
   959
  then have "divmod_nat_rel m n (q, m - n * q)"
a165b97f3658 moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 33318
diff changeset
   960
    unfolding divmod_nat_rel_def by (auto simp add: mult_ac)
a165b97f3658 moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 33318
diff changeset
   961
  with divmod_nat_rel_unique divmod_nat_rel [of m n]
30923
2697a1d1d34a more coherent developement in Divides.thy and IntDiv.thy
haftmann
parents: 30840
diff changeset
   962
  have "(q, m - n * q) = (m div n, m mod n)" by auto
2697a1d1d34a more coherent developement in Divides.thy and IntDiv.thy
haftmann
parents: 30840
diff changeset
   963
  then show ?rhs by simp
26100
fbc60cd02ae2 using only an relation predicate to construct div and mod
haftmann
parents: 26072
diff changeset
   964
qed
13882
2266550ab316 New theorems split_div' and mod_div_equality'.
berghofe
parents: 13517
diff changeset
   965
2266550ab316 New theorems split_div' and mod_div_equality'.
berghofe
parents: 13517
diff changeset
   966
theorem split_div':
2266550ab316 New theorems split_div' and mod_div_equality'.
berghofe
parents: 13517
diff changeset
   967
  "P ((m::nat) div n) = ((n = 0 \<and> P 0) \<or>
14267
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14208
diff changeset
   968
   (\<exists>q. (n * q \<le> m \<and> m < n * (Suc q)) \<and> P q))"
13882
2266550ab316 New theorems split_div' and mod_div_equality'.
berghofe
parents: 13517
diff changeset
   969
  apply (case_tac "0 < n")
2266550ab316 New theorems split_div' and mod_div_equality'.
berghofe
parents: 13517
diff changeset
   970
  apply (simp only: add: split_div_lemma)
27651
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27540
diff changeset
   971
  apply simp_all
13882
2266550ab316 New theorems split_div' and mod_div_equality'.
berghofe
parents: 13517
diff changeset
   972
  done
2266550ab316 New theorems split_div' and mod_div_equality'.
berghofe
parents: 13517
diff changeset
   973
13189
81ed5c6de890 Now arith can deal with div/mod arbitrary nat numerals.
nipkow
parents: 13152
diff changeset
   974
lemma split_mod:
81ed5c6de890 Now arith can deal with div/mod arbitrary nat numerals.
nipkow
parents: 13152
diff changeset
   975
 "P(n mod k :: nat) =
81ed5c6de890 Now arith can deal with div/mod arbitrary nat numerals.
nipkow
parents: 13152
diff changeset
   976
 ((k = 0 \<longrightarrow> P n) \<and> (k \<noteq> 0 \<longrightarrow> (!i. !j<k. n = k*i + j \<longrightarrow> P j)))"
81ed5c6de890 Now arith can deal with div/mod arbitrary nat numerals.
nipkow
parents: 13152
diff changeset
   977
 (is "?P = ?Q" is "_ = (_ \<and> (_ \<longrightarrow> ?R))")
81ed5c6de890 Now arith can deal with div/mod arbitrary nat numerals.
nipkow
parents: 13152
diff changeset
   978
proof
81ed5c6de890 Now arith can deal with div/mod arbitrary nat numerals.
nipkow
parents: 13152
diff changeset
   979
  assume P: ?P
81ed5c6de890 Now arith can deal with div/mod arbitrary nat numerals.
nipkow
parents: 13152
diff changeset
   980
  show ?Q
81ed5c6de890 Now arith can deal with div/mod arbitrary nat numerals.
nipkow
parents: 13152
diff changeset
   981
  proof (cases)
81ed5c6de890 Now arith can deal with div/mod arbitrary nat numerals.
nipkow
parents: 13152
diff changeset
   982
    assume "k = 0"
27651
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27540
diff changeset
   983
    with P show ?Q by simp
13189
81ed5c6de890 Now arith can deal with div/mod arbitrary nat numerals.
nipkow
parents: 13152
diff changeset
   984
  next
81ed5c6de890 Now arith can deal with div/mod arbitrary nat numerals.
nipkow
parents: 13152
diff changeset
   985
    assume not0: "k \<noteq> 0"
81ed5c6de890 Now arith can deal with div/mod arbitrary nat numerals.
nipkow
parents: 13152
diff changeset
   986
    thus ?Q
81ed5c6de890 Now arith can deal with div/mod arbitrary nat numerals.
nipkow
parents: 13152
diff changeset
   987
    proof (simp, intro allI impI)
81ed5c6de890 Now arith can deal with div/mod arbitrary nat numerals.
nipkow
parents: 13152
diff changeset
   988
      fix i j
81ed5c6de890 Now arith can deal with div/mod arbitrary nat numerals.
nipkow
parents: 13152
diff changeset
   989
      assume "n = k*i + j" "j < k"
81ed5c6de890 Now arith can deal with div/mod arbitrary nat numerals.
nipkow
parents: 13152
diff changeset
   990
      thus "P j" using not0 P by(simp add:add_ac mult_ac)
81ed5c6de890 Now arith can deal with div/mod arbitrary nat numerals.
nipkow
parents: 13152
diff changeset
   991
    qed
81ed5c6de890 Now arith can deal with div/mod arbitrary nat numerals.
nipkow
parents: 13152
diff changeset
   992
  qed
81ed5c6de890 Now arith can deal with div/mod arbitrary nat numerals.
nipkow
parents: 13152
diff changeset
   993
next
81ed5c6de890 Now arith can deal with div/mod arbitrary nat numerals.
nipkow
parents: 13152
diff changeset
   994
  assume Q: ?Q
81ed5c6de890 Now arith can deal with div/mod arbitrary nat numerals.
nipkow
parents: 13152
diff changeset
   995
  show ?P
81ed5c6de890 Now arith can deal with div/mod arbitrary nat numerals.
nipkow
parents: 13152
diff changeset
   996
  proof (cases)
81ed5c6de890 Now arith can deal with div/mod arbitrary nat numerals.
nipkow
parents: 13152
diff changeset
   997
    assume "k = 0"
27651
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27540
diff changeset
   998
    with Q show ?P by simp
13189
81ed5c6de890 Now arith can deal with div/mod arbitrary nat numerals.
nipkow
parents: 13152
diff changeset
   999
  next
81ed5c6de890 Now arith can deal with div/mod arbitrary nat numerals.
nipkow
parents: 13152
diff changeset
  1000
    assume not0: "k \<noteq> 0"
81ed5c6de890 Now arith can deal with div/mod arbitrary nat numerals.
nipkow
parents: 13152
diff changeset
  1001
    with Q have R: ?R by simp
81ed5c6de890 Now arith can deal with div/mod arbitrary nat numerals.
nipkow
parents: 13152
diff changeset
  1002
    from not0 R[THEN spec,of "n div k",THEN spec, of "n mod k"]
13517
42efec18f5b2 Added div+mod cancelling simproc
nipkow
parents: 13189
diff changeset
  1003
    show ?P by simp
13189
81ed5c6de890 Now arith can deal with div/mod arbitrary nat numerals.
nipkow
parents: 13152
diff changeset
  1004
  qed
81ed5c6de890 Now arith can deal with div/mod arbitrary nat numerals.
nipkow
parents: 13152
diff changeset
  1005
qed
81ed5c6de890 Now arith can deal with div/mod arbitrary nat numerals.
nipkow
parents: 13152
diff changeset
  1006
13882
2266550ab316 New theorems split_div' and mod_div_equality'.
berghofe
parents: 13517
diff changeset
  1007
theorem mod_div_equality': "(m::nat) mod n = m - (m div n) * n"
2266550ab316 New theorems split_div' and mod_div_equality'.
berghofe
parents: 13517
diff changeset
  1008
  apply (rule_tac P="%x. m mod n = x - (m div n) * n" in
2266550ab316 New theorems split_div' and mod_div_equality'.
berghofe
parents: 13517
diff changeset
  1009
    subst [OF mod_div_equality [of _ n]])
2266550ab316 New theorems split_div' and mod_div_equality'.
berghofe
parents: 13517
diff changeset
  1010
  apply arith
2266550ab316 New theorems split_div' and mod_div_equality'.
berghofe
parents: 13517
diff changeset
  1011
  done
2266550ab316 New theorems split_div' and mod_div_equality'.
berghofe
parents: 13517
diff changeset
  1012
22800
eaf5e7ef35d9 added lemmatas
haftmann
parents: 22744
diff changeset
  1013
lemma div_mod_equality':
eaf5e7ef35d9 added lemmatas
haftmann
parents: 22744
diff changeset
  1014
  fixes m n :: nat
eaf5e7ef35d9 added lemmatas
haftmann
parents: 22744
diff changeset
  1015
  shows "m div n * n = m - m mod n"
eaf5e7ef35d9 added lemmatas
haftmann
parents: 22744
diff changeset
  1016
proof -
eaf5e7ef35d9 added lemmatas
haftmann
parents: 22744
diff changeset
  1017
  have "m mod n \<le> m mod n" ..
eaf5e7ef35d9 added lemmatas
haftmann
parents: 22744
diff changeset
  1018
  from div_mod_equality have 
eaf5e7ef35d9 added lemmatas
haftmann
parents: 22744
diff changeset
  1019
    "m div n * n + m mod n - m mod n = m - m mod n" by simp
eaf5e7ef35d9 added lemmatas
haftmann
parents: 22744
diff changeset
  1020
  with diff_add_assoc [OF `m mod n \<le> m mod n`, of "m div n * n"] have
eaf5e7ef35d9 added lemmatas
haftmann
parents: 22744
diff changeset
  1021
    "m div n * n + (m mod n - m mod n) = m - m mod n"
eaf5e7ef35d9 added lemmatas
haftmann
parents: 22744
diff changeset
  1022
    by simp
eaf5e7ef35d9 added lemmatas
haftmann
parents: 22744
diff changeset
  1023
  then show ?thesis by simp
eaf5e7ef35d9 added lemmatas
haftmann
parents: 22744
diff changeset
  1024
qed
eaf5e7ef35d9 added lemmatas
haftmann
parents: 22744
diff changeset
  1025
eaf5e7ef35d9 added lemmatas
haftmann
parents: 22744
diff changeset
  1026
46551
866bce5442a3 simplify projections on simultaneous computations of div and mod; tuned structure (from Florian Haftmann)
huffman
parents: 46026
diff changeset
  1027
subsubsection {* An ``induction'' law for modulus arithmetic. *}
14640
b31870c50c68 new lemmas
paulson
parents: 14437
diff changeset
  1028
b31870c50c68 new lemmas
paulson
parents: 14437
diff changeset
  1029
lemma mod_induct_0:
b31870c50c68 new lemmas
paulson
parents: 14437
diff changeset
  1030
  assumes step: "\<forall>i<p. P i \<longrightarrow> P ((Suc i) mod p)"
b31870c50c68 new lemmas
paulson
parents: 14437
diff changeset
  1031
  and base: "P i" and i: "i<p"
b31870c50c68 new lemmas
paulson
parents: 14437
diff changeset
  1032
  shows "P 0"
b31870c50c68 new lemmas
paulson
parents: 14437
diff changeset
  1033
proof (rule ccontr)
b31870c50c68 new lemmas
paulson
parents: 14437
diff changeset
  1034
  assume contra: "\<not>(P 0)"
b31870c50c68 new lemmas
paulson
parents: 14437
diff changeset
  1035
  from i have p: "0<p" by simp
b31870c50c68 new lemmas
paulson
parents: 14437
diff changeset
  1036
  have "\<forall>k. 0<k \<longrightarrow> \<not> P (p-k)" (is "\<forall>k. ?A k")
b31870c50c68 new lemmas
paulson
parents: 14437
diff changeset
  1037
  proof
b31870c50c68 new lemmas
paulson
parents: 14437
diff changeset
  1038
    fix k
b31870c50c68 new lemmas
paulson
parents: 14437
diff changeset
  1039
    show "?A k"
b31870c50c68 new lemmas
paulson
parents: 14437
diff changeset
  1040
    proof (induct k)
b31870c50c68 new lemmas
paulson
parents: 14437
diff changeset
  1041
      show "?A 0" by simp  -- "by contradiction"
b31870c50c68 new lemmas
paulson
parents: 14437
diff changeset
  1042
    next
b31870c50c68 new lemmas
paulson
parents: 14437
diff changeset
  1043
      fix n
b31870c50c68 new lemmas
paulson
parents: 14437
diff changeset
  1044
      assume ih: "?A n"
b31870c50c68 new lemmas
paulson
parents: 14437
diff changeset
  1045
      show "?A (Suc n)"
b31870c50c68 new lemmas
paulson
parents: 14437
diff changeset
  1046
      proof (clarsimp)
22718
936f7580937d tuned proofs;
wenzelm
parents: 22473
diff changeset
  1047
        assume y: "P (p - Suc n)"
936f7580937d tuned proofs;
wenzelm
parents: 22473
diff changeset
  1048
        have n: "Suc n < p"
936f7580937d tuned proofs;
wenzelm
parents: 22473
diff changeset
  1049
        proof (rule ccontr)
936f7580937d tuned proofs;
wenzelm
parents: 22473
diff changeset
  1050
          assume "\<not>(Suc n < p)"
936f7580937d tuned proofs;
wenzelm
parents: 22473
diff changeset
  1051
          hence "p - Suc n = 0"
936f7580937d tuned proofs;
wenzelm
parents: 22473
diff changeset
  1052
            by simp
936f7580937d tuned proofs;
wenzelm
parents: 22473
diff changeset
  1053
          with y contra show "False"
936f7580937d tuned proofs;
wenzelm
parents: 22473
diff changeset
  1054
            by simp
936f7580937d tuned proofs;
wenzelm
parents: 22473
diff changeset
  1055
        qed
936f7580937d tuned proofs;
wenzelm
parents: 22473
diff changeset
  1056
        hence n2: "Suc (p - Suc n) = p-n" by arith
936f7580937d tuned proofs;
wenzelm
parents: 22473
diff changeset
  1057
        from p have "p - Suc n < p" by arith
936f7580937d tuned proofs;
wenzelm
parents: 22473
diff changeset
  1058
        with y step have z: "P ((Suc (p - Suc n)) mod p)"
936f7580937d tuned proofs;
wenzelm
parents: 22473
diff changeset
  1059
          by blast
936f7580937d tuned proofs;
wenzelm
parents: 22473
diff changeset
  1060
        show "False"
936f7580937d tuned proofs;
wenzelm
parents: 22473
diff changeset
  1061
        proof (cases "n=0")
936f7580937d tuned proofs;
wenzelm
parents: 22473
diff changeset
  1062
          case True
936f7580937d tuned proofs;
wenzelm
parents: 22473
diff changeset
  1063
          with z n2 contra show ?thesis by simp
936f7580937d tuned proofs;
wenzelm
parents: 22473
diff changeset
  1064
        next
936f7580937d tuned proofs;
wenzelm
parents: 22473
diff changeset
  1065
          case False
936f7580937d tuned proofs;
wenzelm
parents: 22473
diff changeset
  1066
          with p have "p-n < p" by arith
936f7580937d tuned proofs;
wenzelm
parents: 22473
diff changeset
  1067
          with z n2 False ih show ?thesis by simp
936f7580937d tuned proofs;
wenzelm
parents: 22473
diff changeset
  1068
        qed
14640
b31870c50c68 new lemmas
paulson
parents: 14437
diff changeset
  1069
      qed
b31870c50c68 new lemmas
paulson
parents: 14437
diff changeset
  1070
    qed
b31870c50c68 new lemmas
paulson
parents: 14437
diff changeset
  1071
  qed
b31870c50c68 new lemmas
paulson
parents: 14437
diff changeset
  1072
  moreover
b31870c50c68 new lemmas
paulson
parents: 14437
diff changeset
  1073
  from i obtain k where "0<k \<and> i+k=p"
b31870c50c68 new lemmas
paulson
parents: 14437
diff changeset
  1074
    by (blast dest: less_imp_add_positive)
b31870c50c68 new lemmas
paulson
parents: 14437
diff changeset
  1075
  hence "0<k \<and> i=p-k" by auto
b31870c50c68 new lemmas
paulson
parents: 14437
diff changeset
  1076
  moreover
b31870c50c68 new lemmas
paulson
parents: 14437
diff changeset
  1077
  note base
b31870c50c68 new lemmas
paulson
parents: 14437
diff changeset
  1078
  ultimately
b31870c50c68 new lemmas
paulson
parents: 14437
diff changeset
  1079
  show "False" by blast
b31870c50c68 new lemmas
paulson
parents: 14437
diff changeset
  1080
qed
b31870c50c68 new lemmas
paulson
parents: 14437
diff changeset
  1081
b31870c50c68 new lemmas
paulson
parents: 14437
diff changeset
  1082
lemma mod_induct:
b31870c50c68 new lemmas
paulson
parents: 14437
diff changeset
  1083
  assumes step: "\<forall>i<p. P i \<longrightarrow> P ((Suc i) mod p)"
b31870c50c68 new lemmas
paulson
parents: 14437
diff changeset
  1084
  and base: "P i" and i: "i<p" and j: "j<p"
b31870c50c68 new lemmas
paulson
parents: 14437
diff changeset
  1085
  shows "P j"
b31870c50c68 new lemmas
paulson
parents: 14437
diff changeset
  1086
proof -
b31870c50c68 new lemmas
paulson
parents: 14437
diff changeset
  1087
  have "\<forall>j<p. P j"
b31870c50c68 new lemmas
paulson
parents: 14437
diff changeset
  1088
  proof
b31870c50c68 new lemmas
paulson
parents: 14437
diff changeset
  1089
    fix j
b31870c50c68 new lemmas
paulson
parents: 14437
diff changeset
  1090
    show "j<p \<longrightarrow> P j" (is "?A j")
b31870c50c68 new lemmas
paulson
parents: 14437
diff changeset
  1091
    proof (induct j)
b31870c50c68 new lemmas
paulson
parents: 14437
diff changeset
  1092
      from step base i show "?A 0"
22718
936f7580937d tuned proofs;
wenzelm
parents: 22473
diff changeset
  1093
        by (auto elim: mod_induct_0)
14640
b31870c50c68 new lemmas
paulson
parents: 14437
diff changeset
  1094
    next
b31870c50c68 new lemmas
paulson
parents: 14437
diff changeset
  1095
      fix k
b31870c50c68 new lemmas
paulson
parents: 14437
diff changeset
  1096
      assume ih: "?A k"
b31870c50c68 new lemmas
paulson
parents: 14437
diff changeset
  1097
      show "?A (Suc k)"
b31870c50c68 new lemmas
paulson
parents: 14437
diff changeset
  1098
      proof
22718
936f7580937d tuned proofs;
wenzelm
parents: 22473
diff changeset
  1099
        assume suc: "Suc k < p"
936f7580937d tuned proofs;
wenzelm
parents: 22473
diff changeset
  1100
        hence k: "k<p" by simp
936f7580937d tuned proofs;
wenzelm
parents: 22473
diff changeset
  1101
        with ih have "P k" ..
936f7580937d tuned proofs;
wenzelm
parents: 22473
diff changeset
  1102
        with step k have "P (Suc k mod p)"
936f7580937d tuned proofs;
wenzelm
parents: 22473
diff changeset
  1103
          by blast
936f7580937d tuned proofs;
wenzelm
parents: 22473
diff changeset
  1104
        moreover
936f7580937d tuned proofs;
wenzelm
parents: 22473
diff changeset
  1105
        from suc have "Suc k mod p = Suc k"
936f7580937d tuned proofs;
wenzelm
parents: 22473
diff changeset
  1106
          by simp
936f7580937d tuned proofs;
wenzelm
parents: 22473
diff changeset
  1107
        ultimately
936f7580937d tuned proofs;
wenzelm
parents: 22473
diff changeset
  1108
        show "P (Suc k)" by simp
14640
b31870c50c68 new lemmas
paulson
parents: 14437
diff changeset
  1109
      qed
b31870c50c68 new lemmas
paulson
parents: 14437
diff changeset
  1110
    qed
b31870c50c68 new lemmas
paulson
parents: 14437
diff changeset
  1111
  qed
b31870c50c68 new lemmas
paulson
parents: 14437
diff changeset
  1112
  with j show ?thesis by blast
b31870c50c68 new lemmas
paulson
parents: 14437
diff changeset
  1113
qed
b31870c50c68 new lemmas
paulson
parents: 14437
diff changeset
  1114
33296
a3924d1069e5 moved theory Divides after theory Nat_Numeral; tuned some proof texts
haftmann
parents: 33274
diff changeset
  1115
lemma div2_Suc_Suc [simp]: "Suc (Suc m) div 2 = Suc (m div 2)"
a3924d1069e5 moved theory Divides after theory Nat_Numeral; tuned some proof texts
haftmann
parents: 33274
diff changeset
  1116
by (auto simp add: numeral_2_eq_2 le_div_geq)
a3924d1069e5 moved theory Divides after theory Nat_Numeral; tuned some proof texts
haftmann
parents: 33274
diff changeset
  1117
a3924d1069e5 moved theory Divides after theory Nat_Numeral; tuned some proof texts
haftmann
parents: 33274
diff changeset
  1118
lemma add_self_div_2 [simp]: "(m + m) div 2 = (m::nat)"
a3924d1069e5 moved theory Divides after theory Nat_Numeral; tuned some proof texts
haftmann
parents: 33274
diff changeset
  1119
by (simp add: nat_mult_2 [symmetric])
a3924d1069e5 moved theory Divides after theory Nat_Numeral; tuned some proof texts
haftmann
parents: 33274
diff changeset
  1120
a3924d1069e5 moved theory Divides after theory Nat_Numeral; tuned some proof texts
haftmann
parents: 33274
diff changeset
  1121
lemma mod2_Suc_Suc [simp]: "Suc(Suc(m)) mod 2 = m mod 2"
a3924d1069e5 moved theory Divides after theory Nat_Numeral; tuned some proof texts
haftmann
parents: 33274
diff changeset
  1122
apply (subgoal_tac "m mod 2 < 2")
a3924d1069e5 moved theory Divides after theory Nat_Numeral; tuned some proof texts
haftmann
parents: 33274
diff changeset
  1123
apply (erule less_2_cases [THEN disjE])
35216
7641e8d831d2 get rid of many duplicate simp rule warnings
huffman
parents: 35050
diff changeset
  1124
apply (simp_all (no_asm_simp) add: Let_def mod_Suc)
33296
a3924d1069e5 moved theory Divides after theory Nat_Numeral; tuned some proof texts
haftmann
parents: 33274
diff changeset
  1125
done
a3924d1069e5 moved theory Divides after theory Nat_Numeral; tuned some proof texts
haftmann
parents: 33274
diff changeset
  1126
a3924d1069e5 moved theory Divides after theory Nat_Numeral; tuned some proof texts
haftmann
parents: 33274
diff changeset
  1127
lemma mod2_gr_0 [simp]: "0 < (m\<Colon>nat) mod 2 \<longleftrightarrow> m mod 2 = 1"
a3924d1069e5 moved theory Divides after theory Nat_Numeral; tuned some proof texts
haftmann
parents: 33274
diff changeset
  1128
proof -
35815
10e723e54076 tuned proofs (to avoid linarith error message caused by bootstrapping of HOL)
boehmes
parents: 35673
diff changeset
  1129
  { fix n :: nat have  "(n::nat) < 2 \<Longrightarrow> n = 0 \<or> n = 1" by (cases n) simp_all }
33296
a3924d1069e5 moved theory Divides after theory Nat_Numeral; tuned some proof texts
haftmann
parents: 33274
diff changeset
  1130
  moreover have "m mod 2 < 2" by simp
a3924d1069e5 moved theory Divides after theory Nat_Numeral; tuned some proof texts
haftmann
parents: 33274
diff changeset
  1131
  ultimately have "m mod 2 = 0 \<or> m mod 2 = 1" .
a3924d1069e5 moved theory Divides after theory Nat_Numeral; tuned some proof texts
haftmann
parents: 33274
diff changeset
  1132
  then show ?thesis by auto
a3924d1069e5 moved theory Divides after theory Nat_Numeral; tuned some proof texts
haftmann
parents: 33274
diff changeset
  1133
qed
a3924d1069e5 moved theory Divides after theory Nat_Numeral; tuned some proof texts
haftmann
parents: 33274
diff changeset
  1134
a3924d1069e5 moved theory Divides after theory Nat_Numeral; tuned some proof texts
haftmann
parents: 33274
diff changeset
  1135
text{*These lemmas collapse some needless occurrences of Suc:
a3924d1069e5 moved theory Divides after theory Nat_Numeral; tuned some proof texts
haftmann
parents: 33274
diff changeset
  1136
    at least three Sucs, since two and fewer are rewritten back to Suc again!
a3924d1069e5 moved theory Divides after theory Nat_Numeral; tuned some proof texts
haftmann
parents: 33274
diff changeset
  1137
    We already have some rules to simplify operands smaller than 3.*}
a3924d1069e5 moved theory Divides after theory Nat_Numeral; tuned some proof texts
haftmann
parents: 33274
diff changeset
  1138
a3924d1069e5 moved theory Divides after theory Nat_Numeral; tuned some proof texts
haftmann
parents: 33274
diff changeset
  1139
lemma div_Suc_eq_div_add3 [simp]: "m div (Suc (Suc (Suc n))) = m div (3+n)"
a3924d1069e5 moved theory Divides after theory Nat_Numeral; tuned some proof texts
haftmann
parents: 33274
diff changeset
  1140
by (simp add: Suc3_eq_add_3)
a3924d1069e5 moved theory Divides after theory Nat_Numeral; tuned some proof texts
haftmann
parents: 33274
diff changeset
  1141
a3924d1069e5 moved theory Divides after theory Nat_Numeral; tuned some proof texts
haftmann
parents: 33274
diff changeset
  1142
lemma mod_Suc_eq_mod_add3 [simp]: "m mod (Suc (Suc (Suc n))) = m mod (3+n)"
a3924d1069e5 moved theory Divides after theory Nat_Numeral; tuned some proof texts
haftmann
parents: 33274
diff changeset
  1143
by (simp add: Suc3_eq_add_3)
a3924d1069e5 moved theory Divides after theory Nat_Numeral; tuned some proof texts
haftmann
parents: 33274
diff changeset
  1144
a3924d1069e5 moved theory Divides after theory Nat_Numeral; tuned some proof texts
haftmann
parents: 33274
diff changeset
  1145
lemma Suc_div_eq_add3_div: "(Suc (Suc (Suc m))) div n = (3+m) div n"
a3924d1069e5 moved theory Divides after theory Nat_Numeral; tuned some proof texts
haftmann
parents: 33274
diff changeset
  1146
by (simp add: Suc3_eq_add_3)
a3924d1069e5 moved theory Divides after theory Nat_Numeral; tuned some proof texts
haftmann
parents: 33274
diff changeset
  1147
a3924d1069e5 moved theory Divides after theory Nat_Numeral; tuned some proof texts
haftmann
parents: 33274
diff changeset
  1148
lemma Suc_mod_eq_add3_mod: "(Suc (Suc (Suc m))) mod n = (3+m) mod n"
a3924d1069e5 moved theory Divides after theory Nat_Numeral; tuned some proof texts
haftmann
parents: 33274
diff changeset
  1149
by (simp add: Suc3_eq_add_3)
a3924d1069e5 moved theory Divides after theory Nat_Numeral; tuned some proof texts
haftmann
parents: 33274
diff changeset
  1150
45607
16b4f5774621 eliminated obsolete "standard";
wenzelm
parents: 45530
diff changeset
  1151
lemmas Suc_div_eq_add3_div_number_of [simp] = Suc_div_eq_add3_div [of _ "number_of v"] for v
16b4f5774621 eliminated obsolete "standard";
wenzelm
parents: 45530
diff changeset
  1152
lemmas Suc_mod_eq_add3_mod_number_of [simp] = Suc_mod_eq_add3_mod [of _ "number_of v"] for v
33296
a3924d1069e5 moved theory Divides after theory Nat_Numeral; tuned some proof texts
haftmann
parents: 33274
diff changeset
  1153
33361
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1154
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1155
lemma Suc_times_mod_eq: "1<k ==> Suc (k * m) mod k = 1" 
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1156
apply (induct "m")
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1157
apply (simp_all add: mod_Suc)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1158
done
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1159
45607
16b4f5774621 eliminated obsolete "standard";
wenzelm
parents: 45530
diff changeset
  1160
declare Suc_times_mod_eq [of "number_of w", simp] for w
33361
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1161
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1162
lemma [simp]: "n div k \<le> (Suc n) div k"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1163
by (simp add: div_le_mono) 
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1164
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1165
lemma Suc_n_div_2_gt_zero [simp]: "(0::nat) < n ==> 0 < (n + 1) div 2"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1166
by (cases n) simp_all
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1167
35815
10e723e54076 tuned proofs (to avoid linarith error message caused by bootstrapping of HOL)
boehmes
parents: 35673
diff changeset
  1168
lemma div_2_gt_zero [simp]: assumes A: "(1::nat) < n" shows "0 < n div 2"
10e723e54076 tuned proofs (to avoid linarith error message caused by bootstrapping of HOL)
boehmes
parents: 35673
diff changeset
  1169
proof -
10e723e54076 tuned proofs (to avoid linarith error message caused by bootstrapping of HOL)
boehmes
parents: 35673
diff changeset
  1170
  from A have B: "0 < n - 1" and C: "n - 1 + 1 = n" by simp_all
10e723e54076 tuned proofs (to avoid linarith error message caused by bootstrapping of HOL)
boehmes
parents: 35673
diff changeset
  1171
  from Suc_n_div_2_gt_zero [OF B] C show ?thesis by simp 
10e723e54076 tuned proofs (to avoid linarith error message caused by bootstrapping of HOL)
boehmes
parents: 35673
diff changeset
  1172
qed
33361
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1173
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1174
  (* Potential use of algebra : Equality modulo n*)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1175
lemma mod_mult_self3 [simp]: "(k*n + m) mod n = m mod (n::nat)"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1176
by (simp add: mult_ac add_ac)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1177
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1178
lemma mod_mult_self4 [simp]: "Suc (k*n + m) mod n = Suc m mod n"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1179
proof -
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1180
  have "Suc (k * n + m) mod n = (k * n + Suc m) mod n" by simp
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1181
  also have "... = Suc m mod n" by (rule mod_mult_self3) 
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1182
  finally show ?thesis .
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1183
qed
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1184
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1185
lemma mod_Suc_eq_Suc_mod: "Suc m mod n = Suc (m mod n) mod n"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1186
apply (subst mod_Suc [of m]) 
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1187
apply (subst mod_Suc [of "m mod n"], simp) 
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1188
done
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1189
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1190
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1191
subsection {* Division on @{typ int} *}
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1192
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1193
definition divmod_int_rel :: "int \<Rightarrow> int \<Rightarrow> int \<times> int \<Rightarrow> bool" where
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1194
    --{*definition of quotient and remainder*}
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1195
    [code]: "divmod_int_rel a b = (\<lambda>(q, r). a = b * q + r \<and>
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1196
               (if 0 < b then 0 \<le> r \<and> r < b else b < r \<and> r \<le> 0))"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1197
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1198
definition adjust :: "int \<Rightarrow> int \<times> int \<Rightarrow> int \<times> int" where
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1199
    --{*for the division algorithm*}
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1200
    [code]: "adjust b = (\<lambda>(q, r). if 0 \<le> r - b then (2 * q + 1, r - b)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1201
                         else (2 * q, r))"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1202
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1203
text{*algorithm for the case @{text "a\<ge>0, b>0"}*}
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1204
function posDivAlg :: "int \<Rightarrow> int \<Rightarrow> int \<times> int" where
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1205
  "posDivAlg a b = (if a < b \<or>  b \<le> 0 then (0, a)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1206
     else adjust b (posDivAlg a (2 * b)))"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1207
by auto
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1208
termination by (relation "measure (\<lambda>(a, b). nat (a - b + 1))")
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1209
  (auto simp add: mult_2)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1210
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1211
text{*algorithm for the case @{text "a<0, b>0"}*}
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1212
function negDivAlg :: "int \<Rightarrow> int \<Rightarrow> int \<times> int" where
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1213
  "negDivAlg a b = (if 0 \<le>a + b \<or> b \<le> 0  then (-1, a + b)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1214
     else adjust b (negDivAlg a (2 * b)))"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1215
by auto
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1216
termination by (relation "measure (\<lambda>(a, b). nat (- a - b))")
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1217
  (auto simp add: mult_2)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1218
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1219
text{*algorithm for the general case @{term "b\<noteq>0"}*}
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1220
definition negateSnd :: "int \<times> int \<Rightarrow> int \<times> int" where
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1221
  [code_unfold]: "negateSnd = apsnd uminus"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1222
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1223
definition divmod_int :: "int \<Rightarrow> int \<Rightarrow> int \<times> int" where
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1224
    --{*The full division algorithm considers all possible signs for a, b
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1225
       including the special case @{text "a=0, b<0"} because 
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1226
       @{term negDivAlg} requires @{term "a<0"}.*}
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1227
  "divmod_int a b = (if 0 \<le> a then if 0 \<le> b then posDivAlg a b
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1228
                  else if a = 0 then (0, 0)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1229
                       else negateSnd (negDivAlg (-a) (-b))
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1230
               else 
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1231
                  if 0 < b then negDivAlg a b
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1232
                  else negateSnd (posDivAlg (-a) (-b)))"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1233
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1234
instantiation int :: Divides.div
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1235
begin
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1236
46551
866bce5442a3 simplify projections on simultaneous computations of div and mod; tuned structure (from Florian Haftmann)
huffman
parents: 46026
diff changeset
  1237
definition div_int where
33361
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1238
  "a div b = fst (divmod_int a b)"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1239
46551
866bce5442a3 simplify projections on simultaneous computations of div and mod; tuned structure (from Florian Haftmann)
huffman
parents: 46026
diff changeset
  1240
lemma fst_divmod_int [simp]:
866bce5442a3 simplify projections on simultaneous computations of div and mod; tuned structure (from Florian Haftmann)
huffman
parents: 46026
diff changeset
  1241
  "fst (divmod_int a b) = a div b"
866bce5442a3 simplify projections on simultaneous computations of div and mod; tuned structure (from Florian Haftmann)
huffman
parents: 46026
diff changeset
  1242
  by (simp add: div_int_def)
866bce5442a3 simplify projections on simultaneous computations of div and mod; tuned structure (from Florian Haftmann)
huffman
parents: 46026
diff changeset
  1243
866bce5442a3 simplify projections on simultaneous computations of div and mod; tuned structure (from Florian Haftmann)
huffman
parents: 46026
diff changeset
  1244
definition mod_int where
33361
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1245
 "a mod b = snd (divmod_int a b)"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1246
46551
866bce5442a3 simplify projections on simultaneous computations of div and mod; tuned structure (from Florian Haftmann)
huffman
parents: 46026
diff changeset
  1247
lemma snd_divmod_int [simp]:
866bce5442a3 simplify projections on simultaneous computations of div and mod; tuned structure (from Florian Haftmann)
huffman
parents: 46026
diff changeset
  1248
  "snd (divmod_int a b) = a mod b"
866bce5442a3 simplify projections on simultaneous computations of div and mod; tuned structure (from Florian Haftmann)
huffman
parents: 46026
diff changeset
  1249
  by (simp add: mod_int_def)
866bce5442a3 simplify projections on simultaneous computations of div and mod; tuned structure (from Florian Haftmann)
huffman
parents: 46026
diff changeset
  1250
33361
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1251
instance ..
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1252
3366
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
  1253
end
33361
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1254
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1255
lemma divmod_int_mod_div:
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1256
  "divmod_int p q = (p div q, p mod q)"
46551
866bce5442a3 simplify projections on simultaneous computations of div and mod; tuned structure (from Florian Haftmann)
huffman
parents: 46026
diff changeset
  1257
  by (simp add: prod_eq_iff)
33361
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1258
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1259
text{*
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1260
Here is the division algorithm in ML:
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1261
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1262
\begin{verbatim}
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1263
    fun posDivAlg (a,b) =
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1264
      if a<b then (0,a)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1265
      else let val (q,r) = posDivAlg(a, 2*b)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1266
               in  if 0\<le>r-b then (2*q+1, r-b) else (2*q, r)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1267
           end
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1268
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1269
    fun negDivAlg (a,b) =
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1270
      if 0\<le>a+b then (~1,a+b)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1271
      else let val (q,r) = negDivAlg(a, 2*b)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1272
               in  if 0\<le>r-b then (2*q+1, r-b) else (2*q, r)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1273
           end;
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1274
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1275
    fun negateSnd (q,r:int) = (q,~r);
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1276
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1277
    fun divmod (a,b) = if 0\<le>a then 
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1278
                          if b>0 then posDivAlg (a,b) 
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1279
                           else if a=0 then (0,0)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1280
                                else negateSnd (negDivAlg (~a,~b))
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1281
                       else 
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1282
                          if 0<b then negDivAlg (a,b)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1283
                          else        negateSnd (posDivAlg (~a,~b));
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1284
\end{verbatim}
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1285
*}
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1286
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1287
46551
866bce5442a3 simplify projections on simultaneous computations of div and mod; tuned structure (from Florian Haftmann)
huffman
parents: 46026
diff changeset
  1288
subsubsection {* Uniqueness and Monotonicity of Quotients and Remainders *}
33361
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1289
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1290
lemma unique_quotient_lemma:
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1291
     "[| b*q' + r'  \<le> b*q + r;  0 \<le> r';  r' < b;  r < b |]  
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1292
      ==> q' \<le> (q::int)"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1293
apply (subgoal_tac "r' + b * (q'-q) \<le> r")
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1294
 prefer 2 apply (simp add: right_diff_distrib)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1295
apply (subgoal_tac "0 < b * (1 + q - q') ")
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1296
apply (erule_tac [2] order_le_less_trans)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1297
 prefer 2 apply (simp add: right_diff_distrib right_distrib)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1298
apply (subgoal_tac "b * q' < b * (1 + q) ")
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1299
 prefer 2 apply (simp add: right_diff_distrib right_distrib)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1300
apply (simp add: mult_less_cancel_left)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1301
done
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1302
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1303
lemma unique_quotient_lemma_neg:
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1304
     "[| b*q' + r' \<le> b*q + r;  r \<le> 0;  b < r;  b < r' |]  
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1305
      ==> q \<le> (q'::int)"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1306
by (rule_tac b = "-b" and r = "-r'" and r' = "-r" in unique_quotient_lemma, 
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1307
    auto)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1308
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1309
lemma unique_quotient:
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1310
     "[| divmod_int_rel a b (q, r); divmod_int_rel a b (q', r');  b \<noteq> 0 |]  
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1311
      ==> q = q'"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1312
apply (simp add: divmod_int_rel_def linorder_neq_iff split: split_if_asm)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1313
apply (blast intro: order_antisym
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1314
             dest: order_eq_refl [THEN unique_quotient_lemma] 
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1315
             order_eq_refl [THEN unique_quotient_lemma_neg] sym)+
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1316
done
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1317
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1318
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1319
lemma unique_remainder:
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1320
     "[| divmod_int_rel a b (q, r); divmod_int_rel a b (q', r');  b \<noteq> 0 |]  
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1321
      ==> r = r'"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1322
apply (subgoal_tac "q = q'")
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1323
 apply (simp add: divmod_int_rel_def)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1324
apply (blast intro: unique_quotient)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1325
done
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1326
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1327
46551
866bce5442a3 simplify projections on simultaneous computations of div and mod; tuned structure (from Florian Haftmann)
huffman
parents: 46026
diff changeset
  1328
subsubsection {* Correctness of @{term posDivAlg}, the Algorithm for Non-Negative Dividends *}
33361
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1329
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1330
text{*And positive divisors*}
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1331
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1332
lemma adjust_eq [simp]:
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1333
     "adjust b (q,r) = 
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1334
      (let diff = r-b in  
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1335
        if 0 \<le> diff then (2*q + 1, diff)   
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1336
                     else (2*q, r))"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1337
by (simp add: Let_def adjust_def)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1338
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1339
declare posDivAlg.simps [simp del]
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1340
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1341
text{*use with a simproc to avoid repeatedly proving the premise*}
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1342
lemma posDivAlg_eqn:
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1343
     "0 < b ==>  
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1344
      posDivAlg a b = (if a<b then (0,a) else adjust b (posDivAlg a (2*b)))"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1345
by (rule posDivAlg.simps [THEN trans], simp)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1346
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1347
text{*Correctness of @{term posDivAlg}: it computes quotients correctly*}
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1348
theorem posDivAlg_correct:
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1349
  assumes "0 \<le> a" and "0 < b"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1350
  shows "divmod_int_rel a b (posDivAlg a b)"
41550
efa734d9b221 eliminated global prems;
wenzelm
parents: 39489
diff changeset
  1351
  using assms
efa734d9b221 eliminated global prems;
wenzelm
parents: 39489
diff changeset
  1352
  apply (induct a b rule: posDivAlg.induct)
efa734d9b221 eliminated global prems;
wenzelm
parents: 39489
diff changeset
  1353
  apply auto
efa734d9b221 eliminated global prems;
wenzelm
parents: 39489
diff changeset
  1354
  apply (simp add: divmod_int_rel_def)
efa734d9b221 eliminated global prems;
wenzelm
parents: 39489
diff changeset
  1355
  apply (subst posDivAlg_eqn, simp add: right_distrib)
efa734d9b221 eliminated global prems;
wenzelm
parents: 39489
diff changeset
  1356
  apply (case_tac "a < b")
efa734d9b221 eliminated global prems;
wenzelm
parents: 39489
diff changeset
  1357
  apply simp_all
efa734d9b221 eliminated global prems;
wenzelm
parents: 39489
diff changeset
  1358
  apply (erule splitE)
efa734d9b221 eliminated global prems;
wenzelm
parents: 39489
diff changeset
  1359
  apply (auto simp add: right_distrib Let_def mult_ac mult_2_right)
efa734d9b221 eliminated global prems;
wenzelm
parents: 39489
diff changeset
  1360
  done
33361
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1361
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1362
46551
866bce5442a3 simplify projections on simultaneous computations of div and mod; tuned structure (from Florian Haftmann)
huffman
parents: 46026
diff changeset
  1363
subsubsection {* Correctness of @{term negDivAlg}, the Algorithm for Negative Dividends *}
33361
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1364
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1365
text{*And positive divisors*}
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1366
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1367
declare negDivAlg.simps [simp del]
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1368
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1369
text{*use with a simproc to avoid repeatedly proving the premise*}
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1370
lemma negDivAlg_eqn:
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1371
     "0 < b ==>  
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1372
      negDivAlg a b =       
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1373
       (if 0\<le>a+b then (-1,a+b) else adjust b (negDivAlg a (2*b)))"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1374
by (rule negDivAlg.simps [THEN trans], simp)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1375
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1376
(*Correctness of negDivAlg: it computes quotients correctly
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1377
  It doesn't work if a=0 because the 0/b equals 0, not -1*)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1378
lemma negDivAlg_correct:
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1379
  assumes "a < 0" and "b > 0"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1380
  shows "divmod_int_rel a b (negDivAlg a b)"
41550
efa734d9b221 eliminated global prems;
wenzelm
parents: 39489
diff changeset
  1381
  using assms
efa734d9b221 eliminated global prems;
wenzelm
parents: 39489
diff changeset
  1382
  apply (induct a b rule: negDivAlg.induct)
efa734d9b221 eliminated global prems;
wenzelm
parents: 39489
diff changeset
  1383
  apply (auto simp add: linorder_not_le)
efa734d9b221 eliminated global prems;
wenzelm
parents: 39489
diff changeset
  1384
  apply (simp add: divmod_int_rel_def)
efa734d9b221 eliminated global prems;
wenzelm
parents: 39489
diff changeset
  1385
  apply (subst negDivAlg_eqn, assumption)
efa734d9b221 eliminated global prems;
wenzelm
parents: 39489
diff changeset
  1386
  apply (case_tac "a + b < (0\<Colon>int)")
efa734d9b221 eliminated global prems;
wenzelm
parents: 39489
diff changeset
  1387
  apply simp_all
efa734d9b221 eliminated global prems;
wenzelm
parents: 39489
diff changeset
  1388
  apply (erule splitE)
efa734d9b221 eliminated global prems;
wenzelm
parents: 39489
diff changeset
  1389
  apply (auto simp add: right_distrib Let_def mult_ac mult_2_right)
efa734d9b221 eliminated global prems;
wenzelm
parents: 39489
diff changeset
  1390
  done
33361
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1391
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1392
46551
866bce5442a3 simplify projections on simultaneous computations of div and mod; tuned structure (from Florian Haftmann)
huffman
parents: 46026
diff changeset
  1393
subsubsection {* Existence Shown by Proving the Division Algorithm to be Correct *}
33361
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1394
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1395
(*the case a=0*)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1396
lemma divmod_int_rel_0: "b \<noteq> 0 ==> divmod_int_rel 0 b (0, 0)"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1397
by (auto simp add: divmod_int_rel_def linorder_neq_iff)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1398
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1399
lemma posDivAlg_0 [simp]: "posDivAlg 0 b = (0, 0)"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1400
by (subst posDivAlg.simps, auto)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1401
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1402
lemma negDivAlg_minus1 [simp]: "negDivAlg -1 b = (-1, b - 1)"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1403
by (subst negDivAlg.simps, auto)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1404
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1405
lemma negateSnd_eq [simp]: "negateSnd(q,r) = (q,-r)"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1406
by (simp add: negateSnd_def)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1407
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1408
lemma divmod_int_rel_neg: "divmod_int_rel (-a) (-b) qr ==> divmod_int_rel a b (negateSnd qr)"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1409
by (auto simp add: split_ifs divmod_int_rel_def)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1410
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1411
lemma divmod_int_correct: "b \<noteq> 0 ==> divmod_int_rel a b (divmod_int a b)"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1412
by (force simp add: linorder_neq_iff divmod_int_rel_0 divmod_int_def divmod_int_rel_neg
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1413
                    posDivAlg_correct negDivAlg_correct)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1414
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1415
text{*Arbitrary definitions for division by zero.  Useful to simplify 
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1416
    certain equations.*}
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1417
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1418
lemma DIVISION_BY_ZERO [simp]: "a div (0::int) = 0 & a mod (0::int) = a"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1419
by (simp add: div_int_def mod_int_def divmod_int_def posDivAlg.simps)  
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1420
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1421
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1422
text{*Basic laws about division and remainder*}
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1423
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1424
lemma zmod_zdiv_equality: "(a::int) = b * (a div b) + (a mod b)"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1425
apply (case_tac "b = 0", simp)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1426
apply (cut_tac a = a and b = b in divmod_int_correct)
46551
866bce5442a3 simplify projections on simultaneous computations of div and mod; tuned structure (from Florian Haftmann)
huffman
parents: 46026
diff changeset
  1427
apply (auto simp add: divmod_int_rel_def prod_eq_iff)
33361
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1428
done
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1429
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1430
lemma zdiv_zmod_equality: "(b * (a div b) + (a mod b)) + k = (a::int)+k"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1431
by(simp add: zmod_zdiv_equality[symmetric])
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1432
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1433
lemma zdiv_zmod_equality2: "((a div b) * b + (a mod b)) + k = (a::int)+k"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1434
by(simp add: mult_commute zmod_zdiv_equality[symmetric])
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1435
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1436
text {* Tool setup *}
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1437
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1438
ML {*
43594
ef1ddc59b825 modernized some simproc setup;
wenzelm
parents: 41792
diff changeset
  1439
structure Cancel_Div_Mod_Int = Cancel_Div_Mod
41550
efa734d9b221 eliminated global prems;
wenzelm
parents: 39489
diff changeset
  1440
(
33361
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1441
  val div_name = @{const_name div};
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1442
  val mod_name = @{const_name mod};
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1443
  val mk_binop = HOLogic.mk_binop;
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1444
  val mk_sum = Arith_Data.mk_sum HOLogic.intT;
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1445
  val dest_sum = Arith_Data.dest_sum;
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1446
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1447
  val div_mod_eqs = map mk_meta_eq [@{thm zdiv_zmod_equality}, @{thm zdiv_zmod_equality2}];
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1448
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1449
  val prove_eq_sums = Arith_Data.prove_conv2 all_tac (Arith_Data.simp_all_tac 
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1450
    (@{thm diff_minus} :: @{thms add_0s} @ @{thms add_ac}))
41550
efa734d9b221 eliminated global prems;
wenzelm
parents: 39489
diff changeset
  1451
)
33361
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1452
*}
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1453
43594
ef1ddc59b825 modernized some simproc setup;
wenzelm
parents: 41792
diff changeset
  1454
simproc_setup cancel_div_mod_int ("(k::int) + l") = {* K Cancel_Div_Mod_Int.proc *}
ef1ddc59b825 modernized some simproc setup;
wenzelm
parents: 41792
diff changeset
  1455
33361
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1456
lemma pos_mod_conj : "(0::int) < b ==> 0 \<le> a mod b & a mod b < b"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1457
apply (cut_tac a = a and b = b in divmod_int_correct)
46551
866bce5442a3 simplify projections on simultaneous computations of div and mod; tuned structure (from Florian Haftmann)
huffman
parents: 46026
diff changeset
  1458
apply (auto simp add: divmod_int_rel_def prod_eq_iff)
33361
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1459
done
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1460
45607
16b4f5774621 eliminated obsolete "standard";
wenzelm
parents: 45530
diff changeset
  1461
lemmas pos_mod_sign [simp] = pos_mod_conj [THEN conjunct1]
16b4f5774621 eliminated obsolete "standard";
wenzelm
parents: 45530
diff changeset
  1462
   and pos_mod_bound [simp] = pos_mod_conj [THEN conjunct2]
33361
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1463
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1464
lemma neg_mod_conj : "b < (0::int) ==> a mod b \<le> 0 & b < a mod b"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1465
apply (cut_tac a = a and b = b in divmod_int_correct)
46551
866bce5442a3 simplify projections on simultaneous computations of div and mod; tuned structure (from Florian Haftmann)
huffman
parents: 46026
diff changeset
  1466
apply (auto simp add: divmod_int_rel_def prod_eq_iff)
33361
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1467
done
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1468
45607
16b4f5774621 eliminated obsolete "standard";
wenzelm
parents: 45530
diff changeset
  1469
lemmas neg_mod_sign [simp] = neg_mod_conj [THEN conjunct1]
16b4f5774621 eliminated obsolete "standard";
wenzelm
parents: 45530
diff changeset
  1470
   and neg_mod_bound [simp] = neg_mod_conj [THEN conjunct2]
33361
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1471
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1472
46551
866bce5442a3 simplify projections on simultaneous computations of div and mod; tuned structure (from Florian Haftmann)
huffman
parents: 46026
diff changeset
  1473
subsubsection {* General Properties of div and mod *}
33361
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1474
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1475
lemma divmod_int_rel_div_mod: "b \<noteq> 0 ==> divmod_int_rel a b (a div b, a mod b)"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1476
apply (cut_tac a = a and b = b in zmod_zdiv_equality)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1477
apply (force simp add: divmod_int_rel_def linorder_neq_iff)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1478
done
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1479
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1480
lemma divmod_int_rel_div: "[| divmod_int_rel a b (q, r);  b \<noteq> 0 |] ==> a div b = q"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1481
by (simp add: divmod_int_rel_div_mod [THEN unique_quotient])
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1482
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1483
lemma divmod_int_rel_mod: "[| divmod_int_rel a b (q, r);  b \<noteq> 0 |] ==> a mod b = r"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1484
by (simp add: divmod_int_rel_div_mod [THEN unique_remainder])
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1485
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1486
lemma div_pos_pos_trivial: "[| (0::int) \<le> a;  a < b |] ==> a div b = 0"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1487
apply (rule divmod_int_rel_div)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1488
apply (auto simp add: divmod_int_rel_def)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1489
done
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1490
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1491
lemma div_neg_neg_trivial: "[| a \<le> (0::int);  b < a |] ==> a div b = 0"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1492
apply (rule divmod_int_rel_div)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1493
apply (auto simp add: divmod_int_rel_def)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1494
done
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1495
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1496
lemma div_pos_neg_trivial: "[| (0::int) < a;  a+b \<le> 0 |] ==> a div b = -1"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1497
apply (rule divmod_int_rel_div)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1498
apply (auto simp add: divmod_int_rel_def)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1499
done
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1500
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1501
(*There is no div_neg_pos_trivial because  0 div b = 0 would supersede it*)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1502
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1503
lemma mod_pos_pos_trivial: "[| (0::int) \<le> a;  a < b |] ==> a mod b = a"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1504
apply (rule_tac q = 0 in divmod_int_rel_mod)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1505
apply (auto simp add: divmod_int_rel_def)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1506
done
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1507
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1508
lemma mod_neg_neg_trivial: "[| a \<le> (0::int);  b < a |] ==> a mod b = a"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1509
apply (rule_tac q = 0 in divmod_int_rel_mod)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1510
apply (auto simp add: divmod_int_rel_def)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1511
done
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1512
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1513
lemma mod_pos_neg_trivial: "[| (0::int) < a;  a+b \<le> 0 |] ==> a mod b = a+b"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1514
apply (rule_tac q = "-1" in divmod_int_rel_mod)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1515
apply (auto simp add: divmod_int_rel_def)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1516
done
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1517
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1518
text{*There is no @{text mod_neg_pos_trivial}.*}
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1519
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1520
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1521
(*Simpler laws such as -a div b = -(a div b) FAIL, but see just below*)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1522
lemma zdiv_zminus_zminus [simp]: "(-a) div (-b) = a div (b::int)"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1523
apply (case_tac "b = 0", simp)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1524
apply (simp add: divmod_int_rel_div_mod [THEN divmod_int_rel_neg, simplified, 
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1525
                                 THEN divmod_int_rel_div, THEN sym])
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1526
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1527
done
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1528
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1529
(*Simpler laws such as -a mod b = -(a mod b) FAIL, but see just below*)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1530
lemma zmod_zminus_zminus [simp]: "(-a) mod (-b) = - (a mod (b::int))"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1531
apply (case_tac "b = 0", simp)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1532
apply (subst divmod_int_rel_div_mod [THEN divmod_int_rel_neg, simplified, THEN divmod_int_rel_mod],
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1533
       auto)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1534
done
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1535
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1536
46551
866bce5442a3 simplify projections on simultaneous computations of div and mod; tuned structure (from Florian Haftmann)
huffman
parents: 46026
diff changeset
  1537
subsubsection {* Laws for div and mod with Unary Minus *}
33361
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1538
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1539
lemma zminus1_lemma:
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1540
     "divmod_int_rel a b (q, r)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1541
      ==> divmod_int_rel (-a) b (if r=0 then -q else -q - 1,  
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1542
                          if r=0 then 0 else b-r)"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1543
by (force simp add: split_ifs divmod_int_rel_def linorder_neq_iff right_diff_distrib)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1544
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1545
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1546
lemma zdiv_zminus1_eq_if:
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1547
     "b \<noteq> (0::int)  
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1548
      ==> (-a) div b =  
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1549
          (if a mod b = 0 then - (a div b) else  - (a div b) - 1)"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1550
by (blast intro: divmod_int_rel_div_mod [THEN zminus1_lemma, THEN divmod_int_rel_div])
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1551
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1552
lemma zmod_zminus1_eq_if:
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1553
     "(-a::int) mod b = (if a mod b = 0 then 0 else  b - (a mod b))"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1554
apply (case_tac "b = 0", simp)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1555
apply (blast intro: divmod_int_rel_div_mod [THEN zminus1_lemma, THEN divmod_int_rel_mod])
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1556
done
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1557
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1558
lemma zmod_zminus1_not_zero:
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1559
  fixes k l :: int
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1560
  shows "- k mod l \<noteq> 0 \<Longrightarrow> k mod l \<noteq> 0"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1561
  unfolding zmod_zminus1_eq_if by auto
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1562
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1563
lemma zdiv_zminus2: "a div (-b) = (-a::int) div b"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1564
by (cut_tac a = "-a" in zdiv_zminus_zminus, auto)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1565
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1566
lemma zmod_zminus2: "a mod (-b) = - ((-a::int) mod b)"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1567
by (cut_tac a = "-a" and b = b in zmod_zminus_zminus, auto)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1568
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1569
lemma zdiv_zminus2_eq_if:
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1570
     "b \<noteq> (0::int)  
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1571
      ==> a div (-b) =  
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1572
          (if a mod b = 0 then - (a div b) else  - (a div b) - 1)"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1573
by (simp add: zdiv_zminus1_eq_if zdiv_zminus2)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1574
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1575
lemma zmod_zminus2_eq_if:
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1576
     "a mod (-b::int) = (if a mod b = 0 then 0 else  (a mod b) - b)"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1577
by (simp add: zmod_zminus1_eq_if zmod_zminus2)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1578
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1579
lemma zmod_zminus2_not_zero:
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1580
  fixes k l :: int
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1581
  shows "k mod - l \<noteq> 0 \<Longrightarrow> k mod l \<noteq> 0"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1582
  unfolding zmod_zminus2_eq_if by auto 
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1583
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1584
46551
866bce5442a3 simplify projections on simultaneous computations of div and mod; tuned structure (from Florian Haftmann)
huffman
parents: 46026
diff changeset
  1585
subsubsection {* Division of a Number by Itself *}
33361
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1586
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1587
lemma self_quotient_aux1: "[| (0::int) < a; a = r + a*q; r < a |] ==> 1 \<le> q"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1588
apply (subgoal_tac "0 < a*q")
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1589
 apply (simp add: zero_less_mult_iff, arith)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1590
done
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1591
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1592
lemma self_quotient_aux2: "[| (0::int) < a; a = r + a*q; 0 \<le> r |] ==> q \<le> 1"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1593
apply (subgoal_tac "0 \<le> a* (1-q) ")
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1594
 apply (simp add: zero_le_mult_iff)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1595
apply (simp add: right_diff_distrib)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1596
done
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1597
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1598
lemma self_quotient: "[| divmod_int_rel a a (q, r);  a \<noteq> (0::int) |] ==> q = 1"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1599
apply (simp add: split_ifs divmod_int_rel_def linorder_neq_iff)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1600
apply (rule order_antisym, safe, simp_all)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1601
apply (rule_tac [3] a = "-a" and r = "-r" in self_quotient_aux1)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1602
apply (rule_tac a = "-a" and r = "-r" in self_quotient_aux2)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1603
apply (force intro: self_quotient_aux1 self_quotient_aux2 simp add: add_commute)+
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1604
done
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1605
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1606
lemma self_remainder: "[| divmod_int_rel a a (q, r);  a \<noteq> (0::int) |] ==> r = 0"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1607
apply (frule self_quotient, assumption)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1608
apply (simp add: divmod_int_rel_def)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1609
done
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1610
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1611
lemma zdiv_self [simp]: "a \<noteq> 0 ==> a div a = (1::int)"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1612
by (simp add: divmod_int_rel_div_mod [THEN self_quotient])
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1613
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1614
(*Here we have 0 mod 0 = 0, also assumed by Knuth (who puts m mod 0 = 0) *)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1615
lemma zmod_self [simp]: "a mod a = (0::int)"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1616
apply (case_tac "a = 0", simp)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1617
apply (simp add: divmod_int_rel_div_mod [THEN self_remainder])
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1618
done
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1619
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1620
46551
866bce5442a3 simplify projections on simultaneous computations of div and mod; tuned structure (from Florian Haftmann)
huffman
parents: 46026
diff changeset
  1621
subsubsection {* Computation of Division and Remainder *}
33361
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1622
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1623
lemma zdiv_zero [simp]: "(0::int) div b = 0"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1624
by (simp add: div_int_def divmod_int_def)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1625
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1626
lemma div_eq_minus1: "(0::int) < b ==> -1 div b = -1"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1627
by (simp add: div_int_def divmod_int_def)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1628
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1629
lemma zmod_zero [simp]: "(0::int) mod b = 0"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1630
by (simp add: mod_int_def divmod_int_def)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1631
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1632
lemma zmod_minus1: "(0::int) < b ==> -1 mod b = b - 1"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1633
by (simp add: mod_int_def divmod_int_def)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1634
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1635
text{*a positive, b positive *}
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1636
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1637
lemma div_pos_pos: "[| 0 < a;  0 \<le> b |] ==> a div b = fst (posDivAlg a b)"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1638
by (simp add: div_int_def divmod_int_def)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1639
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1640
lemma mod_pos_pos: "[| 0 < a;  0 \<le> b |] ==> a mod b = snd (posDivAlg a b)"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1641
by (simp add: mod_int_def divmod_int_def)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1642
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1643
text{*a negative, b positive *}
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1644
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1645
lemma div_neg_pos: "[| a < 0;  0 < b |] ==> a div b = fst (negDivAlg a b)"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1646
by (simp add: div_int_def divmod_int_def)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1647
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1648
lemma mod_neg_pos: "[| a < 0;  0 < b |] ==> a mod b = snd (negDivAlg a b)"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1649
by (simp add: mod_int_def divmod_int_def)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1650
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1651
text{*a positive, b negative *}
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1652
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1653
lemma div_pos_neg:
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1654
     "[| 0 < a;  b < 0 |] ==> a div b = fst (negateSnd (negDivAlg (-a) (-b)))"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1655
by (simp add: div_int_def divmod_int_def)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1656
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1657
lemma mod_pos_neg:
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1658
     "[| 0 < a;  b < 0 |] ==> a mod b = snd (negateSnd (negDivAlg (-a) (-b)))"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1659
by (simp add: mod_int_def divmod_int_def)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1660
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1661
text{*a negative, b negative *}
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1662
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1663
lemma div_neg_neg:
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1664
     "[| a < 0;  b \<le> 0 |] ==> a div b = fst (negateSnd (posDivAlg (-a) (-b)))"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1665
by (simp add: div_int_def divmod_int_def)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1666
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1667
lemma mod_neg_neg:
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1668
     "[| a < 0;  b \<le> 0 |] ==> a mod b = snd (negateSnd (posDivAlg (-a) (-b)))"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1669
by (simp add: mod_int_def divmod_int_def)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1670
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1671
text {*Simplify expresions in which div and mod combine numerical constants*}
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1672
45530
0c4853bb77bf rewrite integer numeral div/mod simprocs to not return conditional rewrites; add regression tests
huffman
parents: 45231
diff changeset
  1673
lemma int_div_pos_eq: "\<lbrakk>(a::int) = b * q + r; 0 \<le> r; r < b\<rbrakk> \<Longrightarrow> a div b = q"
0c4853bb77bf rewrite integer numeral div/mod simprocs to not return conditional rewrites; add regression tests
huffman
parents: 45231
diff changeset
  1674
  by (rule divmod_int_rel_div [of a b q r],
0c4853bb77bf rewrite integer numeral div/mod simprocs to not return conditional rewrites; add regression tests
huffman
parents: 45231
diff changeset
  1675
    simp add: divmod_int_rel_def, simp)
0c4853bb77bf rewrite integer numeral div/mod simprocs to not return conditional rewrites; add regression tests
huffman
parents: 45231
diff changeset
  1676
0c4853bb77bf rewrite integer numeral div/mod simprocs to not return conditional rewrites; add regression tests
huffman
parents: 45231
diff changeset
  1677
lemma int_div_neg_eq: "\<lbrakk>(a::int) = b * q + r; r \<le> 0; b < r\<rbrakk> \<Longrightarrow> a div b = q"
0c4853bb77bf rewrite integer numeral div/mod simprocs to not return conditional rewrites; add regression tests
huffman
parents: 45231
diff changeset
  1678
  by (rule divmod_int_rel_div [of a b q r],
0c4853bb77bf rewrite integer numeral div/mod simprocs to not return conditional rewrites; add regression tests
huffman
parents: 45231
diff changeset
  1679
    simp add: divmod_int_rel_def, simp)
0c4853bb77bf rewrite integer numeral div/mod simprocs to not return conditional rewrites; add regression tests
huffman
parents: 45231
diff changeset
  1680
0c4853bb77bf rewrite integer numeral div/mod simprocs to not return conditional rewrites; add regression tests
huffman
parents: 45231
diff changeset
  1681
lemma int_mod_pos_eq: "\<lbrakk>(a::int) = b * q + r; 0 \<le> r; r < b\<rbrakk> \<Longrightarrow> a mod b = r"
0c4853bb77bf rewrite integer numeral div/mod simprocs to not return conditional rewrites; add regression tests
huffman
parents: 45231
diff changeset
  1682
  by (rule divmod_int_rel_mod [of a b q r],
0c4853bb77bf rewrite integer numeral div/mod simprocs to not return conditional rewrites; add regression tests
huffman
parents: 45231
diff changeset
  1683
    simp add: divmod_int_rel_def, simp)
0c4853bb77bf rewrite integer numeral div/mod simprocs to not return conditional rewrites; add regression tests
huffman
parents: 45231
diff changeset
  1684
0c4853bb77bf rewrite integer numeral div/mod simprocs to not return conditional rewrites; add regression tests
huffman
parents: 45231
diff changeset
  1685
lemma int_mod_neg_eq: "\<lbrakk>(a::int) = b * q + r; r \<le> 0; b < r\<rbrakk> \<Longrightarrow> a mod b = r"
0c4853bb77bf rewrite integer numeral div/mod simprocs to not return conditional rewrites; add regression tests
huffman
parents: 45231
diff changeset
  1686
  by (rule divmod_int_rel_mod [of a b q r],
0c4853bb77bf rewrite integer numeral div/mod simprocs to not return conditional rewrites; add regression tests
huffman
parents: 45231
diff changeset
  1687
    simp add: divmod_int_rel_def, simp)
0c4853bb77bf rewrite integer numeral div/mod simprocs to not return conditional rewrites; add regression tests
huffman
parents: 45231
diff changeset
  1688
33361
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1689
lemmas arithmetic_simps =
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1690
  arith_simps
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1691
  add_special
35050
9f841f20dca6 renamed OrderedGroup to Groups; split theory Ring_and_Field into Rings Fields
haftmann
parents: 34982
diff changeset
  1692
  add_0_left
9f841f20dca6 renamed OrderedGroup to Groups; split theory Ring_and_Field into Rings Fields
haftmann
parents: 34982
diff changeset
  1693
  add_0_right
33361
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1694
  mult_zero_left
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1695
  mult_zero_right
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1696
  mult_1_left
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1697
  mult_1_right
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1698
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1699
(* simprocs adapted from HOL/ex/Binary.thy *)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1700
ML {*
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1701
local
45530
0c4853bb77bf rewrite integer numeral div/mod simprocs to not return conditional rewrites; add regression tests
huffman
parents: 45231
diff changeset
  1702
  val mk_number = HOLogic.mk_number HOLogic.intT
0c4853bb77bf rewrite integer numeral div/mod simprocs to not return conditional rewrites; add regression tests
huffman
parents: 45231
diff changeset
  1703
  val plus = @{term "plus :: int \<Rightarrow> int \<Rightarrow> int"}
0c4853bb77bf rewrite integer numeral div/mod simprocs to not return conditional rewrites; add regression tests
huffman
parents: 45231
diff changeset
  1704
  val times = @{term "times :: int \<Rightarrow> int \<Rightarrow> int"}
0c4853bb77bf rewrite integer numeral div/mod simprocs to not return conditional rewrites; add regression tests
huffman
parents: 45231
diff changeset
  1705
  val zero = @{term "0 :: int"}
0c4853bb77bf rewrite integer numeral div/mod simprocs to not return conditional rewrites; add regression tests
huffman
parents: 45231
diff changeset
  1706
  val less = @{term "op < :: int \<Rightarrow> int \<Rightarrow> bool"}
0c4853bb77bf rewrite integer numeral div/mod simprocs to not return conditional rewrites; add regression tests
huffman
parents: 45231
diff changeset
  1707
  val le = @{term "op \<le> :: int \<Rightarrow> int \<Rightarrow> bool"}
0c4853bb77bf rewrite integer numeral div/mod simprocs to not return conditional rewrites; add regression tests
huffman
parents: 45231
diff changeset
  1708
  val simps = @{thms arith_simps} @ @{thms rel_simps} @
0c4853bb77bf rewrite integer numeral div/mod simprocs to not return conditional rewrites; add regression tests
huffman
parents: 45231
diff changeset
  1709
    map (fn th => th RS sym) [@{thm numeral_0_eq_0}, @{thm numeral_1_eq_1}]
0c4853bb77bf rewrite integer numeral div/mod simprocs to not return conditional rewrites; add regression tests
huffman
parents: 45231
diff changeset
  1710
  fun prove ctxt goal = Goal.prove ctxt [] [] (HOLogic.mk_Trueprop goal)
0c4853bb77bf rewrite integer numeral div/mod simprocs to not return conditional rewrites; add regression tests
huffman
parents: 45231
diff changeset
  1711
    (K (ALLGOALS (full_simp_tac (HOL_basic_ss addsimps simps))));
33361
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1712
  fun binary_proc proc ss ct =
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1713
    (case Thm.term_of ct of
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1714
      _ $ t $ u =>
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1715
      (case try (pairself (`(snd o HOLogic.dest_number))) (t, u) of
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1716
        SOME args => proc (Simplifier.the_context ss) args
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1717
      | NONE => NONE)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1718
    | _ => NONE);
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1719
in
45530
0c4853bb77bf rewrite integer numeral div/mod simprocs to not return conditional rewrites; add regression tests
huffman
parents: 45231
diff changeset
  1720
  fun divmod_proc posrule negrule =
0c4853bb77bf rewrite integer numeral div/mod simprocs to not return conditional rewrites; add regression tests
huffman
parents: 45231
diff changeset
  1721
    binary_proc (fn ctxt => fn ((a, t), (b, u)) =>
0c4853bb77bf rewrite integer numeral div/mod simprocs to not return conditional rewrites; add regression tests
huffman
parents: 45231
diff changeset
  1722
      if b = 0 then NONE else let
0c4853bb77bf rewrite integer numeral div/mod simprocs to not return conditional rewrites; add regression tests
huffman
parents: 45231
diff changeset
  1723
        val (q, r) = pairself mk_number (Integer.div_mod a b)
0c4853bb77bf rewrite integer numeral div/mod simprocs to not return conditional rewrites; add regression tests
huffman
parents: 45231
diff changeset
  1724
        val goal1 = HOLogic.mk_eq (t, plus $ (times $ u $ q) $ r)
0c4853bb77bf rewrite integer numeral div/mod simprocs to not return conditional rewrites; add regression tests
huffman
parents: 45231
diff changeset
  1725
        val (goal2, goal3, rule) = if b > 0
0c4853bb77bf rewrite integer numeral div/mod simprocs to not return conditional rewrites; add regression tests
huffman
parents: 45231
diff changeset
  1726
          then (le $ zero $ r, less $ r $ u, posrule RS eq_reflection)
0c4853bb77bf rewrite integer numeral div/mod simprocs to not return conditional rewrites; add regression tests
huffman
parents: 45231
diff changeset
  1727
          else (le $ r $ zero, less $ u $ r, negrule RS eq_reflection)
0c4853bb77bf rewrite integer numeral div/mod simprocs to not return conditional rewrites; add regression tests
huffman
parents: 45231
diff changeset
  1728
      in SOME (rule OF map (prove ctxt) [goal1, goal2, goal3]) end)
33361
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1729
end
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1730
*}
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1731
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1732
simproc_setup binary_int_div ("number_of m div number_of n :: int") =
45530
0c4853bb77bf rewrite integer numeral div/mod simprocs to not return conditional rewrites; add regression tests
huffman
parents: 45231
diff changeset
  1733
  {* K (divmod_proc @{thm int_div_pos_eq} @{thm int_div_neg_eq}) *}
33361
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1734
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1735
simproc_setup binary_int_mod ("number_of m mod number_of n :: int") =
45530
0c4853bb77bf rewrite integer numeral div/mod simprocs to not return conditional rewrites; add regression tests
huffman
parents: 45231
diff changeset
  1736
  {* K (divmod_proc @{thm int_mod_pos_eq} @{thm int_mod_neg_eq}) *}
33361
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1737
45607
16b4f5774621 eliminated obsolete "standard";
wenzelm
parents: 45530
diff changeset
  1738
lemmas posDivAlg_eqn_number_of [simp] = posDivAlg_eqn [of "number_of v" "number_of w"] for v w
16b4f5774621 eliminated obsolete "standard";
wenzelm
parents: 45530
diff changeset
  1739
lemmas negDivAlg_eqn_number_of [simp] = negDivAlg_eqn [of "number_of v" "number_of w"] for v w
33361
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1740
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1741
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1742
text{*Special-case simplification *}
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1743
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1744
lemma zmod_minus1_right [simp]: "a mod (-1::int) = 0"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1745
apply (cut_tac a = a and b = "-1" in neg_mod_sign)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1746
apply (cut_tac [2] a = a and b = "-1" in neg_mod_bound)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1747
apply (auto simp del: neg_mod_sign neg_mod_bound)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1748
done
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1749
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1750
lemma zdiv_minus1_right [simp]: "a div (-1::int) = -a"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1751
by (cut_tac a = a and b = "-1" in zmod_zdiv_equality, auto)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1752
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1753
(** The last remaining special cases for constant arithmetic:
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1754
    1 div z and 1 mod z **)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1755
45607
16b4f5774621 eliminated obsolete "standard";
wenzelm
parents: 45530
diff changeset
  1756
lemmas div_pos_pos_1_number_of [simp] = div_pos_pos [OF zero_less_one, of "number_of w"] for w
16b4f5774621 eliminated obsolete "standard";
wenzelm
parents: 45530
diff changeset
  1757
lemmas div_pos_neg_1_number_of [simp] = div_pos_neg [OF zero_less_one, of "number_of w"] for w
16b4f5774621 eliminated obsolete "standard";
wenzelm
parents: 45530
diff changeset
  1758
lemmas mod_pos_pos_1_number_of [simp] = mod_pos_pos [OF zero_less_one, of "number_of w"] for w
16b4f5774621 eliminated obsolete "standard";
wenzelm
parents: 45530
diff changeset
  1759
lemmas mod_pos_neg_1_number_of [simp] = mod_pos_neg [OF zero_less_one, of "number_of w"] for w
16b4f5774621 eliminated obsolete "standard";
wenzelm
parents: 45530
diff changeset
  1760
lemmas posDivAlg_eqn_1_number_of [simp] = posDivAlg_eqn [of concl: 1 "number_of w"] for w
16b4f5774621 eliminated obsolete "standard";
wenzelm
parents: 45530
diff changeset
  1761
lemmas negDivAlg_eqn_1_number_of [simp] = negDivAlg_eqn [of concl: 1 "number_of w"] for w
33361
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1762
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1763
46551
866bce5442a3 simplify projections on simultaneous computations of div and mod; tuned structure (from Florian Haftmann)
huffman
parents: 46026
diff changeset
  1764
subsubsection {* Monotonicity in the First Argument (Dividend) *}
33361
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1765
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1766
lemma zdiv_mono1: "[| a \<le> a';  0 < (b::int) |] ==> a div b \<le> a' div b"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1767
apply (cut_tac a = a and b = b in zmod_zdiv_equality)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1768
apply (cut_tac a = a' and b = b in zmod_zdiv_equality)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1769
apply (rule unique_quotient_lemma)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1770
apply (erule subst)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1771
apply (erule subst, simp_all)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1772
done
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1773
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1774
lemma zdiv_mono1_neg: "[| a \<le> a';  (b::int) < 0 |] ==> a' div b \<le> a div b"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1775
apply (cut_tac a = a and b = b in zmod_zdiv_equality)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1776
apply (cut_tac a = a' and b = b in zmod_zdiv_equality)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1777
apply (rule unique_quotient_lemma_neg)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1778
apply (erule subst)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1779
apply (erule subst, simp_all)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1780
done
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1781
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1782
46551
866bce5442a3 simplify projections on simultaneous computations of div and mod; tuned structure (from Florian Haftmann)
huffman
parents: 46026
diff changeset
  1783
subsubsection {* Monotonicity in the Second Argument (Divisor) *}
33361
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1784
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1785
lemma q_pos_lemma:
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1786
     "[| 0 \<le> b'*q' + r'; r' < b';  0 < b' |] ==> 0 \<le> (q'::int)"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1787
apply (subgoal_tac "0 < b'* (q' + 1) ")
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1788
 apply (simp add: zero_less_mult_iff)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1789
apply (simp add: right_distrib)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1790
done
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1791
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1792
lemma zdiv_mono2_lemma:
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1793
     "[| b*q + r = b'*q' + r';  0 \<le> b'*q' + r';   
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1794
         r' < b';  0 \<le> r;  0 < b';  b' \<le> b |]   
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1795
      ==> q \<le> (q'::int)"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1796
apply (frule q_pos_lemma, assumption+) 
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1797
apply (subgoal_tac "b*q < b* (q' + 1) ")
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1798
 apply (simp add: mult_less_cancel_left)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1799
apply (subgoal_tac "b*q = r' - r + b'*q'")
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1800
 prefer 2 apply simp
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1801
apply (simp (no_asm_simp) add: right_distrib)
44766
d4d33a4d7548 avoid using legacy theorem names
huffman
parents: 43594
diff changeset
  1802
apply (subst add_commute, rule add_less_le_mono, arith)
33361
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1803
apply (rule mult_right_mono, auto)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1804
done
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1805
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1806
lemma zdiv_mono2:
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1807
     "[| (0::int) \<le> a;  0 < b';  b' \<le> b |] ==> a div b \<le> a div b'"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1808
apply (subgoal_tac "b \<noteq> 0")
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1809
 prefer 2 apply arith
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1810
apply (cut_tac a = a and b = b in zmod_zdiv_equality)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1811
apply (cut_tac a = a and b = b' in zmod_zdiv_equality)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1812
apply (rule zdiv_mono2_lemma)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1813
apply (erule subst)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1814
apply (erule subst, simp_all)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1815
done
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1816
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1817
lemma q_neg_lemma:
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1818
     "[| b'*q' + r' < 0;  0 \<le> r';  0 < b' |] ==> q' \<le> (0::int)"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1819
apply (subgoal_tac "b'*q' < 0")
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1820
 apply (simp add: mult_less_0_iff, arith)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1821
done
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1822
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1823
lemma zdiv_mono2_neg_lemma:
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1824
     "[| b*q + r = b'*q' + r';  b'*q' + r' < 0;   
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1825
         r < b;  0 \<le> r';  0 < b';  b' \<le> b |]   
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1826
      ==> q' \<le> (q::int)"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1827
apply (frule q_neg_lemma, assumption+) 
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1828
apply (subgoal_tac "b*q' < b* (q + 1) ")
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1829
 apply (simp add: mult_less_cancel_left)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1830
apply (simp add: right_distrib)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1831
apply (subgoal_tac "b*q' \<le> b'*q'")
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1832
 prefer 2 apply (simp add: mult_right_mono_neg, arith)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1833
done
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1834
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1835
lemma zdiv_mono2_neg:
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1836
     "[| a < (0::int);  0 < b';  b' \<le> b |] ==> a div b' \<le> a div b"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1837
apply (cut_tac a = a and b = b in zmod_zdiv_equality)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1838
apply (cut_tac a = a and b = b' in zmod_zdiv_equality)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1839
apply (rule zdiv_mono2_neg_lemma)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1840
apply (erule subst)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1841
apply (erule subst, simp_all)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1842
done
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1843
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1844
46551
866bce5442a3 simplify projections on simultaneous computations of div and mod; tuned structure (from Florian Haftmann)
huffman
parents: 46026
diff changeset
  1845
subsubsection {* More Algebraic Laws for div and mod *}
33361
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1846
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1847
text{*proving (a*b) div c = a * (b div c) + a * (b mod c) *}
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1848
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1849
lemma zmult1_lemma:
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1850
     "[| divmod_int_rel b c (q, r);  c \<noteq> 0 |]  
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1851
      ==> divmod_int_rel (a * b) c (a*q + a*r div c, a*r mod c)"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1852
by (auto simp add: split_ifs divmod_int_rel_def linorder_neq_iff right_distrib mult_ac)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1853
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1854
lemma zdiv_zmult1_eq: "(a*b) div c = a*(b div c) + a*(b mod c) div (c::int)"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1855
apply (case_tac "c = 0", simp)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1856
apply (blast intro: divmod_int_rel_div_mod [THEN zmult1_lemma, THEN divmod_int_rel_div])
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1857
done
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1858
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1859
lemma zmod_zmult1_eq: "(a*b) mod c = a*(b mod c) mod (c::int)"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1860
apply (case_tac "c = 0", simp)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1861
apply (blast intro: divmod_int_rel_div_mod [THEN zmult1_lemma, THEN divmod_int_rel_mod])
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1862
done
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1863
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1864
lemma zmod_zdiv_trivial: "(a mod b) div b = (0::int)"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1865
apply (case_tac "b = 0", simp)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1866
apply (auto simp add: linorder_neq_iff div_pos_pos_trivial div_neg_neg_trivial)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1867
done
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1868
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1869
text{*proving (a+b) div c = a div c + b div c + ((a mod c + b mod c) div c) *}
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1870
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1871
lemma zadd1_lemma:
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1872
     "[| divmod_int_rel a c (aq, ar);  divmod_int_rel b c (bq, br);  c \<noteq> 0 |]  
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1873
      ==> divmod_int_rel (a+b) c (aq + bq + (ar+br) div c, (ar+br) mod c)"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1874
by (force simp add: split_ifs divmod_int_rel_def linorder_neq_iff right_distrib)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1875
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1876
(*NOT suitable for rewriting: the RHS has an instance of the LHS*)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1877
lemma zdiv_zadd1_eq:
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1878
     "(a+b) div (c::int) = a div c + b div c + ((a mod c + b mod c) div c)"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1879
apply (case_tac "c = 0", simp)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1880
apply (blast intro: zadd1_lemma [OF divmod_int_rel_div_mod divmod_int_rel_div_mod] divmod_int_rel_div)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1881
done
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1882
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1883
instance int :: ring_div
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1884
proof
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1885
  fix a b c :: int
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1886
  assume not0: "b \<noteq> 0"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1887
  show "(a + c * b) div b = c + a div b"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1888
    unfolding zdiv_zadd1_eq [of a "c * b"] using not0 
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1889
      by (simp add: zmod_zmult1_eq zmod_zdiv_trivial zdiv_zmult1_eq)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1890
next
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1891
  fix a b c :: int
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1892
  assume "a \<noteq> 0"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1893
  then show "(a * b) div (a * c) = b div c"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1894
  proof (cases "b \<noteq> 0 \<and> c \<noteq> 0")
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1895
    case False then show ?thesis by auto
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1896
  next
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1897
    case True then have "b \<noteq> 0" and "c \<noteq> 0" by auto
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1898
    with `a \<noteq> 0`
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1899
    have "\<And>q r. divmod_int_rel b c (q, r) \<Longrightarrow> divmod_int_rel (a * b) (a * c) (q, a * r)"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1900
      apply (auto simp add: divmod_int_rel_def) 
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1901
      apply (auto simp add: algebra_simps)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1902
      apply (auto simp add: zero_less_mult_iff zero_le_mult_iff mult_le_0_iff mult_commute [of a] mult_less_cancel_right)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1903
      done
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1904
    moreover with `c \<noteq> 0` divmod_int_rel_div_mod have "divmod_int_rel b c (b div c, b mod c)" by auto
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1905
    ultimately have "divmod_int_rel (a * b) (a * c) (b div c, a * (b mod c))" .
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1906
    moreover from  `a \<noteq> 0` `c \<noteq> 0` have "a * c \<noteq> 0" by simp
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1907
    ultimately show ?thesis by (rule divmod_int_rel_div)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1908
  qed
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1909
qed auto
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1910
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1911
lemma posDivAlg_div_mod:
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1912
  assumes "k \<ge> 0"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1913
  and "l \<ge> 0"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1914
  shows "posDivAlg k l = (k div l, k mod l)"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1915
proof (cases "l = 0")
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1916
  case True then show ?thesis by (simp add: posDivAlg.simps)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1917
next
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1918
  case False with assms posDivAlg_correct
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1919
    have "divmod_int_rel k l (fst (posDivAlg k l), snd (posDivAlg k l))"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1920
    by simp
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1921
  from divmod_int_rel_div [OF this `l \<noteq> 0`] divmod_int_rel_mod [OF this `l \<noteq> 0`]
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1922
  show ?thesis by simp
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1923
qed
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1924
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1925
lemma negDivAlg_div_mod:
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1926
  assumes "k < 0"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1927
  and "l > 0"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1928
  shows "negDivAlg k l = (k div l, k mod l)"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1929
proof -
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1930
  from assms have "l \<noteq> 0" by simp
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1931
  from assms negDivAlg_correct
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1932
    have "divmod_int_rel k l (fst (negDivAlg k l), snd (negDivAlg k l))"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1933
    by simp
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1934
  from divmod_int_rel_div [OF this `l \<noteq> 0`] divmod_int_rel_mod [OF this `l \<noteq> 0`]
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1935
  show ?thesis by simp
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1936
qed
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1937
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1938
lemma zmod_eq_0_iff: "(m mod d = 0) = (EX q::int. m = d*q)"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1939
by (simp add: dvd_eq_mod_eq_0 [symmetric] dvd_def)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1940
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1941
(* REVISIT: should this be generalized to all semiring_div types? *)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1942
lemmas zmod_eq_0D [dest!] = zmod_eq_0_iff [THEN iffD1]
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1943
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1944
46551
866bce5442a3 simplify projections on simultaneous computations of div and mod; tuned structure (from Florian Haftmann)
huffman
parents: 46026
diff changeset
  1945
subsubsection {* Proving  @{term "a div (b*c) = (a div b) div c"} *}
33361
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1946
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1947
(*The condition c>0 seems necessary.  Consider that 7 div ~6 = ~2 but
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1948
  7 div 2 div ~3 = 3 div ~3 = ~1.  The subcase (a div b) mod c = 0 seems
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1949
  to cause particular problems.*)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1950
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1951
text{*first, four lemmas to bound the remainder for the cases b<0 and b>0 *}
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1952
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1953
lemma zmult2_lemma_aux1: "[| (0::int) < c;  b < r;  r \<le> 0 |] ==> b*c < b*(q mod c) + r"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1954
apply (subgoal_tac "b * (c - q mod c) < r * 1")
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1955
 apply (simp add: algebra_simps)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1956
apply (rule order_le_less_trans)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1957
 apply (erule_tac [2] mult_strict_right_mono)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1958
 apply (rule mult_left_mono_neg)
35216
7641e8d831d2 get rid of many duplicate simp rule warnings
huffman
parents: 35050
diff changeset
  1959
  using add1_zle_eq[of "q mod c"]apply(simp add: algebra_simps)
33361
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1960
 apply (simp)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1961
apply (simp)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1962
done
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1963
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1964
lemma zmult2_lemma_aux2:
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1965
     "[| (0::int) < c;   b < r;  r \<le> 0 |] ==> b * (q mod c) + r \<le> 0"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1966
apply (subgoal_tac "b * (q mod c) \<le> 0")
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1967
 apply arith
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1968
apply (simp add: mult_le_0_iff)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1969
done
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1970
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1971
lemma zmult2_lemma_aux3: "[| (0::int) < c;  0 \<le> r;  r < b |] ==> 0 \<le> b * (q mod c) + r"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1972
apply (subgoal_tac "0 \<le> b * (q mod c) ")
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1973
apply arith
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1974
apply (simp add: zero_le_mult_iff)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1975
done
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1976
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1977
lemma zmult2_lemma_aux4: "[| (0::int) < c; 0 \<le> r; r < b |] ==> b * (q mod c) + r < b * c"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1978
apply (subgoal_tac "r * 1 < b * (c - q mod c) ")
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1979
 apply (simp add: right_diff_distrib)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1980
apply (rule order_less_le_trans)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1981
 apply (erule mult_strict_right_mono)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1982
 apply (rule_tac [2] mult_left_mono)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1983
  apply simp
35216
7641e8d831d2 get rid of many duplicate simp rule warnings
huffman
parents: 35050
diff changeset
  1984
 using add1_zle_eq[of "q mod c"] apply (simp add: algebra_simps)
33361
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1985
apply simp
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1986
done
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1987
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1988
lemma zmult2_lemma: "[| divmod_int_rel a b (q, r);  b \<noteq> 0;  0 < c |]  
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1989
      ==> divmod_int_rel a (b * c) (q div c, b*(q mod c) + r)"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1990
by (auto simp add: mult_ac divmod_int_rel_def linorder_neq_iff
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1991
                   zero_less_mult_iff right_distrib [symmetric] 
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1992
                   zmult2_lemma_aux1 zmult2_lemma_aux2 zmult2_lemma_aux3 zmult2_lemma_aux4)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1993
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1994
lemma zdiv_zmult2_eq: "(0::int) < c ==> a div (b*c) = (a div b) div c"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1995
apply (case_tac "b = 0", simp)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1996
apply (force simp add: divmod_int_rel_div_mod [THEN zmult2_lemma, THEN divmod_int_rel_div])
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1997
done
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1998
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  1999
lemma zmod_zmult2_eq:
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2000
     "(0::int) < c ==> a mod (b*c) = b*(a div b mod c) + a mod b"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2001
apply (case_tac "b = 0", simp)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2002
apply (force simp add: divmod_int_rel_div_mod [THEN zmult2_lemma, THEN divmod_int_rel_mod])
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2003
done
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2004
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2005
46551
866bce5442a3 simplify projections on simultaneous computations of div and mod; tuned structure (from Florian Haftmann)
huffman
parents: 46026
diff changeset
  2006
subsubsection {* Splitting Rules for div and mod *}
33361
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2007
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2008
text{*The proofs of the two lemmas below are essentially identical*}
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2009
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2010
lemma split_pos_lemma:
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2011
 "0<k ==> 
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2012
    P(n div k :: int)(n mod k) = (\<forall>i j. 0\<le>j & j<k & n = k*i + j --> P i j)"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2013
apply (rule iffI, clarify)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2014
 apply (erule_tac P="P ?x ?y" in rev_mp)  
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2015
 apply (subst mod_add_eq) 
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2016
 apply (subst zdiv_zadd1_eq) 
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2017
 apply (simp add: div_pos_pos_trivial mod_pos_pos_trivial)  
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2018
txt{*converse direction*}
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2019
apply (drule_tac x = "n div k" in spec) 
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2020
apply (drule_tac x = "n mod k" in spec, simp)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2021
done
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2022
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2023
lemma split_neg_lemma:
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2024
 "k<0 ==>
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2025
    P(n div k :: int)(n mod k) = (\<forall>i j. k<j & j\<le>0 & n = k*i + j --> P i j)"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2026
apply (rule iffI, clarify)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2027
 apply (erule_tac P="P ?x ?y" in rev_mp)  
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2028
 apply (subst mod_add_eq) 
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2029
 apply (subst zdiv_zadd1_eq) 
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2030
 apply (simp add: div_neg_neg_trivial mod_neg_neg_trivial)  
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2031
txt{*converse direction*}
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2032
apply (drule_tac x = "n div k" in spec) 
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2033
apply (drule_tac x = "n mod k" in spec, simp)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2034
done
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2035
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2036
lemma split_zdiv:
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2037
 "P(n div k :: int) =
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2038
  ((k = 0 --> P 0) & 
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2039
   (0<k --> (\<forall>i j. 0\<le>j & j<k & n = k*i + j --> P i)) & 
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2040
   (k<0 --> (\<forall>i j. k<j & j\<le>0 & n = k*i + j --> P i)))"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2041
apply (case_tac "k=0", simp)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2042
apply (simp only: linorder_neq_iff)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2043
apply (erule disjE) 
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2044
 apply (simp_all add: split_pos_lemma [of concl: "%x y. P x"] 
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2045
                      split_neg_lemma [of concl: "%x y. P x"])
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2046
done
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2047
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2048
lemma split_zmod:
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2049
 "P(n mod k :: int) =
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2050
  ((k = 0 --> P n) & 
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2051
   (0<k --> (\<forall>i j. 0\<le>j & j<k & n = k*i + j --> P j)) & 
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2052
   (k<0 --> (\<forall>i j. k<j & j\<le>0 & n = k*i + j --> P j)))"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2053
apply (case_tac "k=0", simp)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2054
apply (simp only: linorder_neq_iff)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2055
apply (erule disjE) 
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2056
 apply (simp_all add: split_pos_lemma [of concl: "%x y. P y"] 
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2057
                      split_neg_lemma [of concl: "%x y. P y"])
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2058
done
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2059
33730
1755ca4ec022 Fixed splitting of div and mod on integers (split theorem differed from implementation).
webertj
parents: 33728
diff changeset
  2060
text {* Enable (lin)arith to deal with @{const div} and @{const mod}
1755ca4ec022 Fixed splitting of div and mod on integers (split theorem differed from implementation).
webertj
parents: 33728
diff changeset
  2061
  when these are applied to some constant that is of the form
1755ca4ec022 Fixed splitting of div and mod on integers (split theorem differed from implementation).
webertj
parents: 33728
diff changeset
  2062
  @{term "number_of k"}: *}
45607
16b4f5774621 eliminated obsolete "standard";
wenzelm
parents: 45530
diff changeset
  2063
declare split_zdiv [of _ _ "number_of k", arith_split] for k
16b4f5774621 eliminated obsolete "standard";
wenzelm
parents: 45530
diff changeset
  2064
declare split_zmod [of _ _ "number_of k", arith_split] for k
33361
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2065
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2066
46551
866bce5442a3 simplify projections on simultaneous computations of div and mod; tuned structure (from Florian Haftmann)
huffman
parents: 46026
diff changeset
  2067
subsubsection {* Speeding up the Division Algorithm with Shifting *}
33361
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2068
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2069
text{*computing div by shifting *}
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2070
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2071
lemma pos_zdiv_mult_2: "(0::int) \<le> a ==> (1 + 2*b) div (2*a) = b div a"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2072
proof cases
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2073
  assume "a=0"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2074
    thus ?thesis by simp
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2075
next
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2076
  assume "a\<noteq>0" and le_a: "0\<le>a"   
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2077
  hence a_pos: "1 \<le> a" by arith
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2078
  hence one_less_a2: "1 < 2 * a" by arith
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2079
  hence le_2a: "2 * (1 + b mod a) \<le> 2 * a"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2080
    unfolding mult_le_cancel_left
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2081
    by (simp add: add1_zle_eq add_commute [of 1])
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2082
  with a_pos have "0 \<le> b mod a" by simp
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2083
  hence le_addm: "0 \<le> 1 mod (2*a) + 2*(b mod a)"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2084
    by (simp add: mod_pos_pos_trivial one_less_a2)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2085
  with  le_2a
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2086
  have "(1 mod (2*a) + 2*(b mod a)) div (2*a) = 0"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2087
    by (simp add: div_pos_pos_trivial le_addm mod_pos_pos_trivial one_less_a2
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2088
                  right_distrib) 
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2089
  thus ?thesis
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2090
    by (subst zdiv_zadd1_eq,
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2091
        simp add: mod_mult_mult1 one_less_a2
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2092
                  div_pos_pos_trivial)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2093
qed
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2094
35815
10e723e54076 tuned proofs (to avoid linarith error message caused by bootstrapping of HOL)
boehmes
parents: 35673
diff changeset
  2095
lemma neg_zdiv_mult_2: 
10e723e54076 tuned proofs (to avoid linarith error message caused by bootstrapping of HOL)
boehmes
parents: 35673
diff changeset
  2096
  assumes A: "a \<le> (0::int)" shows "(1 + 2*b) div (2*a) = (b+1) div a"
10e723e54076 tuned proofs (to avoid linarith error message caused by bootstrapping of HOL)
boehmes
parents: 35673
diff changeset
  2097
proof -
10e723e54076 tuned proofs (to avoid linarith error message caused by bootstrapping of HOL)
boehmes
parents: 35673
diff changeset
  2098
  have R: "1 + - (2 * (b + 1)) = - (1 + 2 * b)" by simp
10e723e54076 tuned proofs (to avoid linarith error message caused by bootstrapping of HOL)
boehmes
parents: 35673
diff changeset
  2099
  have "(1 + 2 * (-b - 1)) div (2 * (-a)) = (-b - 1) div (-a)"
10e723e54076 tuned proofs (to avoid linarith error message caused by bootstrapping of HOL)
boehmes
parents: 35673
diff changeset
  2100
    by (rule pos_zdiv_mult_2, simp add: A)
10e723e54076 tuned proofs (to avoid linarith error message caused by bootstrapping of HOL)
boehmes
parents: 35673
diff changeset
  2101
  thus ?thesis
10e723e54076 tuned proofs (to avoid linarith error message caused by bootstrapping of HOL)
boehmes
parents: 35673
diff changeset
  2102
    by (simp only: R zdiv_zminus_zminus diff_minus
10e723e54076 tuned proofs (to avoid linarith error message caused by bootstrapping of HOL)
boehmes
parents: 35673
diff changeset
  2103
      minus_add_distrib [symmetric] mult_minus_right)
10e723e54076 tuned proofs (to avoid linarith error message caused by bootstrapping of HOL)
boehmes
parents: 35673
diff changeset
  2104
qed
33361
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2105
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2106
lemma zdiv_number_of_Bit0 [simp]:
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2107
     "number_of (Int.Bit0 v) div number_of (Int.Bit0 w) =  
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2108
          number_of v div (number_of w :: int)"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2109
by (simp only: number_of_eq numeral_simps) (simp add: mult_2 [symmetric])
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2110
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2111
lemma zdiv_number_of_Bit1 [simp]:
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2112
     "number_of (Int.Bit1 v) div number_of (Int.Bit0 w) =  
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2113
          (if (0::int) \<le> number_of w                    
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2114
           then number_of v div (number_of w)     
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2115
           else (number_of v + (1::int)) div (number_of w))"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2116
apply (simp only: number_of_eq numeral_simps UNIV_I split: split_if) 
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2117
apply (simp add: pos_zdiv_mult_2 neg_zdiv_mult_2 add_ac mult_2 [symmetric])
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2118
done
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2119
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2120
46551
866bce5442a3 simplify projections on simultaneous computations of div and mod; tuned structure (from Florian Haftmann)
huffman
parents: 46026
diff changeset
  2121
subsubsection {* Computing mod by Shifting (proofs resemble those for div) *}
33361
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2122
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2123
lemma pos_zmod_mult_2:
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2124
  fixes a b :: int
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2125
  assumes "0 \<le> a"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2126
  shows "(1 + 2 * b) mod (2 * a) = 1 + 2 * (b mod a)"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2127
proof (cases "0 < a")
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2128
  case False with assms show ?thesis by simp
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2129
next
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2130
  case True
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2131
  then have "b mod a < a" by (rule pos_mod_bound)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2132
  then have "1 + b mod a \<le> a" by simp
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2133
  then have A: "2 * (1 + b mod a) \<le> 2 * a" by simp
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2134
  from `0 < a` have "0 \<le> b mod a" by (rule pos_mod_sign)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2135
  then have B: "0 \<le> 1 + 2 * (b mod a)" by simp
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2136
  have "((1\<Colon>int) mod ((2\<Colon>int) * a) + (2\<Colon>int) * b mod ((2\<Colon>int) * a)) mod ((2\<Colon>int) * a) = (1\<Colon>int) + (2\<Colon>int) * (b mod a)"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2137
    using `0 < a` and A
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2138
    by (auto simp add: mod_mult_mult1 mod_pos_pos_trivial ring_distribs intro!: mod_pos_pos_trivial B)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2139
  then show ?thesis by (subst mod_add_eq)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2140
qed
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2141
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2142
lemma neg_zmod_mult_2:
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2143
  fixes a b :: int
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2144
  assumes "a \<le> 0"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2145
  shows "(1 + 2 * b) mod (2 * a) = 2 * ((b + 1) mod a) - 1"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2146
proof -
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2147
  from assms have "0 \<le> - a" by auto
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2148
  then have "(1 + 2 * (- b - 1)) mod (2 * (- a)) = 1 + 2 * ((- b - 1) mod (- a))"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2149
    by (rule pos_zmod_mult_2)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2150
  then show ?thesis by (simp add: zmod_zminus2 algebra_simps)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2151
     (simp add: diff_minus add_ac)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2152
qed
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2153
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2154
lemma zmod_number_of_Bit0 [simp]:
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2155
     "number_of (Int.Bit0 v) mod number_of (Int.Bit0 w) =  
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2156
      (2::int) * (number_of v mod number_of w)"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2157
apply (simp only: number_of_eq numeral_simps) 
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2158
apply (simp add: mod_mult_mult1 pos_zmod_mult_2 
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2159
                 neg_zmod_mult_2 add_ac mult_2 [symmetric])
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2160
done
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2161
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2162
lemma zmod_number_of_Bit1 [simp]:
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2163
     "number_of (Int.Bit1 v) mod number_of (Int.Bit0 w) =  
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2164
      (if (0::int) \<le> number_of w  
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2165
                then 2 * (number_of v mod number_of w) + 1     
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2166
                else 2 * ((number_of v + (1::int)) mod number_of w) - 1)"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2167
apply (simp only: number_of_eq numeral_simps) 
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2168
apply (simp add: mod_mult_mult1 pos_zmod_mult_2 
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2169
                 neg_zmod_mult_2 add_ac mult_2 [symmetric])
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2170
done
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2171
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2172
39489
8bb7f32a3a08 added lemmas
nipkow
parents: 38715
diff changeset
  2173
lemma zdiv_eq_0_iff:
8bb7f32a3a08 added lemmas
nipkow
parents: 38715
diff changeset
  2174
 "(i::int) div k = 0 \<longleftrightarrow> k=0 \<or> 0\<le>i \<and> i<k \<or> i\<le>0 \<and> k<i" (is "?L = ?R")
8bb7f32a3a08 added lemmas
nipkow
parents: 38715
diff changeset
  2175
proof
8bb7f32a3a08 added lemmas
nipkow
parents: 38715
diff changeset
  2176
  assume ?L
8bb7f32a3a08 added lemmas
nipkow
parents: 38715
diff changeset
  2177
  have "?L \<longrightarrow> ?R" by (rule split_zdiv[THEN iffD2]) simp
8bb7f32a3a08 added lemmas
nipkow
parents: 38715
diff changeset
  2178
  with `?L` show ?R by blast
8bb7f32a3a08 added lemmas
nipkow
parents: 38715
diff changeset
  2179
next
8bb7f32a3a08 added lemmas
nipkow
parents: 38715
diff changeset
  2180
  assume ?R thus ?L
8bb7f32a3a08 added lemmas
nipkow
parents: 38715
diff changeset
  2181
    by(auto simp: div_pos_pos_trivial div_neg_neg_trivial)
8bb7f32a3a08 added lemmas
nipkow
parents: 38715
diff changeset
  2182
qed
8bb7f32a3a08 added lemmas
nipkow
parents: 38715
diff changeset
  2183
8bb7f32a3a08 added lemmas
nipkow
parents: 38715
diff changeset
  2184
46551
866bce5442a3 simplify projections on simultaneous computations of div and mod; tuned structure (from Florian Haftmann)
huffman
parents: 46026
diff changeset
  2185
subsubsection {* Quotients of Signs *}
33361
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2186
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2187
lemma div_neg_pos_less0: "[| a < (0::int);  0 < b |] ==> a div b < 0"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2188
apply (subgoal_tac "a div b \<le> -1", force)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2189
apply (rule order_trans)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2190
apply (rule_tac a' = "-1" in zdiv_mono1)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2191
apply (auto simp add: div_eq_minus1)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2192
done
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2193
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2194
lemma div_nonneg_neg_le0: "[| (0::int) \<le> a; b < 0 |] ==> a div b \<le> 0"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2195
by (drule zdiv_mono1_neg, auto)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2196
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2197
lemma div_nonpos_pos_le0: "[| (a::int) \<le> 0; b > 0 |] ==> a div b \<le> 0"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2198
by (drule zdiv_mono1, auto)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2199
33804
39b494e8c055 added lemma
nipkow
parents: 33730
diff changeset
  2200
text{* Now for some equivalences of the form @{text"a div b >=< 0 \<longleftrightarrow> \<dots>"}
39b494e8c055 added lemma
nipkow
parents: 33730
diff changeset
  2201
conditional upon the sign of @{text a} or @{text b}. There are many more.
39b494e8c055 added lemma
nipkow
parents: 33730
diff changeset
  2202
They should all be simp rules unless that causes too much search. *}
39b494e8c055 added lemma
nipkow
parents: 33730
diff changeset
  2203
33361
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2204
lemma pos_imp_zdiv_nonneg_iff: "(0::int) < b ==> (0 \<le> a div b) = (0 \<le> a)"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2205
apply auto
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2206
apply (drule_tac [2] zdiv_mono1)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2207
apply (auto simp add: linorder_neq_iff)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2208
apply (simp (no_asm_use) add: linorder_not_less [symmetric])
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2209
apply (blast intro: div_neg_pos_less0)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2210
done
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2211
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2212
lemma neg_imp_zdiv_nonneg_iff:
33804
39b494e8c055 added lemma
nipkow
parents: 33730
diff changeset
  2213
  "b < (0::int) ==> (0 \<le> a div b) = (a \<le> (0::int))"
33361
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2214
apply (subst zdiv_zminus_zminus [symmetric])
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2215
apply (subst pos_imp_zdiv_nonneg_iff, auto)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2216
done
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2217
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2218
(*But not (a div b \<le> 0 iff a\<le>0); consider a=1, b=2 when a div b = 0.*)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2219
lemma pos_imp_zdiv_neg_iff: "(0::int) < b ==> (a div b < 0) = (a < 0)"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2220
by (simp add: linorder_not_le [symmetric] pos_imp_zdiv_nonneg_iff)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2221
39489
8bb7f32a3a08 added lemmas
nipkow
parents: 38715
diff changeset
  2222
lemma pos_imp_zdiv_pos_iff:
8bb7f32a3a08 added lemmas
nipkow
parents: 38715
diff changeset
  2223
  "0<k \<Longrightarrow> 0 < (i::int) div k \<longleftrightarrow> k \<le> i"
8bb7f32a3a08 added lemmas
nipkow
parents: 38715
diff changeset
  2224
using pos_imp_zdiv_nonneg_iff[of k i] zdiv_eq_0_iff[of i k]
8bb7f32a3a08 added lemmas
nipkow
parents: 38715
diff changeset
  2225
by arith
8bb7f32a3a08 added lemmas
nipkow
parents: 38715
diff changeset
  2226
33361
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2227
(*Again the law fails for \<le>: consider a = -1, b = -2 when a div b = 0*)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2228
lemma neg_imp_zdiv_neg_iff: "b < (0::int) ==> (a div b < 0) = (0 < a)"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2229
by (simp add: linorder_not_le [symmetric] neg_imp_zdiv_nonneg_iff)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2230
33804
39b494e8c055 added lemma
nipkow
parents: 33730
diff changeset
  2231
lemma nonneg1_imp_zdiv_pos_iff:
39b494e8c055 added lemma
nipkow
parents: 33730
diff changeset
  2232
  "(0::int) <= a \<Longrightarrow> (a div b > 0) = (a >= b & b>0)"
39b494e8c055 added lemma
nipkow
parents: 33730
diff changeset
  2233
apply rule
39b494e8c055 added lemma
nipkow
parents: 33730
diff changeset
  2234
 apply rule
39b494e8c055 added lemma
nipkow
parents: 33730
diff changeset
  2235
  using div_pos_pos_trivial[of a b]apply arith
39b494e8c055 added lemma
nipkow
parents: 33730
diff changeset
  2236
 apply(cases "b=0")apply simp
39b494e8c055 added lemma
nipkow
parents: 33730
diff changeset
  2237
 using div_nonneg_neg_le0[of a b]apply arith
39b494e8c055 added lemma
nipkow
parents: 33730
diff changeset
  2238
using int_one_le_iff_zero_less[of "a div b"] zdiv_mono1[of b a b]apply simp
39b494e8c055 added lemma
nipkow
parents: 33730
diff changeset
  2239
done
39b494e8c055 added lemma
nipkow
parents: 33730
diff changeset
  2240
39489
8bb7f32a3a08 added lemmas
nipkow
parents: 38715
diff changeset
  2241
lemma zmod_le_nonneg_dividend: "(m::int) \<ge> 0 ==> m mod k \<le> m"
8bb7f32a3a08 added lemmas
nipkow
parents: 38715
diff changeset
  2242
apply (rule split_zmod[THEN iffD2])
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 44766
diff changeset
  2243
apply(fastforce dest: q_pos_lemma intro: split_mult_pos_le)
39489
8bb7f32a3a08 added lemmas
nipkow
parents: 38715
diff changeset
  2244
done
8bb7f32a3a08 added lemmas
nipkow
parents: 38715
diff changeset
  2245
8bb7f32a3a08 added lemmas
nipkow
parents: 38715
diff changeset
  2246
33361
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2247
subsubsection {* The Divides Relation *}
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2248
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2249
lemmas zdvd_iff_zmod_eq_0_number_of [simp] =
45607
16b4f5774621 eliminated obsolete "standard";
wenzelm
parents: 45530
diff changeset
  2250
  dvd_eq_mod_eq_0 [of "number_of x" "number_of y"] for x y :: int
33361
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2251
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2252
lemma zdvd_zmod: "f dvd m ==> f dvd (n::int) ==> f dvd m mod n"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2253
  by (rule dvd_mod) (* TODO: remove *)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2254
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2255
lemma zdvd_zmod_imp_zdvd: "k dvd m mod n ==> k dvd n ==> k dvd (m::int)"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2256
  by (rule dvd_mod_imp_dvd) (* TODO: remove *)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2257
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2258
lemma zmult_div_cancel: "(n::int) * (m div n) = m - (m mod n)"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2259
  using zmod_zdiv_equality[where a="m" and b="n"]
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2260
  by (simp add: algebra_simps)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2261
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2262
lemma zpower_zmod: "((x::int) mod m)^y mod m = x^y mod m"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2263
apply (induct "y", auto)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2264
apply (rule zmod_zmult1_eq [THEN trans])
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2265
apply (simp (no_asm_simp))
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2266
apply (rule mod_mult_eq [symmetric])
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2267
done
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2268
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2269
lemma zdiv_int: "int (a div b) = (int a) div (int b)"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2270
apply (subst split_div, auto)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2271
apply (subst split_zdiv, auto)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2272
apply (rule_tac a="int (b * i) + int j" and b="int b" and r="int j" and r'=ja in unique_quotient)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2273
apply (auto simp add: divmod_int_rel_def of_nat_mult)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2274
done
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2275
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2276
lemma zmod_int: "int (a mod b) = (int a) mod (int b)"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2277
apply (subst split_mod, auto)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2278
apply (subst split_zmod, auto)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2279
apply (rule_tac a="int (b * i) + int j" and b="int b" and q="int i" and q'=ia 
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2280
       in unique_remainder)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2281
apply (auto simp add: divmod_int_rel_def of_nat_mult)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2282
done
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2283
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2284
lemma abs_div: "(y::int) dvd x \<Longrightarrow> abs (x div y) = abs x div abs y"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2285
by (unfold dvd_def, cases "y=0", auto simp add: abs_mult)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2286
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2287
lemma zdvd_mult_div_cancel:"(n::int) dvd m \<Longrightarrow> n * (m div n) = m"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2288
apply (subgoal_tac "m mod n = 0")
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2289
 apply (simp add: zmult_div_cancel)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2290
apply (simp only: dvd_eq_mod_eq_0)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2291
done
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2292
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2293
text{*Suggested by Matthias Daum*}
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2294
lemma int_power_div_base:
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2295
     "\<lbrakk>0 < m; 0 < k\<rbrakk> \<Longrightarrow> k ^ m div k = (k::int) ^ (m - Suc 0)"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2296
apply (subgoal_tac "k ^ m = k ^ ((m - Suc 0) + Suc 0)")
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2297
 apply (erule ssubst)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2298
 apply (simp only: power_add)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2299
 apply simp_all
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2300
done
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2301
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2302
text {* by Brian Huffman *}
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2303
lemma zminus_zmod: "- ((x::int) mod m) mod m = - x mod m"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2304
by (rule mod_minus_eq [symmetric])
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2305
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2306
lemma zdiff_zmod_left: "(x mod m - y) mod m = (x - y) mod (m::int)"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2307
by (rule mod_diff_left_eq [symmetric])
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2308
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2309
lemma zdiff_zmod_right: "(x - y mod m) mod m = (x - y) mod (m::int)"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2310
by (rule mod_diff_right_eq [symmetric])
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2311
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2312
lemmas zmod_simps =
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2313
  mod_add_left_eq  [symmetric]
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2314
  mod_add_right_eq [symmetric]
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2315
  zmod_zmult1_eq   [symmetric]
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2316
  mod_mult_left_eq [symmetric]
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2317
  zpower_zmod
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2318
  zminus_zmod zdiff_zmod_left zdiff_zmod_right
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2319
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2320
text {* Distributive laws for function @{text nat}. *}
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2321
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2322
lemma nat_div_distrib: "0 \<le> x \<Longrightarrow> nat (x div y) = nat x div nat y"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2323
apply (rule linorder_cases [of y 0])
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2324
apply (simp add: div_nonneg_neg_le0)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2325
apply simp
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2326
apply (simp add: nat_eq_iff pos_imp_zdiv_nonneg_iff zdiv_int)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2327
done
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2328
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2329
(*Fails if y<0: the LHS collapses to (nat z) but the RHS doesn't*)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2330
lemma nat_mod_distrib:
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2331
  "\<lbrakk>0 \<le> x; 0 \<le> y\<rbrakk> \<Longrightarrow> nat (x mod y) = nat x mod nat y"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2332
apply (case_tac "y = 0", simp)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2333
apply (simp add: nat_eq_iff zmod_int)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2334
done
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2335
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2336
text  {* transfer setup *}
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2337
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2338
lemma transfer_nat_int_functions:
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2339
    "(x::int) >= 0 \<Longrightarrow> y >= 0 \<Longrightarrow> (nat x) div (nat y) = nat (x div y)"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2340
    "(x::int) >= 0 \<Longrightarrow> y >= 0 \<Longrightarrow> (nat x) mod (nat y) = nat (x mod y)"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2341
  by (auto simp add: nat_div_distrib nat_mod_distrib)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2342
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2343
lemma transfer_nat_int_function_closures:
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2344
    "(x::int) >= 0 \<Longrightarrow> y >= 0 \<Longrightarrow> x div y >= 0"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2345
    "(x::int) >= 0 \<Longrightarrow> y >= 0 \<Longrightarrow> x mod y >= 0"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2346
  apply (cases "y = 0")
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2347
  apply (auto simp add: pos_imp_zdiv_nonneg_iff)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2348
  apply (cases "y = 0")
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2349
  apply auto
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2350
done
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2351
35644
d20cf282342e transfer: avoid camel case
haftmann
parents: 35367
diff changeset
  2352
declare transfer_morphism_nat_int [transfer add return:
33361
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2353
  transfer_nat_int_functions
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2354
  transfer_nat_int_function_closures
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2355
]
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2356
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2357
lemma transfer_int_nat_functions:
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2358
    "(int x) div (int y) = int (x div y)"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2359
    "(int x) mod (int y) = int (x mod y)"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2360
  by (auto simp add: zdiv_int zmod_int)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2361
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2362
lemma transfer_int_nat_function_closures:
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2363
    "is_nat x \<Longrightarrow> is_nat y \<Longrightarrow> is_nat (x div y)"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2364
    "is_nat x \<Longrightarrow> is_nat y \<Longrightarrow> is_nat (x mod y)"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2365
  by (simp_all only: is_nat_def transfer_nat_int_function_closures)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2366
35644
d20cf282342e transfer: avoid camel case
haftmann
parents: 35367
diff changeset
  2367
declare transfer_morphism_int_nat [transfer add return:
33361
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2368
  transfer_int_nat_functions
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2369
  transfer_int_nat_function_closures
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2370
]
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2371
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2372
text{*Suggested by Matthias Daum*}
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2373
lemma int_div_less_self: "\<lbrakk>0 < x; 1 < k\<rbrakk> \<Longrightarrow> x div k < (x::int)"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2374
apply (subgoal_tac "nat x div nat k < nat x")
34225
21c5405deb6b removed legacy asm_lr
nipkow
parents: 34126
diff changeset
  2375
 apply (simp add: nat_div_distrib [symmetric])
33361
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2376
apply (rule Divides.div_less_dividend, simp_all)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2377
done
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2378
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2379
lemma zmod_eq_dvd_iff: "(x::int) mod n = y mod n \<longleftrightarrow> n dvd x - y"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2380
proof
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2381
  assume H: "x mod n = y mod n"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2382
  hence "x mod n - y mod n = 0" by simp
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2383
  hence "(x mod n - y mod n) mod n = 0" by simp 
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2384
  hence "(x - y) mod n = 0" by (simp add: mod_diff_eq[symmetric])
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2385
  thus "n dvd x - y" by (simp add: dvd_eq_mod_eq_0)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2386
next
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2387
  assume H: "n dvd x - y"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2388
  then obtain k where k: "x-y = n*k" unfolding dvd_def by blast
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2389
  hence "x = n*k + y" by simp
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2390
  hence "x mod n = (n*k + y) mod n" by simp
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2391
  thus "x mod n = y mod n" by (simp add: mod_add_left_eq)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2392
qed
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2393
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2394
lemma nat_mod_eq_lemma: assumes xyn: "(x::nat) mod n = y  mod n" and xy:"y \<le> x"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2395
  shows "\<exists>q. x = y + n * q"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2396
proof-
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2397
  from xy have th: "int x - int y = int (x - y)" by simp 
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2398
  from xyn have "int x mod int n = int y mod int n" 
46551
866bce5442a3 simplify projections on simultaneous computations of div and mod; tuned structure (from Florian Haftmann)
huffman
parents: 46026
diff changeset
  2399
    by (simp add: zmod_int [symmetric])
33361
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2400
  hence "int n dvd int x - int y" by (simp only: zmod_eq_dvd_iff[symmetric]) 
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2401
  hence "n dvd x - y" by (simp add: th zdvd_int)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2402
  then show ?thesis using xy unfolding dvd_def apply clarsimp apply (rule_tac x="k" in exI) by arith
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2403
qed
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2404
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2405
lemma nat_mod_eq_iff: "(x::nat) mod n = y mod n \<longleftrightarrow> (\<exists>q1 q2. x + n * q1 = y + n * q2)" 
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2406
  (is "?lhs = ?rhs")
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2407
proof
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2408
  assume H: "x mod n = y mod n"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2409
  {assume xy: "x \<le> y"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2410
    from H have th: "y mod n = x mod n" by simp
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2411
    from nat_mod_eq_lemma[OF th xy] have ?rhs 
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2412
      apply clarify  apply (rule_tac x="q" in exI) by (rule exI[where x="0"], simp)}
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2413
  moreover
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2414
  {assume xy: "y \<le> x"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2415
    from nat_mod_eq_lemma[OF H xy] have ?rhs 
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2416
      apply clarify  apply (rule_tac x="0" in exI) by (rule_tac x="q" in exI, simp)}
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2417
  ultimately  show ?rhs using linear[of x y] by blast  
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2418
next
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2419
  assume ?rhs then obtain q1 q2 where q12: "x + n * q1 = y + n * q2" by blast
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2420
  hence "(x + n * q1) mod n = (y + n * q2) mod n" by simp
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2421
  thus  ?lhs by simp
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2422
qed
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2423
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2424
lemma div_nat_number_of [simp]:
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2425
     "(number_of v :: nat)  div  number_of v' =  
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2426
          (if neg (number_of v :: int) then 0  
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2427
           else nat (number_of v div number_of v'))"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2428
  unfolding nat_number_of_def number_of_is_id neg_def
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2429
  by (simp add: nat_div_distrib)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2430
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2431
lemma one_div_nat_number_of [simp]:
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2432
     "Suc 0 div number_of v' = nat (1 div number_of v')" 
46026
83caa4f4bd56 semiring_numeral_0_eq_0, semiring_numeral_1_eq_1 now [simp], superseeding corresponding simp rules on type nat
haftmann
parents: 45607
diff changeset
  2433
  by (simp del: semiring_numeral_1_eq_1 add: numeral_1_eq_Suc_0 [symmetric] semiring_numeral_1_eq_1 [symmetric]) 
33361
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2434
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2435
lemma mod_nat_number_of [simp]:
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2436
     "(number_of v :: nat)  mod  number_of v' =  
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2437
        (if neg (number_of v :: int) then 0  
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2438
         else if neg (number_of v' :: int) then number_of v  
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2439
         else nat (number_of v mod number_of v'))"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2440
  unfolding nat_number_of_def number_of_is_id neg_def
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2441
  by (simp add: nat_mod_distrib)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2442
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2443
lemma one_mod_nat_number_of [simp]:
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2444
     "Suc 0 mod number_of v' =  
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2445
        (if neg (number_of v' :: int) then Suc 0
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2446
         else nat (1 mod number_of v'))"
46026
83caa4f4bd56 semiring_numeral_0_eq_0, semiring_numeral_1_eq_1 now [simp], superseeding corresponding simp rules on type nat
haftmann
parents: 45607
diff changeset
  2447
by (simp del: semiring_numeral_1_eq_1 add: numeral_1_eq_Suc_0 [symmetric] semiring_numeral_1_eq_1 [symmetric]) 
33361
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2448
45607
16b4f5774621 eliminated obsolete "standard";
wenzelm
parents: 45530
diff changeset
  2449
lemmas dvd_eq_mod_eq_0_number_of [simp] =
16b4f5774621 eliminated obsolete "standard";
wenzelm
parents: 45530
diff changeset
  2450
  dvd_eq_mod_eq_0 [of "number_of x" "number_of y"] for x y
33361
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2451
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2452
34126
8a2c5d7aff51 polished Nitpick's binary integer support etc.;
blanchet
parents: 33804
diff changeset
  2453
subsubsection {* Nitpick *}
8a2c5d7aff51 polished Nitpick's binary integer support etc.;
blanchet
parents: 33804
diff changeset
  2454
8a2c5d7aff51 polished Nitpick's binary integer support etc.;
blanchet
parents: 33804
diff changeset
  2455
lemma zmod_zdiv_equality':
8a2c5d7aff51 polished Nitpick's binary integer support etc.;
blanchet
parents: 33804
diff changeset
  2456
"(m\<Colon>int) mod n = m - (m div n) * n"
8a2c5d7aff51 polished Nitpick's binary integer support etc.;
blanchet
parents: 33804
diff changeset
  2457
by (rule_tac P="%x. m mod n = x - (m div n) * n"
8a2c5d7aff51 polished Nitpick's binary integer support etc.;
blanchet
parents: 33804
diff changeset
  2458
    in subst [OF mod_div_equality [of _ n]])
8a2c5d7aff51 polished Nitpick's binary integer support etc.;
blanchet
parents: 33804
diff changeset
  2459
   arith
8a2c5d7aff51 polished Nitpick's binary integer support etc.;
blanchet
parents: 33804
diff changeset
  2460
41792
ff3cb0c418b7 renamed "nitpick\_def" to "nitpick_unfold" to reflect its new semantics
blanchet
parents: 41550
diff changeset
  2461
lemmas [nitpick_unfold] = dvd_eq_mod_eq_0 mod_div_equality' zmod_zdiv_equality'
34126
8a2c5d7aff51 polished Nitpick's binary integer support etc.;
blanchet
parents: 33804
diff changeset
  2462
35673
178caf872f95 weakend class ring_div; tuned
haftmann
parents: 35644
diff changeset
  2463
33361
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2464
subsubsection {* Code generation *}
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2465
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2466
definition pdivmod :: "int \<Rightarrow> int \<Rightarrow> int \<times> int" where
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2467
  "pdivmod k l = (\<bar>k\<bar> div \<bar>l\<bar>, \<bar>k\<bar> mod \<bar>l\<bar>)"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2468
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2469
lemma pdivmod_posDivAlg [code]:
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2470
  "pdivmod k l = (if l = 0 then (0, \<bar>k\<bar>) else posDivAlg \<bar>k\<bar> \<bar>l\<bar>)"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2471
by (subst posDivAlg_div_mod) (simp_all add: pdivmod_def)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2472
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2473
lemma divmod_int_pdivmod: "divmod_int k l = (if k = 0 then (0, 0) else if l = 0 then (0, k) else
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2474
  apsnd ((op *) (sgn l)) (if 0 < l \<and> 0 \<le> k \<or> l < 0 \<and> k < 0
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2475
    then pdivmod k l
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2476
    else (let (r, s) = pdivmod k l in
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2477
      if s = 0 then (- r, 0) else (- r - 1, \<bar>l\<bar> - s))))"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2478
proof -
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2479
  have aux: "\<And>q::int. - k = l * q \<longleftrightarrow> k = l * - q" by auto
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2480
  show ?thesis
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2481
    by (simp add: divmod_int_mod_div pdivmod_def)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2482
      (auto simp add: aux not_less not_le zdiv_zminus1_eq_if
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2483
      zmod_zminus1_eq_if zdiv_zminus2_eq_if zmod_zminus2_eq_if)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2484
qed
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2485
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2486
lemma divmod_int_code [code]: "divmod_int k l = (if k = 0 then (0, 0) else if l = 0 then (0, k) else
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2487
  apsnd ((op *) (sgn l)) (if sgn k = sgn l
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2488
    then pdivmod k l
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2489
    else (let (r, s) = pdivmod k l in
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2490
      if s = 0 then (- r, 0) else (- r - 1, \<bar>l\<bar> - s))))"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2491
proof -
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2492
  have "k \<noteq> 0 \<Longrightarrow> l \<noteq> 0 \<Longrightarrow> 0 < l \<and> 0 \<le> k \<or> l < 0 \<and> k < 0 \<longleftrightarrow> sgn k = sgn l"
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2493
    by (auto simp add: not_less sgn_if)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2494
  then show ?thesis by (simp add: divmod_int_pdivmod)
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2495
qed
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2496
35673
178caf872f95 weakend class ring_div; tuned
haftmann
parents: 35644
diff changeset
  2497
context ring_1
178caf872f95 weakend class ring_div; tuned
haftmann
parents: 35644
diff changeset
  2498
begin
178caf872f95 weakend class ring_div; tuned
haftmann
parents: 35644
diff changeset
  2499
178caf872f95 weakend class ring_div; tuned
haftmann
parents: 35644
diff changeset
  2500
lemma of_int_num [code]:
178caf872f95 weakend class ring_div; tuned
haftmann
parents: 35644
diff changeset
  2501
  "of_int k = (if k = 0 then 0 else if k < 0 then
178caf872f95 weakend class ring_div; tuned
haftmann
parents: 35644
diff changeset
  2502
     - of_int (- k) else let
178caf872f95 weakend class ring_div; tuned
haftmann
parents: 35644
diff changeset
  2503
       (l, m) = divmod_int k 2;
178caf872f95 weakend class ring_div; tuned
haftmann
parents: 35644
diff changeset
  2504
       l' = of_int l
178caf872f95 weakend class ring_div; tuned
haftmann
parents: 35644
diff changeset
  2505
     in if m = 0 then l' + l' else l' + l' + 1)"
178caf872f95 weakend class ring_div; tuned
haftmann
parents: 35644
diff changeset
  2506
proof -
178caf872f95 weakend class ring_div; tuned
haftmann
parents: 35644
diff changeset
  2507
  have aux1: "k mod (2\<Colon>int) \<noteq> (0\<Colon>int) \<Longrightarrow> 
178caf872f95 weakend class ring_div; tuned
haftmann
parents: 35644
diff changeset
  2508
    of_int k = of_int (k div 2 * 2 + 1)"
178caf872f95 weakend class ring_div; tuned
haftmann
parents: 35644
diff changeset
  2509
  proof -
178caf872f95 weakend class ring_div; tuned
haftmann
parents: 35644
diff changeset
  2510
    have "k mod 2 < 2" by (auto intro: pos_mod_bound)
178caf872f95 weakend class ring_div; tuned
haftmann
parents: 35644
diff changeset
  2511
    moreover have "0 \<le> k mod 2" by (auto intro: pos_mod_sign)
178caf872f95 weakend class ring_div; tuned
haftmann
parents: 35644
diff changeset
  2512
    moreover assume "k mod 2 \<noteq> 0"
178caf872f95 weakend class ring_div; tuned
haftmann
parents: 35644
diff changeset
  2513
    ultimately have "k mod 2 = 1" by arith
178caf872f95 weakend class ring_div; tuned
haftmann
parents: 35644
diff changeset
  2514
    moreover have "of_int k = of_int (k div 2 * 2 + k mod 2)" by simp
178caf872f95 weakend class ring_div; tuned
haftmann
parents: 35644
diff changeset
  2515
    ultimately show ?thesis by auto
178caf872f95 weakend class ring_div; tuned
haftmann
parents: 35644
diff changeset
  2516
  qed
178caf872f95 weakend class ring_div; tuned
haftmann
parents: 35644
diff changeset
  2517
  have aux2: "\<And>x. of_int 2 * x = x + x"
178caf872f95 weakend class ring_div; tuned
haftmann
parents: 35644
diff changeset
  2518
  proof -
178caf872f95 weakend class ring_div; tuned
haftmann
parents: 35644
diff changeset
  2519
    fix x
178caf872f95 weakend class ring_div; tuned
haftmann
parents: 35644
diff changeset
  2520
    have int2: "(2::int) = 1 + 1" by arith
178caf872f95 weakend class ring_div; tuned
haftmann
parents: 35644
diff changeset
  2521
    show "of_int 2 * x = x + x"
178caf872f95 weakend class ring_div; tuned
haftmann
parents: 35644
diff changeset
  2522
    unfolding int2 of_int_add left_distrib by simp
178caf872f95 weakend class ring_div; tuned
haftmann
parents: 35644
diff changeset
  2523
  qed
178caf872f95 weakend class ring_div; tuned
haftmann
parents: 35644
diff changeset
  2524
  have aux3: "\<And>x. x * of_int 2 = x + x"
178caf872f95 weakend class ring_div; tuned
haftmann
parents: 35644
diff changeset
  2525
  proof -
178caf872f95 weakend class ring_div; tuned
haftmann
parents: 35644
diff changeset
  2526
    fix x
178caf872f95 weakend class ring_div; tuned
haftmann
parents: 35644
diff changeset
  2527
    have int2: "(2::int) = 1 + 1" by arith
178caf872f95 weakend class ring_div; tuned
haftmann
parents: 35644
diff changeset
  2528
    show "x * of_int 2 = x + x" 
178caf872f95 weakend class ring_div; tuned
haftmann
parents: 35644
diff changeset
  2529
    unfolding int2 of_int_add right_distrib by simp
178caf872f95 weakend class ring_div; tuned
haftmann
parents: 35644
diff changeset
  2530
  qed
178caf872f95 weakend class ring_div; tuned
haftmann
parents: 35644
diff changeset
  2531
  from aux1 show ?thesis by (auto simp add: divmod_int_mod_div Let_def aux2 aux3)
178caf872f95 weakend class ring_div; tuned
haftmann
parents: 35644
diff changeset
  2532
qed
178caf872f95 weakend class ring_div; tuned
haftmann
parents: 35644
diff changeset
  2533
178caf872f95 weakend class ring_div; tuned
haftmann
parents: 35644
diff changeset
  2534
end
178caf872f95 weakend class ring_div; tuned
haftmann
parents: 35644
diff changeset
  2535
33364
2bd12592c5e8 tuned code setup
haftmann
parents: 33361
diff changeset
  2536
code_modulename SML
2bd12592c5e8 tuned code setup
haftmann
parents: 33361
diff changeset
  2537
  Divides Arith
2bd12592c5e8 tuned code setup
haftmann
parents: 33361
diff changeset
  2538
2bd12592c5e8 tuned code setup
haftmann
parents: 33361
diff changeset
  2539
code_modulename OCaml
2bd12592c5e8 tuned code setup
haftmann
parents: 33361
diff changeset
  2540
  Divides Arith
2bd12592c5e8 tuned code setup
haftmann
parents: 33361
diff changeset
  2541
2bd12592c5e8 tuned code setup
haftmann
parents: 33361
diff changeset
  2542
code_modulename Haskell
2bd12592c5e8 tuned code setup
haftmann
parents: 33361
diff changeset
  2543
  Divides Arith
2bd12592c5e8 tuned code setup
haftmann
parents: 33361
diff changeset
  2544
33361
1f18de40b43f combined former theories Divides and IntDiv to one theory Divides
haftmann
parents: 33340
diff changeset
  2545
end