src/HOL/Euclidean_Division.thy
author nipkow
Tue, 11 Oct 2022 12:13:47 +0200
changeset 76260 5fd8ba24ca48
parent 76246 c9ea813f92f2
child 76387 8cb141384682
permissions -rw-r--r--
removed redundant lemma
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
64785
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
     1
(*  Title:      HOL/Euclidean_Division.thy
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
     2
    Author:     Manuel Eberl, TU Muenchen
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
     3
    Author:     Florian Haftmann, TU Muenchen
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
     4
*)
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
     5
66817
0b12755ccbb2 euclidean rings need no normalization
haftmann
parents: 66816
diff changeset
     6
section \<open>Division in euclidean (semi)rings\<close>
64785
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
     7
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
     8
theory Euclidean_Division
66817
0b12755ccbb2 euclidean rings need no normalization
haftmann
parents: 66816
diff changeset
     9
  imports Int Lattices_Big
64785
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
    10
begin
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
    11
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
    12
subsection \<open>Euclidean (semi)rings with explicit division and remainder\<close>
76224
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
    13
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
    14
class euclidean_semiring = semidom_modulo +
64785
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
    15
  fixes euclidean_size :: "'a \<Rightarrow> nat"
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
    16
  assumes size_0 [simp]: "euclidean_size 0 = 0"
76224
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
    17
  assumes mod_size_less:
64785
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
    18
    "b \<noteq> 0 \<Longrightarrow> euclidean_size (a mod b) < euclidean_size b"
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
    19
  assumes size_mult_mono:
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
    20
    "b \<noteq> 0 \<Longrightarrow> euclidean_size a \<le> euclidean_size (a * b)"
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
    21
begin
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
    22
66840
0d689d71dbdc canonical multiplicative euclidean size
haftmann
parents: 66839
diff changeset
    23
lemma euclidean_size_eq_0_iff [simp]:
0d689d71dbdc canonical multiplicative euclidean size
haftmann
parents: 66839
diff changeset
    24
  "euclidean_size b = 0 \<longleftrightarrow> b = 0"
0d689d71dbdc canonical multiplicative euclidean size
haftmann
parents: 66839
diff changeset
    25
proof
0d689d71dbdc canonical multiplicative euclidean size
haftmann
parents: 66839
diff changeset
    26
  assume "b = 0"
0d689d71dbdc canonical multiplicative euclidean size
haftmann
parents: 66839
diff changeset
    27
  then show "euclidean_size b = 0"
0d689d71dbdc canonical multiplicative euclidean size
haftmann
parents: 66839
diff changeset
    28
    by simp
0d689d71dbdc canonical multiplicative euclidean size
haftmann
parents: 66839
diff changeset
    29
next
0d689d71dbdc canonical multiplicative euclidean size
haftmann
parents: 66839
diff changeset
    30
  assume "euclidean_size b = 0"
0d689d71dbdc canonical multiplicative euclidean size
haftmann
parents: 66839
diff changeset
    31
  show "b = 0"
0d689d71dbdc canonical multiplicative euclidean size
haftmann
parents: 66839
diff changeset
    32
  proof (rule ccontr)
0d689d71dbdc canonical multiplicative euclidean size
haftmann
parents: 66839
diff changeset
    33
    assume "b \<noteq> 0"
0d689d71dbdc canonical multiplicative euclidean size
haftmann
parents: 66839
diff changeset
    34
    with mod_size_less have "euclidean_size (b mod b) < euclidean_size b" .
0d689d71dbdc canonical multiplicative euclidean size
haftmann
parents: 66839
diff changeset
    35
    with \<open>euclidean_size b = 0\<close> show False
0d689d71dbdc canonical multiplicative euclidean size
haftmann
parents: 66839
diff changeset
    36
      by simp
0d689d71dbdc canonical multiplicative euclidean size
haftmann
parents: 66839
diff changeset
    37
  qed
0d689d71dbdc canonical multiplicative euclidean size
haftmann
parents: 66839
diff changeset
    38
qed
0d689d71dbdc canonical multiplicative euclidean size
haftmann
parents: 66839
diff changeset
    39
0d689d71dbdc canonical multiplicative euclidean size
haftmann
parents: 66839
diff changeset
    40
lemma euclidean_size_greater_0_iff [simp]:
0d689d71dbdc canonical multiplicative euclidean size
haftmann
parents: 66839
diff changeset
    41
  "euclidean_size b > 0 \<longleftrightarrow> b \<noteq> 0"
0d689d71dbdc canonical multiplicative euclidean size
haftmann
parents: 66839
diff changeset
    42
  using euclidean_size_eq_0_iff [symmetric, of b] by safe simp
0d689d71dbdc canonical multiplicative euclidean size
haftmann
parents: 66839
diff changeset
    43
64785
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
    44
lemma size_mult_mono': "b \<noteq> 0 \<Longrightarrow> euclidean_size a \<le> euclidean_size (b * a)"
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
    45
  by (subst mult.commute) (rule size_mult_mono)
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
    46
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
    47
lemma dvd_euclidean_size_eq_imp_dvd:
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
    48
  assumes "a \<noteq> 0" and "euclidean_size a = euclidean_size b"
76224
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
    49
    and "b dvd a"
64785
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
    50
  shows "a dvd b"
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
    51
proof (rule ccontr)
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
    52
  assume "\<not> a dvd b"
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
    53
  hence "b mod a \<noteq> 0" using mod_0_imp_dvd [of b a] by blast
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
    54
  then have "b mod a \<noteq> 0" by (simp add: mod_eq_0_iff_dvd)
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
    55
  from \<open>b dvd a\<close> have "b dvd b mod a" by (simp add: dvd_mod_iff)
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
    56
  then obtain c where "b mod a = b * c" unfolding dvd_def by blast
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
    57
    with \<open>b mod a \<noteq> 0\<close> have "c \<noteq> 0" by auto
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
    58
  with \<open>b mod a = b * c\<close> have "euclidean_size (b mod a) \<ge> euclidean_size b"
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
    59
    using size_mult_mono by force
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
    60
  moreover from \<open>\<not> a dvd b\<close> and \<open>a \<noteq> 0\<close>
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
    61
  have "euclidean_size (b mod a) < euclidean_size a"
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
    62
    using mod_size_less by blast
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
    63
  ultimately show False using \<open>euclidean_size a = euclidean_size b\<close>
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
    64
    by simp
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
    65
qed
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
    66
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
    67
lemma euclidean_size_times_unit:
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
    68
  assumes "is_unit a"
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
    69
  shows   "euclidean_size (a * b) = euclidean_size b"
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
    70
proof (rule antisym)
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
    71
  from assms have [simp]: "a \<noteq> 0" by auto
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
    72
  thus "euclidean_size (a * b) \<ge> euclidean_size b" by (rule size_mult_mono')
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
    73
  from assms have "is_unit (1 div a)" by simp
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
    74
  hence "1 div a \<noteq> 0" by (intro notI) simp_all
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
    75
  hence "euclidean_size (a * b) \<le> euclidean_size ((1 div a) * (a * b))"
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
    76
    by (rule size_mult_mono')
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
    77
  also from assms have "(1 div a) * (a * b) = b"
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
    78
    by (simp add: algebra_simps unit_div_mult_swap)
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
    79
  finally show "euclidean_size (a * b) \<le> euclidean_size b" .
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
    80
qed
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
    81
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
    82
lemma euclidean_size_unit:
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
    83
  "is_unit a \<Longrightarrow> euclidean_size a = euclidean_size 1"
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
    84
  using euclidean_size_times_unit [of a 1] by simp
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
    85
76224
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
    86
lemma unit_iff_euclidean_size:
64785
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
    87
  "is_unit a \<longleftrightarrow> euclidean_size a = euclidean_size 1 \<and> a \<noteq> 0"
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
    88
proof safe
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
    89
  assume A: "a \<noteq> 0" and B: "euclidean_size a = euclidean_size 1"
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
    90
  show "is_unit a"
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
    91
    by (rule dvd_euclidean_size_eq_imp_dvd [OF A B]) simp_all
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
    92
qed (auto intro: euclidean_size_unit)
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
    93
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
    94
lemma euclidean_size_times_nonunit:
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
    95
  assumes "a \<noteq> 0" "b \<noteq> 0" "\<not> is_unit a"
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
    96
  shows   "euclidean_size b < euclidean_size (a * b)"
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
    97
proof (rule ccontr)
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
    98
  assume "\<not>euclidean_size b < euclidean_size (a * b)"
76224
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
    99
  with size_mult_mono'[OF assms(1), of b]
64785
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   100
    have eq: "euclidean_size (a * b) = euclidean_size b" by simp
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   101
  have "a * b dvd b"
75669
43f5dfb7fa35 tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents: 74592
diff changeset
   102
    by (rule dvd_euclidean_size_eq_imp_dvd [OF _ eq])
43f5dfb7fa35 tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents: 74592
diff changeset
   103
       (use assms in simp_all)
64785
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   104
  hence "a * b dvd 1 * b" by simp
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   105
  with \<open>b \<noteq> 0\<close> have "is_unit a" by (subst (asm) dvd_times_right_cancel_iff)
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   106
  with assms(3) show False by contradiction
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   107
qed
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   108
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   109
lemma dvd_imp_size_le:
76224
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
   110
  assumes "a dvd b" "b \<noteq> 0"
64785
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   111
  shows   "euclidean_size a \<le> euclidean_size b"
75669
43f5dfb7fa35 tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents: 74592
diff changeset
   112
  using assms by (auto simp: size_mult_mono)
64785
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   113
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   114
lemma dvd_proper_imp_size_less:
76224
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
   115
  assumes "a dvd b" "\<not> b dvd a" "b \<noteq> 0"
64785
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   116
  shows   "euclidean_size a < euclidean_size b"
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   117
proof -
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   118
  from assms(1) obtain c where "b = a * c" by (erule dvdE)
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   119
  hence z: "b = c * a" by (simp add: mult.commute)
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   120
  from z assms have "\<not>is_unit c" by (auto simp: mult.commute mult_unit_dvd_iff)
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   121
  with z assms show ?thesis
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   122
    by (auto intro!: euclidean_size_times_nonunit)
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   123
qed
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   124
66798
39bb2462e681 fundamental property of division by units
haftmann
parents: 64785
diff changeset
   125
lemma unit_imp_mod_eq_0:
39bb2462e681 fundamental property of division by units
haftmann
parents: 64785
diff changeset
   126
  "a mod b = 0" if "is_unit b"
39bb2462e681 fundamental property of division by units
haftmann
parents: 64785
diff changeset
   127
  using that by (simp add: mod_eq_0_iff_dvd unit_imp_dvd)
39bb2462e681 fundamental property of division by units
haftmann
parents: 64785
diff changeset
   128
69695
753ae9e9773d algebraized more material from theory Divides
haftmann
parents: 69593
diff changeset
   129
lemma mod_eq_self_iff_div_eq_0:
753ae9e9773d algebraized more material from theory Divides
haftmann
parents: 69593
diff changeset
   130
  "a mod b = a \<longleftrightarrow> a div b = 0" (is "?P \<longleftrightarrow> ?Q")
753ae9e9773d algebraized more material from theory Divides
haftmann
parents: 69593
diff changeset
   131
proof
753ae9e9773d algebraized more material from theory Divides
haftmann
parents: 69593
diff changeset
   132
  assume ?P
753ae9e9773d algebraized more material from theory Divides
haftmann
parents: 69593
diff changeset
   133
  with div_mult_mod_eq [of a b] show ?Q
753ae9e9773d algebraized more material from theory Divides
haftmann
parents: 69593
diff changeset
   134
    by auto
753ae9e9773d algebraized more material from theory Divides
haftmann
parents: 69593
diff changeset
   135
next
753ae9e9773d algebraized more material from theory Divides
haftmann
parents: 69593
diff changeset
   136
  assume ?Q
753ae9e9773d algebraized more material from theory Divides
haftmann
parents: 69593
diff changeset
   137
  with div_mult_mod_eq [of a b] show ?P
753ae9e9773d algebraized more material from theory Divides
haftmann
parents: 69593
diff changeset
   138
    by simp
753ae9e9773d algebraized more material from theory Divides
haftmann
parents: 69593
diff changeset
   139
qed
753ae9e9773d algebraized more material from theory Divides
haftmann
parents: 69593
diff changeset
   140
67051
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
   141
lemma coprime_mod_left_iff [simp]:
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
   142
  "coprime (a mod b) b \<longleftrightarrow> coprime a b" if "b \<noteq> 0"
75669
43f5dfb7fa35 tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents: 74592
diff changeset
   143
  by (rule iffI; rule coprimeI)
67051
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
   144
    (use that in \<open>auto dest!: dvd_mod_imp_dvd coprime_common_divisor simp add: dvd_mod_iff\<close>)
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
   145
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
   146
lemma coprime_mod_right_iff [simp]:
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
   147
  "coprime a (b mod a) \<longleftrightarrow> coprime a b" if "a \<noteq> 0"
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
   148
  using that coprime_mod_left_iff [of a b] by (simp add: ac_simps)
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
   149
64785
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   150
end
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   151
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   152
class euclidean_ring = idom_modulo + euclidean_semiring
66886
960509bfd47e added lemmas and tuned proofs
haftmann
parents: 66840
diff changeset
   153
begin
960509bfd47e added lemmas and tuned proofs
haftmann
parents: 66840
diff changeset
   154
67087
733017b19de9 generalized more lemmas
haftmann
parents: 67083
diff changeset
   155
lemma dvd_diff_commute [ac_simps]:
66886
960509bfd47e added lemmas and tuned proofs
haftmann
parents: 66840
diff changeset
   156
  "a dvd c - b \<longleftrightarrow> a dvd b - c"
960509bfd47e added lemmas and tuned proofs
haftmann
parents: 66840
diff changeset
   157
proof -
960509bfd47e added lemmas and tuned proofs
haftmann
parents: 66840
diff changeset
   158
  have "a dvd c - b \<longleftrightarrow> a dvd (c - b) * - 1"
960509bfd47e added lemmas and tuned proofs
haftmann
parents: 66840
diff changeset
   159
    by (subst dvd_mult_unit_iff) simp_all
960509bfd47e added lemmas and tuned proofs
haftmann
parents: 66840
diff changeset
   160
  then show ?thesis
960509bfd47e added lemmas and tuned proofs
haftmann
parents: 66840
diff changeset
   161
    by simp
960509bfd47e added lemmas and tuned proofs
haftmann
parents: 66840
diff changeset
   162
qed
76224
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
   163
66886
960509bfd47e added lemmas and tuned proofs
haftmann
parents: 66840
diff changeset
   164
end
64785
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   165
66840
0d689d71dbdc canonical multiplicative euclidean size
haftmann
parents: 66839
diff changeset
   166
66806
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   167
subsection \<open>Euclidean (semi)rings with cancel rules\<close>
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   168
70147
1657688a6406 backed out a93e6472ac9c, which does not bring anything substantial: division_ring is not commutative in multiplication but semidom_divide is
haftmann
parents: 70094
diff changeset
   169
class euclidean_semiring_cancel = euclidean_semiring +
1657688a6406 backed out a93e6472ac9c, which does not bring anything substantial: division_ring is not commutative in multiplication but semidom_divide is
haftmann
parents: 70094
diff changeset
   170
  assumes div_mult_self1 [simp]: "b \<noteq> 0 \<Longrightarrow> (a + c * b) div b = c + a div b"
1657688a6406 backed out a93e6472ac9c, which does not bring anything substantial: division_ring is not commutative in multiplication but semidom_divide is
haftmann
parents: 70094
diff changeset
   171
  and div_mult_mult1 [simp]: "c \<noteq> 0 \<Longrightarrow> (c * a) div (c * b) = a div b"
66806
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   172
begin
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   173
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   174
lemma div_mult_self2 [simp]:
70147
1657688a6406 backed out a93e6472ac9c, which does not bring anything substantial: division_ring is not commutative in multiplication but semidom_divide is
haftmann
parents: 70094
diff changeset
   175
  assumes "b \<noteq> 0"
1657688a6406 backed out a93e6472ac9c, which does not bring anything substantial: division_ring is not commutative in multiplication but semidom_divide is
haftmann
parents: 70094
diff changeset
   176
  shows "(a + b * c) div b = c + a div b"
1657688a6406 backed out a93e6472ac9c, which does not bring anything substantial: division_ring is not commutative in multiplication but semidom_divide is
haftmann
parents: 70094
diff changeset
   177
  using assms div_mult_self1 [of b a c] by (simp add: mult.commute)
66806
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   178
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   179
lemma div_mult_self3 [simp]:
70147
1657688a6406 backed out a93e6472ac9c, which does not bring anything substantial: division_ring is not commutative in multiplication but semidom_divide is
haftmann
parents: 70094
diff changeset
   180
  assumes "b \<noteq> 0"
1657688a6406 backed out a93e6472ac9c, which does not bring anything substantial: division_ring is not commutative in multiplication but semidom_divide is
haftmann
parents: 70094
diff changeset
   181
  shows "(c * b + a) div b = c + a div b"
1657688a6406 backed out a93e6472ac9c, which does not bring anything substantial: division_ring is not commutative in multiplication but semidom_divide is
haftmann
parents: 70094
diff changeset
   182
  using assms by (simp add: add.commute)
66806
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   183
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   184
lemma div_mult_self4 [simp]:
70147
1657688a6406 backed out a93e6472ac9c, which does not bring anything substantial: division_ring is not commutative in multiplication but semidom_divide is
haftmann
parents: 70094
diff changeset
   185
  assumes "b \<noteq> 0"
1657688a6406 backed out a93e6472ac9c, which does not bring anything substantial: division_ring is not commutative in multiplication but semidom_divide is
haftmann
parents: 70094
diff changeset
   186
  shows "(b * c + a) div b = c + a div b"
1657688a6406 backed out a93e6472ac9c, which does not bring anything substantial: division_ring is not commutative in multiplication but semidom_divide is
haftmann
parents: 70094
diff changeset
   187
  using assms by (simp add: add.commute)
66806
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   188
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   189
lemma mod_mult_self1 [simp]: "(a + c * b) mod b = a mod b"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   190
proof (cases "b = 0")
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   191
  case True then show ?thesis by simp
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   192
next
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   193
  case False
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   194
  have "a + c * b = (a + c * b) div b * b + (a + c * b) mod b"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   195
    by (simp add: div_mult_mod_eq)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   196
  also from False div_mult_self1 [of b a c] have
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   197
    "\<dots> = (c + a div b) * b + (a + c * b) mod b"
70147
1657688a6406 backed out a93e6472ac9c, which does not bring anything substantial: division_ring is not commutative in multiplication but semidom_divide is
haftmann
parents: 70094
diff changeset
   198
      by (simp add: algebra_simps)
66806
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   199
  finally have "a = a div b * b + (a + c * b) mod b"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   200
    by (simp add: add.commute [of a] add.assoc distrib_right)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   201
  then have "a div b * b + (a + c * b) mod b = a div b * b + a mod b"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   202
    by (simp add: div_mult_mod_eq)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   203
  then show ?thesis by simp
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   204
qed
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   205
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   206
lemma mod_mult_self2 [simp]:
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   207
  "(a + b * c) mod b = a mod b"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   208
  by (simp add: mult.commute [of b])
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   209
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   210
lemma mod_mult_self3 [simp]:
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   211
  "(c * b + a) mod b = a mod b"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   212
  by (simp add: add.commute)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   213
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   214
lemma mod_mult_self4 [simp]:
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   215
  "(b * c + a) mod b = a mod b"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   216
  by (simp add: add.commute)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   217
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   218
lemma mod_mult_self1_is_0 [simp]:
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   219
  "b * a mod b = 0"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   220
  using mod_mult_self2 [of 0 b a] by simp
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   221
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   222
lemma mod_mult_self2_is_0 [simp]:
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   223
  "a * b mod b = 0"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   224
  using mod_mult_self1 [of 0 a b] by simp
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   225
70147
1657688a6406 backed out a93e6472ac9c, which does not bring anything substantial: division_ring is not commutative in multiplication but semidom_divide is
haftmann
parents: 70094
diff changeset
   226
lemma div_add_self1:
1657688a6406 backed out a93e6472ac9c, which does not bring anything substantial: division_ring is not commutative in multiplication but semidom_divide is
haftmann
parents: 70094
diff changeset
   227
  assumes "b \<noteq> 0"
1657688a6406 backed out a93e6472ac9c, which does not bring anything substantial: division_ring is not commutative in multiplication but semidom_divide is
haftmann
parents: 70094
diff changeset
   228
  shows "(b + a) div b = a div b + 1"
1657688a6406 backed out a93e6472ac9c, which does not bring anything substantial: division_ring is not commutative in multiplication but semidom_divide is
haftmann
parents: 70094
diff changeset
   229
  using assms div_mult_self1 [of b a 1] by (simp add: add.commute)
1657688a6406 backed out a93e6472ac9c, which does not bring anything substantial: division_ring is not commutative in multiplication but semidom_divide is
haftmann
parents: 70094
diff changeset
   230
1657688a6406 backed out a93e6472ac9c, which does not bring anything substantial: division_ring is not commutative in multiplication but semidom_divide is
haftmann
parents: 70094
diff changeset
   231
lemma div_add_self2:
1657688a6406 backed out a93e6472ac9c, which does not bring anything substantial: division_ring is not commutative in multiplication but semidom_divide is
haftmann
parents: 70094
diff changeset
   232
  assumes "b \<noteq> 0"
1657688a6406 backed out a93e6472ac9c, which does not bring anything substantial: division_ring is not commutative in multiplication but semidom_divide is
haftmann
parents: 70094
diff changeset
   233
  shows "(a + b) div b = a div b + 1"
1657688a6406 backed out a93e6472ac9c, which does not bring anything substantial: division_ring is not commutative in multiplication but semidom_divide is
haftmann
parents: 70094
diff changeset
   234
  using assms div_add_self1 [of b a] by (simp add: add.commute)
1657688a6406 backed out a93e6472ac9c, which does not bring anything substantial: division_ring is not commutative in multiplication but semidom_divide is
haftmann
parents: 70094
diff changeset
   235
66806
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   236
lemma mod_add_self1 [simp]:
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   237
  "(b + a) mod b = a mod b"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   238
  using mod_mult_self1 [of a 1 b] by (simp add: add.commute)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   239
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   240
lemma mod_add_self2 [simp]:
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   241
  "(a + b) mod b = a mod b"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   242
  using mod_mult_self1 [of a 1 b] by simp
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   243
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   244
lemma mod_div_trivial [simp]:
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   245
  "a mod b div b = 0"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   246
proof (cases "b = 0")
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   247
  assume "b = 0"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   248
  thus ?thesis by simp
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   249
next
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   250
  assume "b \<noteq> 0"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   251
  hence "a div b + a mod b div b = (a mod b + a div b * b) div b"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   252
    by (rule div_mult_self1 [symmetric])
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   253
  also have "\<dots> = a div b"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   254
    by (simp only: mod_div_mult_eq)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   255
  also have "\<dots> = a div b + 0"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   256
    by simp
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   257
  finally show ?thesis
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   258
    by (rule add_left_imp_eq)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   259
qed
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   260
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   261
lemma mod_mod_trivial [simp]:
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   262
  "a mod b mod b = a mod b"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   263
proof -
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   264
  have "a mod b mod b = (a mod b + a div b * b) mod b"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   265
    by (simp only: mod_mult_self1)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   266
  also have "\<dots> = a mod b"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   267
    by (simp only: mod_div_mult_eq)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   268
  finally show ?thesis .
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   269
qed
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   270
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   271
lemma mod_mod_cancel:
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   272
  assumes "c dvd b"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   273
  shows "a mod b mod c = a mod c"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   274
proof -
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   275
  from \<open>c dvd b\<close> obtain k where "b = c * k"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   276
    by (rule dvdE)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   277
  have "a mod b mod c = a mod (c * k) mod c"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   278
    by (simp only: \<open>b = c * k\<close>)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   279
  also have "\<dots> = (a mod (c * k) + a div (c * k) * k * c) mod c"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   280
    by (simp only: mod_mult_self1)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   281
  also have "\<dots> = (a div (c * k) * (c * k) + a mod (c * k)) mod c"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   282
    by (simp only: ac_simps)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   283
  also have "\<dots> = a mod c"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   284
    by (simp only: div_mult_mod_eq)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   285
  finally show ?thesis .
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   286
qed
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   287
70147
1657688a6406 backed out a93e6472ac9c, which does not bring anything substantial: division_ring is not commutative in multiplication but semidom_divide is
haftmann
parents: 70094
diff changeset
   288
lemma div_mult_mult2 [simp]:
1657688a6406 backed out a93e6472ac9c, which does not bring anything substantial: division_ring is not commutative in multiplication but semidom_divide is
haftmann
parents: 70094
diff changeset
   289
  "c \<noteq> 0 \<Longrightarrow> (a * c) div (b * c) = a div b"
1657688a6406 backed out a93e6472ac9c, which does not bring anything substantial: division_ring is not commutative in multiplication but semidom_divide is
haftmann
parents: 70094
diff changeset
   290
  by (drule div_mult_mult1) (simp add: mult.commute)
1657688a6406 backed out a93e6472ac9c, which does not bring anything substantial: division_ring is not commutative in multiplication but semidom_divide is
haftmann
parents: 70094
diff changeset
   291
1657688a6406 backed out a93e6472ac9c, which does not bring anything substantial: division_ring is not commutative in multiplication but semidom_divide is
haftmann
parents: 70094
diff changeset
   292
lemma div_mult_mult1_if [simp]:
1657688a6406 backed out a93e6472ac9c, which does not bring anything substantial: division_ring is not commutative in multiplication but semidom_divide is
haftmann
parents: 70094
diff changeset
   293
  "(c * a) div (c * b) = (if c = 0 then 0 else a div b)"
1657688a6406 backed out a93e6472ac9c, which does not bring anything substantial: division_ring is not commutative in multiplication but semidom_divide is
haftmann
parents: 70094
diff changeset
   294
  by simp_all
1657688a6406 backed out a93e6472ac9c, which does not bring anything substantial: division_ring is not commutative in multiplication but semidom_divide is
haftmann
parents: 70094
diff changeset
   295
66806
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   296
lemma mod_mult_mult1:
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   297
  "(c * a) mod (c * b) = c * (a mod b)"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   298
proof (cases "c = 0")
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   299
  case True then show ?thesis by simp
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   300
next
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   301
  case False
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   302
  from div_mult_mod_eq
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   303
  have "((c * a) div (c * b)) * (c * b) + (c * a) mod (c * b) = c * a" .
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   304
  with False have "c * ((a div b) * b + a mod b) + (c * a) mod (c * b)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   305
    = c * a + c * (a mod b)" by (simp add: algebra_simps)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   306
  with div_mult_mod_eq show ?thesis by simp
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   307
qed
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   308
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   309
lemma mod_mult_mult2:
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   310
  "(a * c) mod (b * c) = (a mod b) * c"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   311
  using mod_mult_mult1 [of c a b] by (simp add: mult.commute)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   312
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   313
lemma mult_mod_left: "(a mod b) * c = (a * c) mod (b * c)"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   314
  by (fact mod_mult_mult2 [symmetric])
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   315
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   316
lemma mult_mod_right: "c * (a mod b) = (c * a) mod (c * b)"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   317
  by (fact mod_mult_mult1 [symmetric])
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   318
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   319
lemma dvd_mod: "k dvd m \<Longrightarrow> k dvd n \<Longrightarrow> k dvd (m mod n)"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   320
  unfolding dvd_def by (auto simp add: mod_mult_mult1)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   321
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   322
lemma div_plus_div_distrib_dvd_left:
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   323
  "c dvd a \<Longrightarrow> (a + b) div c = a div c + b div c"
75875
48d032035744 streamlined primitive definitions for integer division
haftmann
parents: 75669
diff changeset
   324
  by (cases "c = 0") auto
66806
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   325
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   326
lemma div_plus_div_distrib_dvd_right:
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   327
  "c dvd b \<Longrightarrow> (a + b) div c = a div c + b div c"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   328
  using div_plus_div_distrib_dvd_left [of c b a]
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   329
  by (simp add: ac_simps)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   330
71413
65ffe9e910d4 more specific class assumptions
haftmann
parents: 71412
diff changeset
   331
lemma sum_div_partition:
65ffe9e910d4 more specific class assumptions
haftmann
parents: 71412
diff changeset
   332
  \<open>(\<Sum>a\<in>A. f a) div b = (\<Sum>a\<in>A \<inter> {a. b dvd f a}. f a div b) + (\<Sum>a\<in>A \<inter> {a. \<not> b dvd f a}. f a) div b\<close>
65ffe9e910d4 more specific class assumptions
haftmann
parents: 71412
diff changeset
   333
    if \<open>finite A\<close>
65ffe9e910d4 more specific class assumptions
haftmann
parents: 71412
diff changeset
   334
proof -
65ffe9e910d4 more specific class assumptions
haftmann
parents: 71412
diff changeset
   335
  have \<open>A = A \<inter> {a. b dvd f a} \<union> A \<inter> {a. \<not> b dvd f a}\<close>
65ffe9e910d4 more specific class assumptions
haftmann
parents: 71412
diff changeset
   336
    by auto
65ffe9e910d4 more specific class assumptions
haftmann
parents: 71412
diff changeset
   337
  then have \<open>(\<Sum>a\<in>A. f a) = (\<Sum>a\<in>A \<inter> {a. b dvd f a} \<union> A \<inter> {a. \<not> b dvd f a}. f a)\<close>
65ffe9e910d4 more specific class assumptions
haftmann
parents: 71412
diff changeset
   338
    by simp
65ffe9e910d4 more specific class assumptions
haftmann
parents: 71412
diff changeset
   339
  also have \<open>\<dots> = (\<Sum>a\<in>A \<inter> {a. b dvd f a}. f a) + (\<Sum>a\<in>A \<inter> {a. \<not> b dvd f a}. f a)\<close>
65ffe9e910d4 more specific class assumptions
haftmann
parents: 71412
diff changeset
   340
    using \<open>finite A\<close> by (auto intro: sum.union_inter_neutral)
65ffe9e910d4 more specific class assumptions
haftmann
parents: 71412
diff changeset
   341
  finally have *: \<open>sum f A = sum f (A \<inter> {a. b dvd f a}) + sum f (A \<inter> {a. \<not> b dvd f a})\<close> .
65ffe9e910d4 more specific class assumptions
haftmann
parents: 71412
diff changeset
   342
  define B where B: \<open>B = A \<inter> {a. b dvd f a}\<close>
65ffe9e910d4 more specific class assumptions
haftmann
parents: 71412
diff changeset
   343
  with \<open>finite A\<close> have \<open>finite B\<close> and \<open>a \<in> B \<Longrightarrow> b dvd f a\<close> for a
65ffe9e910d4 more specific class assumptions
haftmann
parents: 71412
diff changeset
   344
    by simp_all
65ffe9e910d4 more specific class assumptions
haftmann
parents: 71412
diff changeset
   345
  then have \<open>(\<Sum>a\<in>B. f a) div b = (\<Sum>a\<in>B. f a div b)\<close> and \<open>b dvd (\<Sum>a\<in>B. f a)\<close>
65ffe9e910d4 more specific class assumptions
haftmann
parents: 71412
diff changeset
   346
    by induction (simp_all add: div_plus_div_distrib_dvd_left)
65ffe9e910d4 more specific class assumptions
haftmann
parents: 71412
diff changeset
   347
  then show ?thesis using *
65ffe9e910d4 more specific class assumptions
haftmann
parents: 71412
diff changeset
   348
    by (simp add: B div_plus_div_distrib_dvd_left)
65ffe9e910d4 more specific class assumptions
haftmann
parents: 71412
diff changeset
   349
qed
65ffe9e910d4 more specific class assumptions
haftmann
parents: 71412
diff changeset
   350
66806
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   351
named_theorems mod_simps
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   352
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   353
text \<open>Addition respects modular equivalence.\<close>
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   354
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   355
lemma mod_add_left_eq [mod_simps]:
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   356
  "(a mod c + b) mod c = (a + b) mod c"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   357
proof -
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   358
  have "(a + b) mod c = (a div c * c + a mod c + b) mod c"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   359
    by (simp only: div_mult_mod_eq)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   360
  also have "\<dots> = (a mod c + b + a div c * c) mod c"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   361
    by (simp only: ac_simps)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   362
  also have "\<dots> = (a mod c + b) mod c"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   363
    by (rule mod_mult_self1)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   364
  finally show ?thesis
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   365
    by (rule sym)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   366
qed
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   367
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   368
lemma mod_add_right_eq [mod_simps]:
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   369
  "(a + b mod c) mod c = (a + b) mod c"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   370
  using mod_add_left_eq [of b c a] by (simp add: ac_simps)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   371
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   372
lemma mod_add_eq:
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   373
  "(a mod c + b mod c) mod c = (a + b) mod c"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   374
  by (simp add: mod_add_left_eq mod_add_right_eq)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   375
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   376
lemma mod_sum_eq [mod_simps]:
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   377
  "(\<Sum>i\<in>A. f i mod a) mod a = sum f A mod a"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   378
proof (induct A rule: infinite_finite_induct)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   379
  case (insert i A)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   380
  then have "(\<Sum>i\<in>insert i A. f i mod a) mod a
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   381
    = (f i mod a + (\<Sum>i\<in>A. f i mod a)) mod a"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   382
    by simp
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   383
  also have "\<dots> = (f i + (\<Sum>i\<in>A. f i mod a) mod a) mod a"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   384
    by (simp add: mod_simps)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   385
  also have "\<dots> = (f i + (\<Sum>i\<in>A. f i) mod a) mod a"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   386
    by (simp add: insert.hyps)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   387
  finally show ?case
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   388
    by (simp add: insert.hyps mod_simps)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   389
qed simp_all
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   390
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   391
lemma mod_add_cong:
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   392
  assumes "a mod c = a' mod c"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   393
  assumes "b mod c = b' mod c"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   394
  shows "(a + b) mod c = (a' + b') mod c"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   395
proof -
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   396
  have "(a mod c + b mod c) mod c = (a' mod c + b' mod c) mod c"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   397
    unfolding assms ..
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   398
  then show ?thesis
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   399
    by (simp add: mod_add_eq)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   400
qed
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   401
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   402
text \<open>Multiplication respects modular equivalence.\<close>
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   403
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   404
lemma mod_mult_left_eq [mod_simps]:
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   405
  "((a mod c) * b) mod c = (a * b) mod c"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   406
proof -
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   407
  have "(a * b) mod c = ((a div c * c + a mod c) * b) mod c"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   408
    by (simp only: div_mult_mod_eq)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   409
  also have "\<dots> = (a mod c * b + a div c * b * c) mod c"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   410
    by (simp only: algebra_simps)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   411
  also have "\<dots> = (a mod c * b) mod c"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   412
    by (rule mod_mult_self1)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   413
  finally show ?thesis
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   414
    by (rule sym)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   415
qed
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   416
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   417
lemma mod_mult_right_eq [mod_simps]:
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   418
  "(a * (b mod c)) mod c = (a * b) mod c"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   419
  using mod_mult_left_eq [of b c a] by (simp add: ac_simps)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   420
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   421
lemma mod_mult_eq:
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   422
  "((a mod c) * (b mod c)) mod c = (a * b) mod c"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   423
  by (simp add: mod_mult_left_eq mod_mult_right_eq)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   424
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   425
lemma mod_prod_eq [mod_simps]:
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   426
  "(\<Prod>i\<in>A. f i mod a) mod a = prod f A mod a"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   427
proof (induct A rule: infinite_finite_induct)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   428
  case (insert i A)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   429
  then have "(\<Prod>i\<in>insert i A. f i mod a) mod a
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   430
    = (f i mod a * (\<Prod>i\<in>A. f i mod a)) mod a"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   431
    by simp
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   432
  also have "\<dots> = (f i * ((\<Prod>i\<in>A. f i mod a) mod a)) mod a"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   433
    by (simp add: mod_simps)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   434
  also have "\<dots> = (f i * ((\<Prod>i\<in>A. f i) mod a)) mod a"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   435
    by (simp add: insert.hyps)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   436
  finally show ?case
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   437
    by (simp add: insert.hyps mod_simps)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   438
qed simp_all
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   439
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   440
lemma mod_mult_cong:
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   441
  assumes "a mod c = a' mod c"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   442
  assumes "b mod c = b' mod c"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   443
  shows "(a * b) mod c = (a' * b') mod c"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   444
proof -
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   445
  have "(a mod c * (b mod c)) mod c = (a' mod c * (b' mod c)) mod c"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   446
    unfolding assms ..
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   447
  then show ?thesis
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   448
    by (simp add: mod_mult_eq)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   449
qed
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   450
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   451
text \<open>Exponentiation respects modular equivalence.\<close>
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   452
76224
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
   453
lemma power_mod [mod_simps]:
66806
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   454
  "((a mod b) ^ n) mod b = (a ^ n) mod b"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   455
proof (induct n)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   456
  case 0
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   457
  then show ?case by simp
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   458
next
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   459
  case (Suc n)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   460
  have "(a mod b) ^ Suc n mod b = (a mod b) * ((a mod b) ^ n mod b) mod b"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   461
    by (simp add: mod_mult_right_eq)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   462
  with Suc show ?case
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   463
    by (simp add: mod_mult_left_eq mod_mult_right_eq)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   464
qed
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   465
71413
65ffe9e910d4 more specific class assumptions
haftmann
parents: 71412
diff changeset
   466
lemma power_diff_power_eq:
65ffe9e910d4 more specific class assumptions
haftmann
parents: 71412
diff changeset
   467
  \<open>a ^ m div a ^ n = (if n \<le> m then a ^ (m - n) else 1 div a ^ (n - m))\<close>
65ffe9e910d4 more specific class assumptions
haftmann
parents: 71412
diff changeset
   468
    if \<open>a \<noteq> 0\<close>
65ffe9e910d4 more specific class assumptions
haftmann
parents: 71412
diff changeset
   469
proof (cases \<open>n \<le> m\<close>)
65ffe9e910d4 more specific class assumptions
haftmann
parents: 71412
diff changeset
   470
  case True
65ffe9e910d4 more specific class assumptions
haftmann
parents: 71412
diff changeset
   471
  with that power_diff [symmetric, of a n m] show ?thesis by simp
65ffe9e910d4 more specific class assumptions
haftmann
parents: 71412
diff changeset
   472
next
65ffe9e910d4 more specific class assumptions
haftmann
parents: 71412
diff changeset
   473
  case False
65ffe9e910d4 more specific class assumptions
haftmann
parents: 71412
diff changeset
   474
  then obtain q where n: \<open>n = m + Suc q\<close>
65ffe9e910d4 more specific class assumptions
haftmann
parents: 71412
diff changeset
   475
    by (auto simp add: not_le dest: less_imp_Suc_add)
65ffe9e910d4 more specific class assumptions
haftmann
parents: 71412
diff changeset
   476
  then have \<open>a ^ m div a ^ n = (a ^ m * 1) div (a ^ m * a ^ Suc q)\<close>
65ffe9e910d4 more specific class assumptions
haftmann
parents: 71412
diff changeset
   477
    by (simp add: power_add ac_simps)
65ffe9e910d4 more specific class assumptions
haftmann
parents: 71412
diff changeset
   478
  moreover from that have \<open>a ^ m \<noteq> 0\<close>
65ffe9e910d4 more specific class assumptions
haftmann
parents: 71412
diff changeset
   479
    by simp
65ffe9e910d4 more specific class assumptions
haftmann
parents: 71412
diff changeset
   480
  ultimately have \<open>a ^ m div a ^ n = 1 div a ^ Suc q\<close>
65ffe9e910d4 more specific class assumptions
haftmann
parents: 71412
diff changeset
   481
    by (subst (asm) div_mult_mult1) simp
65ffe9e910d4 more specific class assumptions
haftmann
parents: 71412
diff changeset
   482
  with False n show ?thesis
65ffe9e910d4 more specific class assumptions
haftmann
parents: 71412
diff changeset
   483
    by simp
65ffe9e910d4 more specific class assumptions
haftmann
parents: 71412
diff changeset
   484
qed
65ffe9e910d4 more specific class assumptions
haftmann
parents: 71412
diff changeset
   485
66806
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   486
end
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   487
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   488
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   489
class euclidean_ring_cancel = euclidean_ring + euclidean_semiring_cancel
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   490
begin
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   491
70147
1657688a6406 backed out a93e6472ac9c, which does not bring anything substantial: division_ring is not commutative in multiplication but semidom_divide is
haftmann
parents: 70094
diff changeset
   492
subclass idom_divide ..
1657688a6406 backed out a93e6472ac9c, which does not bring anything substantial: division_ring is not commutative in multiplication but semidom_divide is
haftmann
parents: 70094
diff changeset
   493
1657688a6406 backed out a93e6472ac9c, which does not bring anything substantial: division_ring is not commutative in multiplication but semidom_divide is
haftmann
parents: 70094
diff changeset
   494
lemma div_minus_minus [simp]: "(- a) div (- b) = a div b"
1657688a6406 backed out a93e6472ac9c, which does not bring anything substantial: division_ring is not commutative in multiplication but semidom_divide is
haftmann
parents: 70094
diff changeset
   495
  using div_mult_mult1 [of "- 1" a b] by simp
66806
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   496
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   497
lemma mod_minus_minus [simp]: "(- a) mod (- b) = - (a mod b)"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   498
  using mod_mult_mult1 [of "- 1" a b] by simp
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   499
70147
1657688a6406 backed out a93e6472ac9c, which does not bring anything substantial: division_ring is not commutative in multiplication but semidom_divide is
haftmann
parents: 70094
diff changeset
   500
lemma div_minus_right: "a div (- b) = (- a) div b"
1657688a6406 backed out a93e6472ac9c, which does not bring anything substantial: division_ring is not commutative in multiplication but semidom_divide is
haftmann
parents: 70094
diff changeset
   501
  using div_minus_minus [of "- a" b] by simp
1657688a6406 backed out a93e6472ac9c, which does not bring anything substantial: division_ring is not commutative in multiplication but semidom_divide is
haftmann
parents: 70094
diff changeset
   502
66806
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   503
lemma mod_minus_right: "a mod (- b) = - ((- a) mod b)"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   504
  using mod_minus_minus [of "- a" b] by simp
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   505
70147
1657688a6406 backed out a93e6472ac9c, which does not bring anything substantial: division_ring is not commutative in multiplication but semidom_divide is
haftmann
parents: 70094
diff changeset
   506
lemma div_minus1_right [simp]: "a div (- 1) = - a"
1657688a6406 backed out a93e6472ac9c, which does not bring anything substantial: division_ring is not commutative in multiplication but semidom_divide is
haftmann
parents: 70094
diff changeset
   507
  using div_minus_right [of a 1] by simp
1657688a6406 backed out a93e6472ac9c, which does not bring anything substantial: division_ring is not commutative in multiplication but semidom_divide is
haftmann
parents: 70094
diff changeset
   508
66806
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   509
lemma mod_minus1_right [simp]: "a mod (- 1) = 0"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   510
  using mod_minus_right [of a 1] by simp
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   511
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   512
text \<open>Negation respects modular equivalence.\<close>
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   513
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   514
lemma mod_minus_eq [mod_simps]:
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   515
  "(- (a mod b)) mod b = (- a) mod b"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   516
proof -
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   517
  have "(- a) mod b = (- (a div b * b + a mod b)) mod b"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   518
    by (simp only: div_mult_mod_eq)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   519
  also have "\<dots> = (- (a mod b) + - (a div b) * b) mod b"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   520
    by (simp add: ac_simps)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   521
  also have "\<dots> = (- (a mod b)) mod b"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   522
    by (rule mod_mult_self1)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   523
  finally show ?thesis
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   524
    by (rule sym)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   525
qed
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   526
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   527
lemma mod_minus_cong:
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   528
  assumes "a mod b = a' mod b"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   529
  shows "(- a) mod b = (- a') mod b"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   530
proof -
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   531
  have "(- (a mod b)) mod b = (- (a' mod b)) mod b"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   532
    unfolding assms ..
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   533
  then show ?thesis
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   534
    by (simp add: mod_minus_eq)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   535
qed
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   536
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   537
text \<open>Subtraction respects modular equivalence.\<close>
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   538
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   539
lemma mod_diff_left_eq [mod_simps]:
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   540
  "(a mod c - b) mod c = (a - b) mod c"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   541
  using mod_add_cong [of a c "a mod c" "- b" "- b"]
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   542
  by simp
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   543
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   544
lemma mod_diff_right_eq [mod_simps]:
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   545
  "(a - b mod c) mod c = (a - b) mod c"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   546
  using mod_add_cong [of a c a "- b" "- (b mod c)"] mod_minus_cong [of "b mod c" c b]
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   547
  by simp
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   548
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   549
lemma mod_diff_eq:
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   550
  "(a mod c - b mod c) mod c = (a - b) mod c"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   551
  using mod_add_cong [of a c "a mod c" "- b" "- (b mod c)"] mod_minus_cong [of "b mod c" c b]
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   552
  by simp
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   553
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   554
lemma mod_diff_cong:
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   555
  assumes "a mod c = a' mod c"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   556
  assumes "b mod c = b' mod c"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   557
  shows "(a - b) mod c = (a' - b') mod c"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   558
  using assms mod_add_cong [of a c a' "- b" "- b'"] mod_minus_cong [of b c "b'"]
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   559
  by simp
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   560
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   561
lemma minus_mod_self2 [simp]:
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   562
  "(a - b) mod b = a mod b"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   563
  using mod_diff_right_eq [of a b b]
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   564
  by (simp add: mod_diff_right_eq)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   565
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   566
lemma minus_mod_self1 [simp]:
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   567
  "(b - a) mod b = - a mod b"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   568
  using mod_add_self2 [of "- a" b] by simp
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   569
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   570
lemma mod_eq_dvd_iff:
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   571
  "a mod c = b mod c \<longleftrightarrow> c dvd a - b" (is "?P \<longleftrightarrow> ?Q")
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   572
proof
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   573
  assume ?P
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   574
  then have "(a mod c - b mod c) mod c = 0"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   575
    by simp
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   576
  then show ?Q
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   577
    by (simp add: dvd_eq_mod_eq_0 mod_simps)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   578
next
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   579
  assume ?Q
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   580
  then obtain d where d: "a - b = c * d" ..
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   581
  then have "a = c * d + b"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   582
    by (simp add: algebra_simps)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   583
  then show ?P by simp
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   584
qed
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   585
66837
6ba663ff2b1c tuned proofs
haftmann
parents: 66817
diff changeset
   586
lemma mod_eqE:
6ba663ff2b1c tuned proofs
haftmann
parents: 66817
diff changeset
   587
  assumes "a mod c = b mod c"
6ba663ff2b1c tuned proofs
haftmann
parents: 66817
diff changeset
   588
  obtains d where "b = a + c * d"
6ba663ff2b1c tuned proofs
haftmann
parents: 66817
diff changeset
   589
proof -
6ba663ff2b1c tuned proofs
haftmann
parents: 66817
diff changeset
   590
  from assms have "c dvd a - b"
6ba663ff2b1c tuned proofs
haftmann
parents: 66817
diff changeset
   591
    by (simp add: mod_eq_dvd_iff)
6ba663ff2b1c tuned proofs
haftmann
parents: 66817
diff changeset
   592
  then obtain d where "a - b = c * d" ..
6ba663ff2b1c tuned proofs
haftmann
parents: 66817
diff changeset
   593
  then have "b = a + c * - d"
6ba663ff2b1c tuned proofs
haftmann
parents: 66817
diff changeset
   594
    by (simp add: algebra_simps)
6ba663ff2b1c tuned proofs
haftmann
parents: 66817
diff changeset
   595
  with that show thesis .
6ba663ff2b1c tuned proofs
haftmann
parents: 66817
diff changeset
   596
qed
6ba663ff2b1c tuned proofs
haftmann
parents: 66817
diff changeset
   597
67051
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
   598
lemma invertible_coprime:
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
   599
  "coprime a c" if "a * b mod c = 1"
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
   600
  by (rule coprimeI) (use that dvd_mod_iff [of _ c "a * b"] in auto)
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
   601
66806
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   602
end
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   603
76224
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
   604
64785
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   605
subsection \<open>Uniquely determined division\<close>
75875
48d032035744 streamlined primitive definitions for integer division
haftmann
parents: 75669
diff changeset
   606
76224
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
   607
class unique_euclidean_semiring = euclidean_semiring +
76053
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   608
  assumes euclidean_size_mult: \<open>euclidean_size (a * b) = euclidean_size a * euclidean_size b\<close>
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   609
  fixes division_segment :: \<open>'a \<Rightarrow> 'a\<close>
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   610
  assumes is_unit_division_segment [simp]: \<open>is_unit (division_segment a)\<close>
66838
17989f6bc7b2 clarified uniqueness criterion for euclidean rings
haftmann
parents: 66837
diff changeset
   611
    and division_segment_mult:
76053
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   612
    \<open>a \<noteq> 0 \<Longrightarrow> b \<noteq> 0 \<Longrightarrow> division_segment (a * b) = division_segment a * division_segment b\<close>
66838
17989f6bc7b2 clarified uniqueness criterion for euclidean rings
haftmann
parents: 66837
diff changeset
   613
    and division_segment_mod:
76053
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   614
    \<open>b \<noteq> 0 \<Longrightarrow> \<not> b dvd a \<Longrightarrow> division_segment (a mod b) = division_segment b\<close>
64785
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   615
  assumes div_bounded:
76053
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   616
    \<open>b \<noteq> 0 \<Longrightarrow> division_segment r = division_segment b
64785
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   617
    \<Longrightarrow> euclidean_size r < euclidean_size b
76053
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   618
    \<Longrightarrow> (q * b + r) div b = q\<close>
64785
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   619
begin
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   620
66839
909ba5ed93dd clarified parity
haftmann
parents: 66838
diff changeset
   621
lemma division_segment_not_0 [simp]:
76053
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   622
  \<open>division_segment a \<noteq> 0\<close>
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   623
  using is_unit_division_segment [of a] is_unitE [of \<open>division_segment a\<close>] by blast
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   624
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   625
lemma euclidean_relationI [case_names by0 divides euclidean_relation]:
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   626
  \<open>(a div b, a mod b) = (q, r)\<close>
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   627
    if by0: \<open>b = 0 \<Longrightarrow> q = 0 \<and> r = a\<close>
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   628
    and divides: \<open>b \<noteq> 0 \<Longrightarrow> b dvd a \<Longrightarrow> r = 0 \<and> a = q * b\<close>
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   629
    and euclidean_relation: \<open>b \<noteq> 0 \<Longrightarrow> \<not> b dvd a \<Longrightarrow> division_segment r = division_segment b
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   630
      \<and> euclidean_size r < euclidean_size b \<and> a = q * b + r\<close>
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   631
proof (cases \<open>b = 0\<close>)
64785
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   632
  case True
76053
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   633
  with by0 show ?thesis
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   634
    by simp
64785
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   635
next
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   636
  case False
76053
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   637
  show ?thesis
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   638
  proof (cases \<open>b dvd a\<close>)
64785
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   639
    case True
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   640
    with \<open>b \<noteq> 0\<close> divides
76053
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   641
    show ?thesis
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   642
      by simp
64785
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   643
  next
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   644
    case False
76053
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   645
    with \<open>b \<noteq> 0\<close> euclidean_relation
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   646
    have \<open>division_segment r = division_segment b\<close>
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   647
      \<open>euclidean_size r < euclidean_size b\<close> \<open>a = q * b + r\<close>
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   648
      by simp_all
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   649
    from \<open>b \<noteq> 0\<close> \<open>division_segment r = division_segment b\<close>
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   650
      \<open>euclidean_size r < euclidean_size b\<close>
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   651
    have \<open>(q * b + r) div b = q\<close>
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   652
      by (rule div_bounded)
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   653
    with \<open>a = q * b + r\<close>
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   654
    have \<open>q = a div b\<close>
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   655
      by simp
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   656
    from \<open>a = q * b + r\<close>
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   657
    have \<open>a div b * b + a mod b = q * b + r\<close>
64785
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   658
      by (simp add: div_mult_mod_eq)
76053
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   659
    with \<open>q = a div b\<close>
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   660
    have \<open>q * b + a mod b = q * b + r\<close>
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   661
      by simp
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   662
    then have \<open>r = a mod b\<close>
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   663
      by simp
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   664
    with \<open>q = a div b\<close>
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   665
    show ?thesis
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   666
      by simp
64785
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   667
  qed
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   668
qed
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   669
66806
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   670
subclass euclidean_semiring_cancel
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   671
proof
76053
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   672
  fix a b c
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   673
  assume \<open>b \<noteq> 0\<close>
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   674
  have \<open>((a + c * b) div b, (a + c * b) mod b) = (c + a div b, a mod b)\<close>
76245
4111c94657b4 slightly less abusive proof pattern
haftmann
parents: 76231
diff changeset
   675
  proof (induction rule: euclidean_relationI)
66806
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   676
    case by0
76053
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   677
    with \<open>b \<noteq> 0\<close>
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   678
    show ?case
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   679
      by simp
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   680
  next
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   681
    case divides
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   682
    then show ?case
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   683
      by (simp add: algebra_simps dvd_add_left_iff)
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   684
  next
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   685
    case euclidean_relation
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   686
    then have \<open>\<not> b dvd a\<close>
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   687
      by (simp add: dvd_add_left_iff)
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   688
    have \<open>a mod b + (b * c + b * (a div b)) = b * c + ((a div b) * b + a mod b)\<close>
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   689
      by (simp add: ac_simps)
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   690
    with \<open>b \<noteq> 0\<close> have *: \<open>a mod b + (b * c + b * (a div b)) = b * c + a\<close>
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   691
      by (simp add: div_mult_mod_eq)
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   692
    from \<open>\<not> b dvd a\<close> euclidean_relation show ?case
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   693
      by (simp_all add: algebra_simps division_segment_mod mod_size_less *)
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   694
  qed
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   695
  then show \<open>(a + c * b) div b = c + a div b\<close>
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   696
    by simp
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   697
next
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   698
  fix a b c
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   699
  assume \<open>c \<noteq> 0\<close>
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   700
  have \<open>((c * a) div (c * b), (c * a) mod (c * b)) = (a div b, c * (a mod b))\<close>
76245
4111c94657b4 slightly less abusive proof pattern
haftmann
parents: 76231
diff changeset
   701
  proof (induction rule: euclidean_relationI)
76053
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   702
    case by0
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   703
    with \<open>c \<noteq> 0\<close> show ?case
66806
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   704
      by simp
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   705
  next
76053
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   706
    case divides
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   707
    then show ?case
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   708
      by (auto simp add: algebra_simps)
66806
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   709
  next
76053
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   710
    case euclidean_relation
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   711
    then have \<open>b \<noteq> 0\<close> \<open>a mod b \<noteq> 0\<close>
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   712
      by (simp_all add: mod_eq_0_iff_dvd)
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   713
    have \<open>c * (a mod b) + b * (c * (a div b)) = c * ((a div b) * b + a mod b)\<close>
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   714
      by (simp add: algebra_simps)
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   715
    with \<open>b \<noteq> 0\<close> have *: \<open>c * (a mod b) + b * (c * (a div b)) = c * a\<close>
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   716
      by (simp add: div_mult_mod_eq)
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   717
    from \<open>b \<noteq> 0\<close> \<open>c \<noteq> 0\<close> have \<open>euclidean_size c * euclidean_size (a mod b)
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   718
      < euclidean_size c * euclidean_size b\<close>
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   719
      using mod_size_less [of b a] by simp
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   720
    with euclidean_relation \<open>b \<noteq> 0\<close> \<open>a mod b \<noteq> 0\<close> show ?case
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   721
      by (simp add: algebra_simps division_segment_mult division_segment_mod euclidean_size_mult *)
66806
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   722
  qed
76053
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   723
  then show \<open>(c * a) div (c * b) = a div b\<close>
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   724
    by simp
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   725
qed
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   726
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   727
lemma div_eq_0_iff:
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   728
  \<open>a div b = 0 \<longleftrightarrow> euclidean_size a < euclidean_size b \<or> b = 0\<close> (is "_ \<longleftrightarrow> ?P")
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   729
  if \<open>division_segment a = division_segment b\<close>
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   730
proof (cases \<open>a = 0 \<or> b = 0\<close>)
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   731
  case True
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   732
  then show ?thesis by auto
66806
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   733
next
76053
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   734
  case False
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   735
  then have \<open>a \<noteq> 0\<close> \<open>b \<noteq> 0\<close>
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   736
    by simp_all
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   737
  have \<open>a div b = 0 \<longleftrightarrow> euclidean_size a < euclidean_size b\<close>
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   738
  proof
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   739
    assume \<open>a div b = 0\<close>
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   740
    then have \<open>a mod b = a\<close>
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   741
      using div_mult_mod_eq [of a b] by simp
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   742
    with \<open>b \<noteq> 0\<close> mod_size_less [of b a]
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   743
    show \<open>euclidean_size a < euclidean_size b\<close>
66806
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   744
      by simp
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   745
  next
76053
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   746
    assume \<open>euclidean_size a < euclidean_size b\<close>
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   747
    have \<open>(a div b, a mod b) = (0, a)\<close>
76245
4111c94657b4 slightly less abusive proof pattern
haftmann
parents: 76231
diff changeset
   748
    proof (induction rule: euclidean_relationI)
76053
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   749
      case by0
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   750
      show ?case
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   751
        by simp
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   752
    next
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   753
      case divides
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   754
      with \<open>euclidean_size a < euclidean_size b\<close> show ?case
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   755
        using dvd_imp_size_le [of b a] \<open>a \<noteq> 0\<close> by simp
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   756
    next
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   757
      case euclidean_relation
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   758
      with \<open>euclidean_size a < euclidean_size b\<close> that
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   759
      show ?case
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   760
        by simp
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   761
    qed
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   762
    then show \<open>a div b = 0\<close>
66806
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   763
      by simp
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   764
  qed
76053
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   765
  with \<open>b \<noteq> 0\<close> show ?thesis
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   766
    by simp
66806
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   767
qed
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   768
66814
a24cde9588bb generalized some rules
haftmann
parents: 66813
diff changeset
   769
lemma div_mult1_eq:
76053
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   770
  \<open>(a * b) div c = a * (b div c) + a * (b mod c) div c\<close>
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   771
proof -
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   772
  have *: \<open>(a * b) mod c + (a * (c * (b div c)) + c * (a * (b mod c) div c)) = a * b\<close> (is \<open>?A + (?B + ?C) = _\<close>)
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   773
  proof -
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   774
    have \<open>?A = a * (b mod c) mod c\<close>
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   775
      by (simp add: mod_mult_right_eq)
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   776
    then have \<open>?C + ?A = a * (b mod c)\<close>
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   777
      by (simp add: mult_div_mod_eq)
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   778
    then have \<open>?B + (?C + ?A) = a * (c * (b div c) + (b mod c))\<close>
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   779
      by (simp add: algebra_simps)
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   780
    also have \<open>\<dots> = a * b\<close>
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   781
      by (simp add: mult_div_mod_eq)
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   782
    finally show ?thesis
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   783
      by (simp add: algebra_simps)
66814
a24cde9588bb generalized some rules
haftmann
parents: 66813
diff changeset
   784
  qed
76053
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   785
  have \<open>((a * b) div c, (a * b) mod c) = (a * (b div c) + a * (b mod c) div c, (a * b) mod c)\<close>
76245
4111c94657b4 slightly less abusive proof pattern
haftmann
parents: 76231
diff changeset
   786
  proof (induction rule: euclidean_relationI)
76053
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   787
    case by0
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   788
    then show ?case by simp
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   789
  next
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   790
    case divides
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   791
    with * show ?case
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   792
      by (simp add: algebra_simps)
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   793
  next
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   794
    case euclidean_relation
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   795
    with * show ?case
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   796
      by (simp add: division_segment_mod mod_size_less algebra_simps)
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   797
  qed
66814
a24cde9588bb generalized some rules
haftmann
parents: 66813
diff changeset
   798
  then show ?thesis
a24cde9588bb generalized some rules
haftmann
parents: 66813
diff changeset
   799
    by simp
a24cde9588bb generalized some rules
haftmann
parents: 66813
diff changeset
   800
qed
a24cde9588bb generalized some rules
haftmann
parents: 66813
diff changeset
   801
a24cde9588bb generalized some rules
haftmann
parents: 66813
diff changeset
   802
lemma div_add1_eq:
76053
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   803
  \<open>(a + b) div c = a div c + b div c + (a mod c + b mod c) div c\<close>
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   804
proof -
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   805
  have *: \<open>(a + b) mod c + (c * (a div c) + (c * (b div c) + c * ((a mod c + b mod c) div c))) = a + b\<close>
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   806
    (is \<open>?A + (?B + (?C + ?D)) = _\<close>)
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   807
  proof -
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   808
    have \<open>?A + (?B + (?C + ?D)) = ?A + ?D + (?B + ?C)\<close>
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   809
      by (simp add: ac_simps)
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   810
    also have \<open>?A + ?D = (a mod c + b mod c) mod c + ?D\<close>
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   811
      by (simp add: mod_add_eq)
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   812
    also have \<open>\<dots> = a mod c + b mod c\<close>
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   813
      by (simp add: mod_mult_div_eq)
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   814
    finally have \<open>?A + (?B + (?C + ?D)) = (a mod c + ?B) + (b mod c + ?C)\<close>
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   815
      by (simp add: ac_simps)
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   816
    then show ?thesis
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   817
      by (simp add: mod_mult_div_eq)
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   818
  qed
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   819
  have \<open>((a + b) div c, (a + b) mod c) = (a div c + b div c + (a mod c + b mod c) div c, (a + b) mod c)\<close>
76245
4111c94657b4 slightly less abusive proof pattern
haftmann
parents: 76231
diff changeset
   820
  proof (induction rule: euclidean_relationI)
76053
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   821
    case by0
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   822
    then show ?case
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   823
      by simp
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   824
  next
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   825
    case divides
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   826
    with * show ?case
66814
a24cde9588bb generalized some rules
haftmann
parents: 66813
diff changeset
   827
      by (simp add: algebra_simps)
76053
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   828
  next
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   829
    case euclidean_relation
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   830
    with * show ?case
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   831
      by (simp add: division_segment_mod mod_size_less algebra_simps)
66814
a24cde9588bb generalized some rules
haftmann
parents: 66813
diff changeset
   832
  qed
a24cde9588bb generalized some rules
haftmann
parents: 66813
diff changeset
   833
  then show ?thesis
a24cde9588bb generalized some rules
haftmann
parents: 66813
diff changeset
   834
    by simp
a24cde9588bb generalized some rules
haftmann
parents: 66813
diff changeset
   835
qed
a24cde9588bb generalized some rules
haftmann
parents: 66813
diff changeset
   836
64785
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   837
end
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   838
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   839
class unique_euclidean_ring = euclidean_ring + unique_euclidean_semiring
66806
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   840
begin
76224
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
   841
66806
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   842
subclass euclidean_ring_cancel ..
64785
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   843
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   844
end
66806
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   845
66808
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   846
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 68536
diff changeset
   847
subsection \<open>Euclidean division on \<^typ>\<open>nat\<close>\<close>
66808
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   848
66816
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
   849
instantiation nat :: normalization_semidom
66808
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   850
begin
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   851
76053
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   852
definition normalize_nat :: \<open>nat \<Rightarrow> nat\<close>
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   853
  where [simp]: \<open>normalize = (id :: nat \<Rightarrow> nat)\<close>
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   854
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   855
definition unit_factor_nat :: \<open>nat \<Rightarrow> nat\<close>
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   856
  where \<open>unit_factor n = of_bool (n > 0)\<close> for n :: nat
66808
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   857
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   858
lemma unit_factor_simps [simp]:
76053
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   859
  \<open>unit_factor 0 = (0::nat)\<close>
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   860
  \<open>unit_factor (Suc n) = 1\<close>
66808
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   861
  by (simp_all add: unit_factor_nat_def)
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   862
76053
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   863
definition divide_nat :: \<open>nat \<Rightarrow> nat \<Rightarrow> nat\<close>
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   864
  where \<open>m div n = (if n = 0 then 0 else Max {k. k * n \<le> m})\<close> for m n :: nat
66816
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
   865
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
   866
instance
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
   867
  by standard (auto simp add: divide_nat_def ac_simps unit_factor_nat_def intro: Max_eqI)
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
   868
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
   869
end
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
   870
67051
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
   871
lemma coprime_Suc_0_left [simp]:
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
   872
  "coprime (Suc 0) n"
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
   873
  using coprime_1_left [of n] by simp
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
   874
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
   875
lemma coprime_Suc_0_right [simp]:
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
   876
  "coprime n (Suc 0)"
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
   877
  using coprime_1_right [of n] by simp
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
   878
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
   879
lemma coprime_common_divisor_nat: "coprime a b \<Longrightarrow> x dvd a \<Longrightarrow> x dvd b \<Longrightarrow> x = 1"
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
   880
  for a b :: nat
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
   881
  by (drule coprime_common_divisor [of _ _ x]) simp_all
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
   882
66816
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
   883
instantiation nat :: unique_euclidean_semiring
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
   884
begin
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
   885
76053
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   886
definition euclidean_size_nat :: \<open>nat \<Rightarrow> nat\<close>
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   887
  where [simp]: \<open>euclidean_size_nat = id\<close>
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   888
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   889
definition division_segment_nat :: \<open>nat \<Rightarrow> nat\<close>
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   890
  where [simp]: \<open>division_segment n = 1\<close> for n :: nat
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   891
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   892
definition modulo_nat :: \<open>nat \<Rightarrow> nat \<Rightarrow> nat\<close>
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   893
  where \<open>m mod n = m - (m div n * n)\<close> for m n :: nat
66808
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   894
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   895
instance proof
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   896
  fix m n :: nat
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   897
  have ex: "\<exists>k. k * n \<le> l" for l :: nat
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   898
    by (rule exI [of _ 0]) simp
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   899
  have fin: "finite {k. k * n \<le> l}" if "n > 0" for l
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   900
  proof -
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   901
    from that have "{k. k * n \<le> l} \<subseteq> {k. k \<le> l}"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   902
      by (cases n) auto
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   903
    then show ?thesis
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   904
      by (rule finite_subset) simp
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   905
  qed
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   906
  have mult_div_unfold: "n * (m div n) = Max {l. l \<le> m \<and> n dvd l}"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   907
  proof (cases "n = 0")
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   908
    case True
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   909
    moreover have "{l. l = 0 \<and> l \<le> m} = {0::nat}"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   910
      by auto
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   911
    ultimately show ?thesis
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   912
      by simp
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   913
  next
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   914
    case False
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   915
    with ex [of m] fin have "n * Max {k. k * n \<le> m} = Max (times n ` {k. k * n \<le> m})"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   916
      by (auto simp add: nat_mult_max_right intro: hom_Max_commute)
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   917
    also have "times n ` {k. k * n \<le> m} = {l. l \<le> m \<and> n dvd l}"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   918
      by (auto simp add: ac_simps elim!: dvdE)
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   919
    finally show ?thesis
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   920
      using False by (simp add: divide_nat_def ac_simps)
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   921
  qed
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   922
  have less_eq: "m div n * n \<le> m"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   923
    by (auto simp add: mult_div_unfold ac_simps intro: Max.boundedI)
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   924
  then show "m div n * n + m mod n = m"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   925
    by (simp add: modulo_nat_def)
76224
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
   926
  assume "n \<noteq> 0"
66808
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   927
  show "euclidean_size (m mod n) < euclidean_size n"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   928
  proof -
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   929
    have "m < Suc (m div n) * n"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   930
    proof (rule ccontr)
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   931
      assume "\<not> m < Suc (m div n) * n"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   932
      then have "Suc (m div n) * n \<le> m"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   933
        by (simp add: not_less)
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   934
      moreover from \<open>n \<noteq> 0\<close> have "Max {k. k * n \<le> m} < Suc (m div n)"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   935
        by (simp add: divide_nat_def)
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   936
      with \<open>n \<noteq> 0\<close> ex fin have "\<And>k. k * n \<le> m \<Longrightarrow> k < Suc (m div n)"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   937
        by auto
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   938
      ultimately have "Suc (m div n) < Suc (m div n)"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   939
        by blast
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   940
      then show False
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   941
        by simp
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   942
    qed
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   943
    with \<open>n \<noteq> 0\<close> show ?thesis
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   944
      by (simp add: modulo_nat_def)
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   945
  qed
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   946
  show "euclidean_size m \<le> euclidean_size (m * n)"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   947
    using \<open>n \<noteq> 0\<close> by (cases n) simp_all
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   948
  fix q r :: nat
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   949
  show "(q * n + r) div n = q" if "euclidean_size r < euclidean_size n"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   950
  proof -
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   951
    from that have "r < n"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   952
      by simp
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   953
    have "k \<le> q" if "k * n \<le> q * n + r" for k
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   954
    proof (rule ccontr)
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   955
      assume "\<not> k \<le> q"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   956
      then have "q < k"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   957
        by simp
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   958
      then obtain l where "k = Suc (q + l)"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   959
        by (auto simp add: less_iff_Suc_add)
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   960
      with \<open>r < n\<close> that show False
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   961
        by (simp add: algebra_simps)
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   962
    qed
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   963
    with \<open>n \<noteq> 0\<close> ex fin show ?thesis
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   964
      by (auto simp add: divide_nat_def Max_eq_iff)
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   965
  qed
66816
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
   966
qed simp_all
66808
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   967
66806
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   968
end
66808
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   969
76141
e7497a1de8b9 more concise instance-specific rules on euclidean relation
haftmann
parents: 76053
diff changeset
   970
lemma euclidean_relation_natI [case_names by0 divides euclidean_relation]:
e7497a1de8b9 more concise instance-specific rules on euclidean relation
haftmann
parents: 76053
diff changeset
   971
  \<open>(m div n, m mod n) = (q, r)\<close>
e7497a1de8b9 more concise instance-specific rules on euclidean relation
haftmann
parents: 76053
diff changeset
   972
    if by0: \<open>n = 0 \<Longrightarrow> q = 0 \<and> r = m\<close>
e7497a1de8b9 more concise instance-specific rules on euclidean relation
haftmann
parents: 76053
diff changeset
   973
    and divides: \<open>n > 0 \<Longrightarrow> n dvd m \<Longrightarrow> r = 0 \<and> m = q * n\<close>
e7497a1de8b9 more concise instance-specific rules on euclidean relation
haftmann
parents: 76053
diff changeset
   974
    and euclidean_relation: \<open>n > 0 \<Longrightarrow> \<not> n dvd m \<Longrightarrow> r < n \<and> m = q * n + r\<close> for m n q r :: nat
e7497a1de8b9 more concise instance-specific rules on euclidean relation
haftmann
parents: 76053
diff changeset
   975
  by (rule euclidean_relationI) (use that in simp_all)
e7497a1de8b9 more concise instance-specific rules on euclidean relation
haftmann
parents: 76053
diff changeset
   976
75881
83e4b6a5e7de streamlined theorems
haftmann
parents: 75880
diff changeset
   977
lemma div_nat_eqI:
76053
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   978
  \<open>m div n = q\<close> if \<open>n * q \<le> m\<close> and \<open>m < n * Suc q\<close> for m n q :: nat
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   979
proof -
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   980
  have \<open>(m div n, m mod n) = (q, m - n * q)\<close>
76245
4111c94657b4 slightly less abusive proof pattern
haftmann
parents: 76231
diff changeset
   981
  proof (induction rule: euclidean_relation_natI)
76053
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   982
    case by0
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   983
    with that show ?case
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   984
      by simp
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   985
  next
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   986
    case divides
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   987
    from \<open>n dvd m\<close> obtain s where \<open>m = n * s\<close> ..
76141
e7497a1de8b9 more concise instance-specific rules on euclidean relation
haftmann
parents: 76053
diff changeset
   988
    with \<open>n > 0\<close> that have \<open>s < Suc q\<close>
76053
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   989
      by (simp only: mult_less_cancel1)
76141
e7497a1de8b9 more concise instance-specific rules on euclidean relation
haftmann
parents: 76053
diff changeset
   990
    with \<open>m = n * s\<close> \<open>n > 0\<close> that have \<open>q = s\<close>
76053
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   991
      by simp
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   992
    with \<open>m = n * s\<close> show ?case
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   993
      by (simp add: ac_simps)
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   994
  next
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   995
    case euclidean_relation
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   996
    with that show ?case
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   997
      by (simp add: ac_simps)
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   998
  qed
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
   999
  then show ?thesis
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
  1000
    by simp
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
  1001
qed
75881
83e4b6a5e7de streamlined theorems
haftmann
parents: 75880
diff changeset
  1002
83e4b6a5e7de streamlined theorems
haftmann
parents: 75880
diff changeset
  1003
lemma mod_nat_eqI:
76053
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
  1004
  \<open>m mod n = r\<close> if \<open>r < n\<close> and \<open>r \<le> m\<close> and \<open>n dvd m - r\<close> for m n r :: nat
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
  1005
proof -
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
  1006
  have \<open>(m div n, m mod n) = ((m - r) div n, r)\<close>
76245
4111c94657b4 slightly less abusive proof pattern
haftmann
parents: 76231
diff changeset
  1007
  proof (induction rule: euclidean_relation_natI)
76053
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
  1008
    case by0
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
  1009
    with that show ?case
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
  1010
      by simp
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
  1011
  next
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
  1012
    case divides
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
  1013
    from that dvd_minus_add [of r \<open>m\<close> 1 n]
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
  1014
    have \<open>n dvd m + (n - r)\<close>
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
  1015
      by simp
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
  1016
    with divides have \<open>n dvd n - r\<close>
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
  1017
      by (simp add: dvd_add_right_iff)
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
  1018
    then have \<open>n \<le> n - r\<close>
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
  1019
      by (rule dvd_imp_le) (use \<open>r < n\<close> in simp)
76141
e7497a1de8b9 more concise instance-specific rules on euclidean relation
haftmann
parents: 76053
diff changeset
  1020
    with \<open>n > 0\<close> have \<open>r = 0\<close>
76053
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
  1021
      by simp
76141
e7497a1de8b9 more concise instance-specific rules on euclidean relation
haftmann
parents: 76053
diff changeset
  1022
    with \<open>n > 0\<close> that show ?case
76053
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
  1023
      by simp
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
  1024
  next
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
  1025
    case euclidean_relation
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
  1026
    with that show ?case
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
  1027
      by (simp add: ac_simps)
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
  1028
  qed
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
  1029
  then show ?thesis
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
  1030
    by simp
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
  1031
qed
75881
83e4b6a5e7de streamlined theorems
haftmann
parents: 75880
diff changeset
  1032
66808
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1033
text \<open>Tool support\<close>
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1034
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1035
ML \<open>
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1036
structure Cancel_Div_Mod_Nat = Cancel_Div_Mod
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1037
(
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 68536
diff changeset
  1038
  val div_name = \<^const_name>\<open>divide\<close>;
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 68536
diff changeset
  1039
  val mod_name = \<^const_name>\<open>modulo\<close>;
66808
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1040
  val mk_binop = HOLogic.mk_binop;
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 68536
diff changeset
  1041
  val dest_plus = HOLogic.dest_bin \<^const_name>\<open>Groups.plus\<close> HOLogic.natT;
66813
351142796345 avoid variant of mk_sum
haftmann
parents: 66810
diff changeset
  1042
  val mk_sum = Arith_Data.mk_sum;
66808
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1043
  fun dest_sum tm =
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1044
    if HOLogic.is_zero tm then []
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1045
    else
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1046
      (case try HOLogic.dest_Suc tm of
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1047
        SOME t => HOLogic.Suc_zero :: dest_sum t
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1048
      | NONE =>
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1049
          (case try dest_plus tm of
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1050
            SOME (t, u) => dest_sum t @ dest_sum u
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1051
          | NONE => [tm]));
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1052
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1053
  val div_mod_eqs = map mk_meta_eq @{thms cancel_div_mod_rules};
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1054
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1055
  val prove_eq_sums = Arith_Data.prove_conv2 all_tac
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1056
    (Arith_Data.simp_all_tac @{thms add_0_left add_0_right ac_simps})
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1057
)
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1058
\<close>
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1059
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1060
simproc_setup cancel_div_mod_nat ("(m::nat) + n") =
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1061
  \<open>K Cancel_Div_Mod_Nat.proc\<close>
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1062
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1063
lemma div_mult_self_is_m [simp]:
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1064
  "m * n div n = m" if "n > 0" for m n :: nat
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1065
  using that by simp
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1066
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1067
lemma div_mult_self1_is_m [simp]:
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1068
  "n * m div n = m" if "n > 0" for m n :: nat
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1069
  using that by simp
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1070
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1071
lemma mod_less_divisor [simp]:
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1072
  "m mod n < n" if "n > 0" for m n :: nat
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1073
  using mod_size_less [of n m] that by simp
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1074
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1075
lemma mod_le_divisor [simp]:
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1076
  "m mod n \<le> n" if "n > 0" for m n :: nat
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1077
  using that by (auto simp add: le_less)
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1078
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1079
lemma div_times_less_eq_dividend [simp]:
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1080
  "m div n * n \<le> m" for m n :: nat
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1081
  by (simp add: minus_mod_eq_div_mult [symmetric])
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1082
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1083
lemma times_div_less_eq_dividend [simp]:
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1084
  "n * (m div n) \<le> m" for m n :: nat
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1085
  using div_times_less_eq_dividend [of m n]
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1086
  by (simp add: ac_simps)
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1087
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1088
lemma dividend_less_div_times:
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1089
  "m < n + (m div n) * n" if "0 < n" for m n :: nat
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1090
proof -
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1091
  from that have "m mod n < n"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1092
    by simp
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1093
  then show ?thesis
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1094
    by (simp add: minus_mod_eq_div_mult [symmetric])
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1095
qed
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1096
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1097
lemma dividend_less_times_div:
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1098
  "m < n + n * (m div n)" if "0 < n" for m n :: nat
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1099
  using dividend_less_div_times [of n m] that
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1100
  by (simp add: ac_simps)
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1101
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1102
lemma mod_Suc_le_divisor [simp]:
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1103
  "m mod Suc n \<le> n"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1104
  using mod_less_divisor [of "Suc n" m] by arith
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1105
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1106
lemma mod_less_eq_dividend [simp]:
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1107
  "m mod n \<le> m" for m n :: nat
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1108
proof (rule add_leD2)
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1109
  from div_mult_mod_eq have "m div n * n + m mod n = m" .
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1110
  then show "m div n * n + m mod n \<le> m" by auto
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1111
qed
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1112
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1113
lemma
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1114
  div_less [simp]: "m div n = 0"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1115
  and mod_less [simp]: "m mod n = m"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1116
  if "m < n" for m n :: nat
76053
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
  1117
  using that by (auto intro: div_nat_eqI mod_nat_eqI)
76224
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  1118
75881
83e4b6a5e7de streamlined theorems
haftmann
parents: 75880
diff changeset
  1119
lemma split_div:
83e4b6a5e7de streamlined theorems
haftmann
parents: 75880
diff changeset
  1120
  \<open>P (m div n) \<longleftrightarrow>
83e4b6a5e7de streamlined theorems
haftmann
parents: 75880
diff changeset
  1121
    (n = 0 \<longrightarrow> P 0) \<and>
83e4b6a5e7de streamlined theorems
haftmann
parents: 75880
diff changeset
  1122
    (n \<noteq> 0 \<longrightarrow> (\<forall>i j. j < n \<and> m = n * i + j \<longrightarrow> P i))\<close> (is ?div)
83e4b6a5e7de streamlined theorems
haftmann
parents: 75880
diff changeset
  1123
  and split_mod:
83e4b6a5e7de streamlined theorems
haftmann
parents: 75880
diff changeset
  1124
  \<open>Q (m mod n) \<longleftrightarrow>
83e4b6a5e7de streamlined theorems
haftmann
parents: 75880
diff changeset
  1125
    (n = 0 \<longrightarrow> Q m) \<and>
83e4b6a5e7de streamlined theorems
haftmann
parents: 75880
diff changeset
  1126
    (n \<noteq> 0 \<longrightarrow> (\<forall>i j. j < n \<and> m = n * i + j \<longrightarrow> Q j))\<close> (is ?mod)
83e4b6a5e7de streamlined theorems
haftmann
parents: 75880
diff changeset
  1127
  for m n :: nat
83e4b6a5e7de streamlined theorems
haftmann
parents: 75880
diff changeset
  1128
proof -
83e4b6a5e7de streamlined theorems
haftmann
parents: 75880
diff changeset
  1129
  have *: \<open>R (m div n) (m mod n) \<longleftrightarrow>
83e4b6a5e7de streamlined theorems
haftmann
parents: 75880
diff changeset
  1130
    (n = 0 \<longrightarrow> R 0 m) \<and>
83e4b6a5e7de streamlined theorems
haftmann
parents: 75880
diff changeset
  1131
    (n \<noteq> 0 \<longrightarrow> (\<forall>i j. j < n \<and> m = n * i + j \<longrightarrow> R i j))\<close> for R
83e4b6a5e7de streamlined theorems
haftmann
parents: 75880
diff changeset
  1132
    by (cases \<open>n = 0\<close>) auto
83e4b6a5e7de streamlined theorems
haftmann
parents: 75880
diff changeset
  1133
  from * [of \<open>\<lambda>q _. P q\<close>] show ?div .
83e4b6a5e7de streamlined theorems
haftmann
parents: 75880
diff changeset
  1134
  from * [of \<open>\<lambda>_ r. Q r\<close>] show ?mod .
83e4b6a5e7de streamlined theorems
haftmann
parents: 75880
diff changeset
  1135
qed
83e4b6a5e7de streamlined theorems
haftmann
parents: 75880
diff changeset
  1136
83e4b6a5e7de streamlined theorems
haftmann
parents: 75880
diff changeset
  1137
declare split_div [of _ _ \<open>numeral n\<close>, linarith_split] for n
83e4b6a5e7de streamlined theorems
haftmann
parents: 75880
diff changeset
  1138
declare split_mod [of _ _ \<open>numeral n\<close>, linarith_split] for n
83e4b6a5e7de streamlined theorems
haftmann
parents: 75880
diff changeset
  1139
83e4b6a5e7de streamlined theorems
haftmann
parents: 75880
diff changeset
  1140
lemma split_div':
83e4b6a5e7de streamlined theorems
haftmann
parents: 75880
diff changeset
  1141
  "P (m div n) \<longleftrightarrow> n = 0 \<and> P 0 \<or> (\<exists>q. (n * q \<le> m \<and> m < n * Suc q) \<and> P q)"
83e4b6a5e7de streamlined theorems
haftmann
parents: 75880
diff changeset
  1142
proof (cases "n = 0")
83e4b6a5e7de streamlined theorems
haftmann
parents: 75880
diff changeset
  1143
  case True
83e4b6a5e7de streamlined theorems
haftmann
parents: 75880
diff changeset
  1144
  then show ?thesis
83e4b6a5e7de streamlined theorems
haftmann
parents: 75880
diff changeset
  1145
    by simp
83e4b6a5e7de streamlined theorems
haftmann
parents: 75880
diff changeset
  1146
next
83e4b6a5e7de streamlined theorems
haftmann
parents: 75880
diff changeset
  1147
  case False
83e4b6a5e7de streamlined theorems
haftmann
parents: 75880
diff changeset
  1148
  then have "n * q \<le> m \<and> m < n * Suc q \<longleftrightarrow> m div n = q" for q
83e4b6a5e7de streamlined theorems
haftmann
parents: 75880
diff changeset
  1149
    by (auto intro: div_nat_eqI dividend_less_times_div)
83e4b6a5e7de streamlined theorems
haftmann
parents: 75880
diff changeset
  1150
  then show ?thesis
83e4b6a5e7de streamlined theorems
haftmann
parents: 75880
diff changeset
  1151
    by auto
83e4b6a5e7de streamlined theorems
haftmann
parents: 75880
diff changeset
  1152
qed
66808
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1153
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1154
lemma le_div_geq:
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1155
  "m div n = Suc ((m - n) div n)" if "0 < n" and "n \<le> m" for m n :: nat
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1156
proof -
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1157
  from \<open>n \<le> m\<close> obtain q where "m = n + q"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1158
    by (auto simp add: le_iff_add)
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1159
  with \<open>0 < n\<close> show ?thesis
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1160
    by (simp add: div_add_self1)
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1161
qed
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1162
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1163
lemma le_mod_geq:
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1164
  "m mod n = (m - n) mod n" if "n \<le> m" for m n :: nat
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1165
proof -
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1166
  from \<open>n \<le> m\<close> obtain q where "m = n + q"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1167
    by (auto simp add: le_iff_add)
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1168
  then show ?thesis
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1169
    by simp
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1170
qed
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1171
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1172
lemma div_if:
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1173
  "m div n = (if m < n \<or> n = 0 then 0 else Suc ((m - n) div n))"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1174
  by (simp add: le_div_geq)
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1175
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1176
lemma mod_if:
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1177
  "m mod n = (if m < n then m else (m - n) mod n)" for m n :: nat
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1178
  by (simp add: le_mod_geq)
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1179
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1180
lemma div_eq_0_iff:
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1181
  "m div n = 0 \<longleftrightarrow> m < n \<or> n = 0" for m n :: nat
66886
960509bfd47e added lemmas and tuned proofs
haftmann
parents: 66840
diff changeset
  1182
  by (simp add: div_eq_0_iff)
66808
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1183
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1184
lemma div_greater_zero_iff:
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1185
  "m div n > 0 \<longleftrightarrow> n \<le> m \<and> n > 0" for m n :: nat
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1186
  using div_eq_0_iff [of m n] by auto
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1187
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1188
lemma mod_greater_zero_iff_not_dvd:
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1189
  "m mod n > 0 \<longleftrightarrow> \<not> n dvd m" for m n :: nat
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1190
  by (simp add: dvd_eq_mod_eq_0)
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1191
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1192
lemma div_by_Suc_0 [simp]:
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1193
  "m div Suc 0 = m"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1194
  using div_by_1 [of m] by simp
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1195
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1196
lemma mod_by_Suc_0 [simp]:
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1197
  "m mod Suc 0 = 0"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1198
  using mod_by_1 [of m] by simp
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1199
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1200
lemma div2_Suc_Suc [simp]:
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1201
  "Suc (Suc m) div 2 = Suc (m div 2)"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1202
  by (simp add: numeral_2_eq_2 le_div_geq)
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1203
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1204
lemma Suc_n_div_2_gt_zero [simp]:
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1205
  "0 < Suc n div 2" if "n > 0" for n :: nat
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1206
  using that by (cases n) simp_all
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1207
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1208
lemma div_2_gt_zero [simp]:
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1209
  "0 < n div 2" if "Suc 0 < n" for n :: nat
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1210
  using that Suc_n_div_2_gt_zero [of "n - 1"] by simp
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1211
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1212
lemma mod2_Suc_Suc [simp]:
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1213
  "Suc (Suc m) mod 2 = m mod 2"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1214
  by (simp add: numeral_2_eq_2 le_mod_geq)
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1215
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1216
lemma add_self_div_2 [simp]:
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1217
  "(m + m) div 2 = m" for m :: nat
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1218
  by (simp add: mult_2 [symmetric])
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1219
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1220
lemma add_self_mod_2 [simp]:
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1221
  "(m + m) mod 2 = 0" for m :: nat
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1222
  by (simp add: mult_2 [symmetric])
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1223
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1224
lemma mod2_gr_0 [simp]:
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1225
  "0 < m mod 2 \<longleftrightarrow> m mod 2 = 1" for m :: nat
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1226
proof -
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1227
  have "m mod 2 < 2"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1228
    by (rule mod_less_divisor) simp
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1229
  then have "m mod 2 = 0 \<or> m mod 2 = 1"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1230
    by arith
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1231
  then show ?thesis
76224
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  1232
    by auto
66808
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1233
qed
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1234
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1235
lemma mod_Suc_eq [mod_simps]:
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1236
  "Suc (m mod n) mod n = Suc m mod n"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1237
proof -
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1238
  have "(m mod n + 1) mod n = (m + 1) mod n"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1239
    by (simp only: mod_simps)
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1240
  then show ?thesis
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1241
    by simp
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1242
qed
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1243
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1244
lemma mod_Suc_Suc_eq [mod_simps]:
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1245
  "Suc (Suc (m mod n)) mod n = Suc (Suc m) mod n"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1246
proof -
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1247
  have "(m mod n + 2) mod n = (m + 2) mod n"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1248
    by (simp only: mod_simps)
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1249
  then show ?thesis
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1250
    by simp
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1251
qed
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1252
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1253
lemma
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1254
  Suc_mod_mult_self1 [simp]: "Suc (m + k * n) mod n = Suc m mod n"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1255
  and Suc_mod_mult_self2 [simp]: "Suc (m + n * k) mod n = Suc m mod n"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1256
  and Suc_mod_mult_self3 [simp]: "Suc (k * n + m) mod n = Suc m mod n"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1257
  and Suc_mod_mult_self4 [simp]: "Suc (n * k + m) mod n = Suc m mod n"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1258
  by (subst mod_Suc_eq [symmetric], simp add: mod_simps)+
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1259
67083
6b2c0681ef28 new simp rule
haftmann
parents: 67051
diff changeset
  1260
lemma Suc_0_mod_eq [simp]:
6b2c0681ef28 new simp rule
haftmann
parents: 67051
diff changeset
  1261
  "Suc 0 mod n = of_bool (n \<noteq> Suc 0)"
6b2c0681ef28 new simp rule
haftmann
parents: 67051
diff changeset
  1262
  by (cases n) simp_all
6b2c0681ef28 new simp rule
haftmann
parents: 67051
diff changeset
  1263
76053
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
  1264
lemma div_mult2_eq:
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
  1265
    \<open>m div (n * q) = (m div n) div q\<close> (is ?Q)
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
  1266
  and mod_mult2_eq:
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
  1267
    \<open>m mod (n * q) = n * (m div n mod q) + m mod n\<close> (is ?R)
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
  1268
  for m n q :: nat
66808
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1269
proof -
76053
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
  1270
  have \<open>(m div (n * q), m mod (n * q)) = ((m div n) div q, n * (m div n mod q) + m mod n)\<close>
76245
4111c94657b4 slightly less abusive proof pattern
haftmann
parents: 76231
diff changeset
  1271
  proof (induction rule: euclidean_relation_natI)
76053
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
  1272
    case by0
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
  1273
    then show ?case
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
  1274
      by auto
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
  1275
  next
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
  1276
    case divides
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
  1277
    from \<open>n * q dvd m\<close> obtain t where \<open>m = n * q * t\<close> ..
76141
e7497a1de8b9 more concise instance-specific rules on euclidean relation
haftmann
parents: 76053
diff changeset
  1278
    with \<open>n * q > 0\<close> show ?case
76053
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
  1279
      by (simp add: algebra_simps)
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
  1280
  next
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
  1281
    case euclidean_relation
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
  1282
    then have \<open>n > 0\<close> \<open>q > 0\<close>
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
  1283
      by simp_all
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
  1284
    from \<open>n > 0\<close> have \<open>m mod n < n\<close>
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
  1285
      by (rule mod_less_divisor)
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
  1286
    from \<open>q > 0\<close> have \<open>m div n mod q < q\<close>
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
  1287
      by (rule mod_less_divisor)
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
  1288
    then obtain s where \<open>q = Suc (m div n mod q + s)\<close>
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
  1289
      by (blast dest: less_imp_Suc_add)
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
  1290
    moreover have \<open>m mod n + n * (m div n mod q) < n * Suc (m div n mod q + s)\<close>
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
  1291
      using \<open>m mod n < n\<close> by (simp add: add_mult_distrib2)
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
  1292
    ultimately have \<open>m mod n + n * (m div n mod q) < n * q\<close>
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
  1293
      by simp
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
  1294
    then show ?case
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
  1295
      by (simp add: algebra_simps flip: add_mult_distrib2)
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
  1296
  qed
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
  1297
  then show ?Q and ?R
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
  1298
    by simp_all
66808
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1299
qed
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1300
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1301
lemma div_le_mono:
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1302
  "m div k \<le> n div k" if "m \<le> n" for m n k :: nat
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1303
proof -
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1304
  from that obtain q where "n = m + q"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1305
    by (auto simp add: le_iff_add)
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1306
  then show ?thesis
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1307
    by (simp add: div_add1_eq [of m q k])
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1308
qed
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1309
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 68536
diff changeset
  1310
text \<open>Antimonotonicity of \<^const>\<open>divide\<close> in second argument\<close>
66808
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1311
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1312
lemma div_le_mono2:
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1313
  "k div n \<le> k div m" if "0 < m" and "m \<le> n" for m n k :: nat
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1314
using that proof (induct k arbitrary: m rule: less_induct)
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1315
  case (less k)
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1316
  show ?case
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1317
  proof (cases "n \<le> k")
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1318
    case False
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1319
    then show ?thesis
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1320
      by simp
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1321
  next
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1322
    case True
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1323
    have "(k - n) div n \<le> (k - m) div n"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1324
      using less.prems
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1325
      by (blast intro: div_le_mono diff_le_mono2)
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1326
    also have "\<dots> \<le> (k - m) div m"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1327
      using \<open>n \<le> k\<close> less.prems less.hyps [of "k - m" m]
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1328
      by simp
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1329
    finally show ?thesis
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1330
      using \<open>n \<le> k\<close> less.prems
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1331
      by (simp add: le_div_geq)
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1332
  qed
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1333
qed
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1334
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1335
lemma div_le_dividend [simp]:
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1336
  "m div n \<le> m" for m n :: nat
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1337
  using div_le_mono2 [of 1 n m] by (cases "n = 0") simp_all
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1338
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1339
lemma div_less_dividend [simp]:
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1340
  "m div n < m" if "1 < n" and "0 < m" for m n :: nat
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1341
using that proof (induct m rule: less_induct)
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1342
  case (less m)
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1343
  show ?case
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1344
  proof (cases "n < m")
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1345
    case False
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1346
    with less show ?thesis
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1347
      by (cases "n = m") simp_all
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1348
  next
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1349
    case True
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1350
    then show ?thesis
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1351
      using less.hyps [of "m - n"] less.prems
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1352
      by (simp add: le_div_geq)
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1353
  qed
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1354
qed
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1355
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1356
lemma div_eq_dividend_iff:
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1357
  "m div n = m \<longleftrightarrow> n = 1" if "m > 0" for m n :: nat
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1358
proof
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1359
  assume "n = 1"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1360
  then show "m div n = m"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1361
    by simp
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1362
next
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1363
  assume P: "m div n = m"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1364
  show "n = 1"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1365
  proof (rule ccontr)
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1366
    have "n \<noteq> 0"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1367
      by (rule ccontr) (use that P in auto)
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1368
    moreover assume "n \<noteq> 1"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1369
    ultimately have "n > 1"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1370
      by simp
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1371
    with that have "m div n < m"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1372
      by simp
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1373
    with P show False
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1374
      by simp
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1375
  qed
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1376
qed
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1377
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1378
lemma less_mult_imp_div_less:
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1379
  "m div n < i" if "m < i * n" for m n i :: nat
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1380
proof -
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1381
  from that have "i * n > 0"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1382
    by (cases "i * n = 0") simp_all
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1383
  then have "i > 0" and "n > 0"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1384
    by simp_all
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1385
  have "m div n * n \<le> m"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1386
    by simp
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1387
  then have "m div n * n < i * n"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1388
    using that by (rule le_less_trans)
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1389
  with \<open>n > 0\<close> show ?thesis
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1390
    by simp
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1391
qed
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1392
73853
52b829b18066 more lemmas
haftmann
parents: 73555
diff changeset
  1393
lemma div_less_iff_less_mult:
52b829b18066 more lemmas
haftmann
parents: 73555
diff changeset
  1394
  \<open>m div q < n \<longleftrightarrow> m < n * q\<close> (is \<open>?P \<longleftrightarrow> ?Q\<close>)
52b829b18066 more lemmas
haftmann
parents: 73555
diff changeset
  1395
  if \<open>q > 0\<close> for m n q :: nat
52b829b18066 more lemmas
haftmann
parents: 73555
diff changeset
  1396
proof
52b829b18066 more lemmas
haftmann
parents: 73555
diff changeset
  1397
  assume ?Q then show ?P
52b829b18066 more lemmas
haftmann
parents: 73555
diff changeset
  1398
    by (rule less_mult_imp_div_less)
52b829b18066 more lemmas
haftmann
parents: 73555
diff changeset
  1399
next
52b829b18066 more lemmas
haftmann
parents: 73555
diff changeset
  1400
  assume ?P
52b829b18066 more lemmas
haftmann
parents: 73555
diff changeset
  1401
  then obtain h where \<open>n = Suc (m div q + h)\<close>
52b829b18066 more lemmas
haftmann
parents: 73555
diff changeset
  1402
    using less_natE by blast
52b829b18066 more lemmas
haftmann
parents: 73555
diff changeset
  1403
  moreover have \<open>m < m + (Suc h * q - m mod q)\<close>
52b829b18066 more lemmas
haftmann
parents: 73555
diff changeset
  1404
    using that by (simp add: trans_less_add1)
52b829b18066 more lemmas
haftmann
parents: 73555
diff changeset
  1405
  ultimately show ?Q
52b829b18066 more lemmas
haftmann
parents: 73555
diff changeset
  1406
    by (simp add: algebra_simps flip: minus_mod_eq_mult_div)
52b829b18066 more lemmas
haftmann
parents: 73555
diff changeset
  1407
qed
52b829b18066 more lemmas
haftmann
parents: 73555
diff changeset
  1408
52b829b18066 more lemmas
haftmann
parents: 73555
diff changeset
  1409
lemma less_eq_div_iff_mult_less_eq:
52b829b18066 more lemmas
haftmann
parents: 73555
diff changeset
  1410
  \<open>m \<le> n div q \<longleftrightarrow> m * q \<le> n\<close> if \<open>q > 0\<close> for m n q :: nat
52b829b18066 more lemmas
haftmann
parents: 73555
diff changeset
  1411
  using div_less_iff_less_mult [of q n m] that by auto
52b829b18066 more lemmas
haftmann
parents: 73555
diff changeset
  1412
76231
8a48e18f081e reduce prominence of facts
haftmann
parents: 76224
diff changeset
  1413
lemma div_Suc:
76246
c9ea813f92f2 tuned proof
haftmann
parents: 76245
diff changeset
  1414
  \<open>Suc m div n = (if Suc m mod n = 0 then Suc (m div n) else m div n)\<close>
76231
8a48e18f081e reduce prominence of facts
haftmann
parents: 76224
diff changeset
  1415
proof (cases \<open>n = 0 \<or> n = 1\<close>)
8a48e18f081e reduce prominence of facts
haftmann
parents: 76224
diff changeset
  1416
  case True
8a48e18f081e reduce prominence of facts
haftmann
parents: 76224
diff changeset
  1417
  then show ?thesis by auto
8a48e18f081e reduce prominence of facts
haftmann
parents: 76224
diff changeset
  1418
next
8a48e18f081e reduce prominence of facts
haftmann
parents: 76224
diff changeset
  1419
  case False
8a48e18f081e reduce prominence of facts
haftmann
parents: 76224
diff changeset
  1420
  then have \<open>n > 1\<close>
8a48e18f081e reduce prominence of facts
haftmann
parents: 76224
diff changeset
  1421
    by simp
76246
c9ea813f92f2 tuned proof
haftmann
parents: 76245
diff changeset
  1422
  then have \<open>Suc m div n = m div n + Suc (m mod n) div n\<close>
c9ea813f92f2 tuned proof
haftmann
parents: 76245
diff changeset
  1423
    using div_add1_eq [of m 1 n] by simp
c9ea813f92f2 tuned proof
haftmann
parents: 76245
diff changeset
  1424
  also have \<open>Suc (m mod n) div n = of_bool (n dvd Suc m)\<close>
76231
8a48e18f081e reduce prominence of facts
haftmann
parents: 76224
diff changeset
  1425
  proof (cases \<open>n dvd Suc m\<close>)
8a48e18f081e reduce prominence of facts
haftmann
parents: 76224
diff changeset
  1426
    case False
76246
c9ea813f92f2 tuned proof
haftmann
parents: 76245
diff changeset
  1427
    moreover have \<open>Suc (m mod n) \<noteq> n\<close>
76231
8a48e18f081e reduce prominence of facts
haftmann
parents: 76224
diff changeset
  1428
    proof (rule ccontr)
8a48e18f081e reduce prominence of facts
haftmann
parents: 76224
diff changeset
  1429
      assume \<open>\<not> Suc (m mod n) \<noteq> n\<close>
76246
c9ea813f92f2 tuned proof
haftmann
parents: 76245
diff changeset
  1430
      then have \<open>m mod n = n - Suc 0\<close>
76231
8a48e18f081e reduce prominence of facts
haftmann
parents: 76224
diff changeset
  1431
        by simp
8a48e18f081e reduce prominence of facts
haftmann
parents: 76224
diff changeset
  1432
      with \<open>n > 1\<close> have \<open>(m + 1) mod n = 0\<close>
8a48e18f081e reduce prominence of facts
haftmann
parents: 76224
diff changeset
  1433
        by (subst mod_add_left_eq [symmetric]) simp
8a48e18f081e reduce prominence of facts
haftmann
parents: 76224
diff changeset
  1434
      then have \<open>n dvd Suc m\<close>
8a48e18f081e reduce prominence of facts
haftmann
parents: 76224
diff changeset
  1435
        by auto
8a48e18f081e reduce prominence of facts
haftmann
parents: 76224
diff changeset
  1436
      with False show False ..
8a48e18f081e reduce prominence of facts
haftmann
parents: 76224
diff changeset
  1437
    qed
8a48e18f081e reduce prominence of facts
haftmann
parents: 76224
diff changeset
  1438
    moreover have \<open>Suc (m mod n) \<le> n\<close>
8a48e18f081e reduce prominence of facts
haftmann
parents: 76224
diff changeset
  1439
      using \<open>n > 1\<close> by (simp add: Suc_le_eq)
76246
c9ea813f92f2 tuned proof
haftmann
parents: 76245
diff changeset
  1440
    ultimately show ?thesis
c9ea813f92f2 tuned proof
haftmann
parents: 76245
diff changeset
  1441
      by (simp add: div_eq_0_iff)
c9ea813f92f2 tuned proof
haftmann
parents: 76245
diff changeset
  1442
  next
c9ea813f92f2 tuned proof
haftmann
parents: 76245
diff changeset
  1443
    case True
c9ea813f92f2 tuned proof
haftmann
parents: 76245
diff changeset
  1444
    then obtain q where q: \<open>Suc m = n * q\<close> ..
c9ea813f92f2 tuned proof
haftmann
parents: 76245
diff changeset
  1445
    moreover have \<open>q > 0\<close> by (rule ccontr)
c9ea813f92f2 tuned proof
haftmann
parents: 76245
diff changeset
  1446
      (use q in simp)
c9ea813f92f2 tuned proof
haftmann
parents: 76245
diff changeset
  1447
    ultimately have \<open>m mod n = n - Suc 0\<close>
c9ea813f92f2 tuned proof
haftmann
parents: 76245
diff changeset
  1448
      using \<open>n > 1\<close> mult_le_cancel1 [of n \<open>Suc 0\<close> q]
c9ea813f92f2 tuned proof
haftmann
parents: 76245
diff changeset
  1449
      by (auto intro: mod_nat_eqI)
c9ea813f92f2 tuned proof
haftmann
parents: 76245
diff changeset
  1450
    with True \<open>n > 1\<close> show ?thesis
76231
8a48e18f081e reduce prominence of facts
haftmann
parents: 76224
diff changeset
  1451
      by simp
8a48e18f081e reduce prominence of facts
haftmann
parents: 76224
diff changeset
  1452
  qed
76246
c9ea813f92f2 tuned proof
haftmann
parents: 76245
diff changeset
  1453
  finally show ?thesis
c9ea813f92f2 tuned proof
haftmann
parents: 76245
diff changeset
  1454
    by (simp add: mod_greater_zero_iff_not_dvd)
76231
8a48e18f081e reduce prominence of facts
haftmann
parents: 76224
diff changeset
  1455
qed
66808
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1456
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1457
lemma mod_Suc:
76246
c9ea813f92f2 tuned proof
haftmann
parents: 76245
diff changeset
  1458
  \<open>Suc m mod n = (if Suc (m mod n) = n then 0 else Suc (m mod n))\<close>
c9ea813f92f2 tuned proof
haftmann
parents: 76245
diff changeset
  1459
proof (cases \<open>n = 0\<close>)
66808
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1460
  case True
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1461
  then show ?thesis
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1462
    by simp
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1463
next
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1464
  case False
76246
c9ea813f92f2 tuned proof
haftmann
parents: 76245
diff changeset
  1465
  moreover have \<open>Suc m mod n = Suc (m mod n) mod n\<close>
66808
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1466
    by (simp add: mod_simps)
76246
c9ea813f92f2 tuned proof
haftmann
parents: 76245
diff changeset
  1467
  ultimately show ?thesis
c9ea813f92f2 tuned proof
haftmann
parents: 76245
diff changeset
  1468
    by (auto intro!: mod_nat_eqI intro: neq_le_trans simp add: Suc_le_eq)
66808
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1469
qed
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1470
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1471
lemma Suc_times_mod_eq:
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1472
  "Suc (m * n) mod m = 1" if "Suc 0 < m"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1473
  using that by (simp add: mod_Suc)
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1474
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1475
lemma Suc_times_numeral_mod_eq [simp]:
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1476
  "Suc (numeral k * n) mod numeral k = 1" if "numeral k \<noteq> (1::nat)"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1477
  by (rule Suc_times_mod_eq) (use that in simp)
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1478
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1479
lemma Suc_div_le_mono [simp]:
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1480
  "m div n \<le> Suc m div n"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1481
  by (simp add: div_le_mono)
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1482
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1483
text \<open>These lemmas collapse some needless occurrences of Suc:
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1484
  at least three Sucs, since two and fewer are rewritten back to Suc again!
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1485
  We already have some rules to simplify operands smaller than 3.\<close>
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1486
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1487
lemma div_Suc_eq_div_add3 [simp]:
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1488
  "m div Suc (Suc (Suc n)) = m div (3 + n)"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1489
  by (simp add: Suc3_eq_add_3)
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1490
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1491
lemma mod_Suc_eq_mod_add3 [simp]:
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1492
  "m mod Suc (Suc (Suc n)) = m mod (3 + n)"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1493
  by (simp add: Suc3_eq_add_3)
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1494
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1495
lemma Suc_div_eq_add3_div:
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1496
  "Suc (Suc (Suc m)) div n = (3 + m) div n"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1497
  by (simp add: Suc3_eq_add_3)
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1498
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1499
lemma Suc_mod_eq_add3_mod:
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1500
  "Suc (Suc (Suc m)) mod n = (3 + m) mod n"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1501
  by (simp add: Suc3_eq_add_3)
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1502
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1503
lemmas Suc_div_eq_add3_div_numeral [simp] =
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1504
  Suc_div_eq_add3_div [of _ "numeral v"] for v
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1505
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1506
lemmas Suc_mod_eq_add3_mod_numeral [simp] =
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1507
  Suc_mod_eq_add3_mod [of _ "numeral v"] for v
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1508
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1509
lemma (in field_char_0) of_nat_div:
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1510
  "of_nat (m div n) = ((of_nat m - of_nat (m mod n)) / of_nat n)"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1511
proof -
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1512
  have "of_nat (m div n) = ((of_nat (m div n * n + m mod n) - of_nat (m mod n)) / of_nat n :: 'a)"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1513
    unfolding of_nat_add by (cases "n = 0") simp_all
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1514
  then show ?thesis
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1515
    by simp
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1516
qed
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1517
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1518
text \<open>An ``induction'' law for modulus arithmetic.\<close>
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1519
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1520
lemma mod_induct [consumes 3, case_names step]:
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1521
  "P m" if "P n" and "n < p" and "m < p"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1522
    and step: "\<And>n. n < p \<Longrightarrow> P n \<Longrightarrow> P (Suc n mod p)"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1523
using \<open>m < p\<close> proof (induct m)
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1524
  case 0
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1525
  show ?case
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1526
  proof (rule ccontr)
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1527
    assume "\<not> P 0"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1528
    from \<open>n < p\<close> have "0 < p"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1529
      by simp
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1530
    from \<open>n < p\<close> obtain m where "0 < m" and "p = n + m"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1531
      by (blast dest: less_imp_add_positive)
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1532
    with \<open>P n\<close> have "P (p - m)"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1533
      by simp
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1534
    moreover have "\<not> P (p - m)"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1535
    using \<open>0 < m\<close> proof (induct m)
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1536
      case 0
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1537
      then show ?case
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1538
        by simp
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1539
    next
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1540
      case (Suc m)
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1541
      show ?case
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1542
      proof
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1543
        assume P: "P (p - Suc m)"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1544
        with \<open>\<not> P 0\<close> have "Suc m < p"
76224
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  1545
          by (auto intro: ccontr)
66808
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1546
        then have "Suc (p - Suc m) = p - m"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1547
          by arith
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1548
        moreover from \<open>0 < p\<close> have "p - Suc m < p"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1549
          by arith
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1550
        with P step have "P ((Suc (p - Suc m)) mod p)"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1551
          by blast
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1552
        ultimately show False
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1553
          using \<open>\<not> P 0\<close> Suc.hyps by (cases "m = 0") simp_all
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1554
      qed
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1555
    qed
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1556
    ultimately show False
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1557
      by blast
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1558
  qed
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1559
next
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1560
  case (Suc m)
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1561
  then have "m < p" and mod: "Suc m mod p = Suc m"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1562
    by simp_all
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1563
  from \<open>m < p\<close> have "P m"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1564
    by (rule Suc.hyps)
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1565
  with \<open>m < p\<close> have "P (Suc m mod p)"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1566
    by (rule step)
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1567
  with mod show ?case
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1568
    by simp
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1569
qed
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1570
73555
92783562ab78 collected combinatorial material
haftmann
parents: 73535
diff changeset
  1571
lemma funpow_mod_eq: \<^marker>\<open>contributor \<open>Lars Noschinski\<close>\<close>
92783562ab78 collected combinatorial material
haftmann
parents: 73535
diff changeset
  1572
  \<open>(f ^^ (m mod n)) x = (f ^^ m) x\<close> if \<open>(f ^^ n) x = x\<close>
92783562ab78 collected combinatorial material
haftmann
parents: 73535
diff changeset
  1573
proof -
92783562ab78 collected combinatorial material
haftmann
parents: 73535
diff changeset
  1574
  have \<open>(f ^^ m) x = (f ^^ (m mod n + m div n * n)) x\<close>
92783562ab78 collected combinatorial material
haftmann
parents: 73535
diff changeset
  1575
    by simp
92783562ab78 collected combinatorial material
haftmann
parents: 73535
diff changeset
  1576
  also have \<open>\<dots> = (f ^^ (m mod n)) (((f ^^ n) ^^ (m div n)) x)\<close>
92783562ab78 collected combinatorial material
haftmann
parents: 73535
diff changeset
  1577
    by (simp only: funpow_add funpow_mult ac_simps) simp
92783562ab78 collected combinatorial material
haftmann
parents: 73535
diff changeset
  1578
  also have \<open>((f ^^ n) ^^ q) x = x\<close> for q
92783562ab78 collected combinatorial material
haftmann
parents: 73535
diff changeset
  1579
    by (induction q) (use \<open>(f ^^ n) x = x\<close> in simp_all)
92783562ab78 collected combinatorial material
haftmann
parents: 73535
diff changeset
  1580
  finally show ?thesis
92783562ab78 collected combinatorial material
haftmann
parents: 73535
diff changeset
  1581
    by simp
92783562ab78 collected combinatorial material
haftmann
parents: 73535
diff changeset
  1582
qed
92783562ab78 collected combinatorial material
haftmann
parents: 73535
diff changeset
  1583
76224
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  1584
lemma mod_eq_dvd_iff_nat:
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  1585
  \<open>m mod q = n mod q \<longleftrightarrow> q dvd m - n\<close> (is \<open>?P \<longleftrightarrow> ?Q\<close>)
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  1586
    if \<open>m \<ge> n\<close> for m n q :: nat
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  1587
proof
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  1588
  assume ?Q
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  1589
  then obtain s where \<open>m - n = q * s\<close> ..
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  1590
  with that have \<open>m = q * s + n\<close>
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  1591
    by simp
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  1592
  then show ?P
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  1593
    by simp
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  1594
next
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  1595
  assume ?P
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  1596
  have \<open>m - n = m div q * q + m mod q - (n div q * q + n mod q)\<close>
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  1597
    by simp
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  1598
  also have \<open>\<dots> = q * (m div q - n div q)\<close>
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  1599
    by (simp only: algebra_simps \<open>?P\<close>)
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  1600
  finally show ?Q ..
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  1601
qed
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  1602
76231
8a48e18f081e reduce prominence of facts
haftmann
parents: 76224
diff changeset
  1603
lemma mod_eq_iff_dvd_symdiff_nat:
8a48e18f081e reduce prominence of facts
haftmann
parents: 76224
diff changeset
  1604
  \<open>m mod q = n mod q \<longleftrightarrow> q dvd nat \<bar>int m - int n\<bar>\<close>
8a48e18f081e reduce prominence of facts
haftmann
parents: 76224
diff changeset
  1605
  by (auto simp add: abs_if mod_eq_dvd_iff_nat nat_diff_distrib dest: sym intro: sym)
8a48e18f081e reduce prominence of facts
haftmann
parents: 76224
diff changeset
  1606
76224
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  1607
lemma mod_eq_nat1E:
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  1608
  fixes m n q :: nat
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  1609
  assumes "m mod q = n mod q" and "m \<ge> n"
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  1610
  obtains s where "m = n + q * s"
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  1611
proof -
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  1612
  from assms have "q dvd m - n"
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  1613
    by (simp add: mod_eq_dvd_iff_nat)
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  1614
  then obtain s where "m - n = q * s" ..
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  1615
  with \<open>m \<ge> n\<close> have "m = n + q * s"
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  1616
    by simp
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  1617
  with that show thesis .
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  1618
qed
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  1619
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  1620
lemma mod_eq_nat2E:
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  1621
  fixes m n q :: nat
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  1622
  assumes "m mod q = n mod q" and "n \<ge> m"
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  1623
  obtains s where "n = m + q * s"
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  1624
  using assms mod_eq_nat1E [of n q m] by (auto simp add: ac_simps)
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  1625
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  1626
lemma nat_mod_eq_iff:
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  1627
  "(x::nat) mod n = y mod n \<longleftrightarrow> (\<exists>q1 q2. x + n * q1 = y + n * q2)"  (is "?lhs = ?rhs")
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  1628
proof
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  1629
  assume H: "x mod n = y mod n"
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  1630
  { assume xy: "x \<le> y"
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  1631
    from H have th: "y mod n = x mod n" by simp
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  1632
    from mod_eq_nat1E [OF th xy] obtain q where "y = x + n * q" .
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  1633
    then have "x + n * q = y + n * 0"
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  1634
      by simp
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  1635
    then have "\<exists>q1 q2. x + n * q1 = y + n * q2"
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  1636
      by blast
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  1637
  }
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  1638
  moreover
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  1639
  { assume xy: "y \<le> x"
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  1640
    from mod_eq_nat1E [OF H xy] obtain q where "x = y + n * q" .
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  1641
    then have "x + n * 0 = y + n * q"
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  1642
      by simp
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  1643
    then have "\<exists>q1 q2. x + n * q1 = y + n * q2"
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  1644
      by blast
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  1645
  }
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  1646
  ultimately show ?rhs using linear[of x y] by blast
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  1647
next
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  1648
  assume ?rhs then obtain q1 q2 where q12: "x + n * q1 = y + n * q2" by blast
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  1649
  hence "(x + n * q1) mod n = (y + n * q2) mod n" by simp
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  1650
  thus  ?lhs by simp
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  1651
qed
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  1652
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  1653
66808
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1654
75876
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  1655
subsection \<open>Elementary euclidean division on \<^typ>\<open>int\<close>\<close>
66816
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1656
75876
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  1657
subsubsection \<open>Basic instantiation\<close>
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  1658
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  1659
instantiation int :: "{normalization_semidom, idom_modulo}"
66816
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1660
begin
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1661
75875
48d032035744 streamlined primitive definitions for integer division
haftmann
parents: 75669
diff changeset
  1662
definition normalize_int :: \<open>int \<Rightarrow> int\<close>
48d032035744 streamlined primitive definitions for integer division
haftmann
parents: 75669
diff changeset
  1663
  where [simp]: \<open>normalize = (abs :: int \<Rightarrow> int)\<close>
66816
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1664
75875
48d032035744 streamlined primitive definitions for integer division
haftmann
parents: 75669
diff changeset
  1665
definition unit_factor_int :: \<open>int \<Rightarrow> int\<close>
48d032035744 streamlined primitive definitions for integer division
haftmann
parents: 75669
diff changeset
  1666
  where [simp]: \<open>unit_factor = (sgn :: int \<Rightarrow> int)\<close>
48d032035744 streamlined primitive definitions for integer division
haftmann
parents: 75669
diff changeset
  1667
48d032035744 streamlined primitive definitions for integer division
haftmann
parents: 75669
diff changeset
  1668
definition divide_int :: \<open>int \<Rightarrow> int \<Rightarrow> int\<close>
48d032035744 streamlined primitive definitions for integer division
haftmann
parents: 75669
diff changeset
  1669
  where \<open>k div l = (sgn k * sgn l * int (nat \<bar>k\<bar> div nat \<bar>l\<bar>)
48d032035744 streamlined primitive definitions for integer division
haftmann
parents: 75669
diff changeset
  1670
    - of_bool (l \<noteq> 0 \<and> sgn k \<noteq> sgn l \<and> \<not> l dvd k))\<close>
66816
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1671
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1672
lemma divide_int_unfold:
75875
48d032035744 streamlined primitive definitions for integer division
haftmann
parents: 75669
diff changeset
  1673
  \<open>(sgn k * int m) div (sgn l * int n) = (sgn k * sgn l * int (m div n)
48d032035744 streamlined primitive definitions for integer division
haftmann
parents: 75669
diff changeset
  1674
    - of_bool ((k = 0 \<longleftrightarrow> m = 0) \<and> l \<noteq> 0 \<and> n \<noteq> 0 \<and> sgn k \<noteq> sgn l \<and> \<not> n dvd m))\<close>
48d032035744 streamlined primitive definitions for integer division
haftmann
parents: 75669
diff changeset
  1675
  by (simp add: divide_int_def sgn_mult nat_mult_distrib abs_mult sgn_eq_0_iff ac_simps)
66816
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1676
75876
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  1677
definition modulo_int :: \<open>int \<Rightarrow> int \<Rightarrow> int\<close>
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  1678
  where \<open>k mod l = sgn k * int (nat \<bar>k\<bar> mod nat \<bar>l\<bar>) + l * of_bool (sgn k \<noteq> sgn l \<and> \<not> l dvd k)\<close>
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  1679
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  1680
lemma modulo_int_unfold:
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  1681
  \<open>(sgn k * int m) mod (sgn l * int n) =
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  1682
    sgn k * int (m mod (of_bool (l \<noteq> 0) * n)) + (sgn l * int n) * of_bool ((k = 0 \<longleftrightarrow> m = 0) \<and> sgn k \<noteq> sgn l \<and> \<not> n dvd m)\<close>
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  1683
  by (auto simp add: modulo_int_def sgn_mult abs_mult)
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  1684
66816
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1685
instance proof
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1686
  fix k :: int show "k div 0 = 0"
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1687
  by (simp add: divide_int_def)
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1688
next
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1689
  fix k l :: int
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1690
  assume "l \<noteq> 0"
76224
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  1691
  obtain n m and s t where k: "k = sgn s * int n" and l: "l = sgn t * int m"
66816
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1692
    by (blast intro: int_sgnE elim: that)
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1693
  then have "k * l = sgn (s * t) * int (n * m)"
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1694
    by (simp add: ac_simps sgn_mult)
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1695
  with k l \<open>l \<noteq> 0\<close> show "k * l div l = k"
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1696
    by (simp only: divide_int_unfold)
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1697
      (auto simp add: algebra_simps sgn_mult sgn_1_pos sgn_0_0)
75876
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  1698
next
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  1699
  fix k l :: int
76224
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  1700
  obtain n m and s t where "k = sgn s * int n" and "l = sgn t * int m"
75876
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  1701
    by (blast intro: int_sgnE elim: that)
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  1702
  then show "k div l * l + k mod l = k"
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  1703
    by (simp add: divide_int_unfold modulo_int_unfold algebra_simps modulo_nat_def of_nat_diff)
66816
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1704
qed (auto simp add: sgn_mult mult_sgn_abs abs_eq_iff')
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1705
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1706
end
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1707
75875
48d032035744 streamlined primitive definitions for integer division
haftmann
parents: 75669
diff changeset
  1708
75876
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  1709
subsubsection \<open>Algebraic foundations\<close>
75875
48d032035744 streamlined primitive definitions for integer division
haftmann
parents: 75669
diff changeset
  1710
67051
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
  1711
lemma coprime_int_iff [simp]:
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
  1712
  "coprime (int m) (int n) \<longleftrightarrow> coprime m n" (is "?P \<longleftrightarrow> ?Q")
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
  1713
proof
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
  1714
  assume ?P
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
  1715
  show ?Q
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
  1716
  proof (rule coprimeI)
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
  1717
    fix q
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
  1718
    assume "q dvd m" "q dvd n"
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
  1719
    then have "int q dvd int m" "int q dvd int n"
67118
ccab07d1196c more simplification rules
haftmann
parents: 67087
diff changeset
  1720
      by simp_all
67051
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
  1721
    with \<open>?P\<close> have "is_unit (int q)"
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
  1722
      by (rule coprime_common_divisor)
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
  1723
    then show "is_unit q"
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
  1724
      by simp
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
  1725
  qed
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
  1726
next
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
  1727
  assume ?Q
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
  1728
  show ?P
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
  1729
  proof (rule coprimeI)
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
  1730
    fix k
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
  1731
    assume "k dvd int m" "k dvd int n"
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
  1732
    then have "nat \<bar>k\<bar> dvd m" "nat \<bar>k\<bar> dvd n"
67118
ccab07d1196c more simplification rules
haftmann
parents: 67087
diff changeset
  1733
      by simp_all
67051
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
  1734
    with \<open>?Q\<close> have "is_unit (nat \<bar>k\<bar>)"
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
  1735
      by (rule coprime_common_divisor)
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
  1736
    then show "is_unit k"
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
  1737
      by simp
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
  1738
  qed
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
  1739
qed
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
  1740
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
  1741
lemma coprime_abs_left_iff [simp]:
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
  1742
  "coprime \<bar>k\<bar> l \<longleftrightarrow> coprime k l" for k l :: int
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
  1743
  using coprime_normalize_left_iff [of k l] by simp
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
  1744
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
  1745
lemma coprime_abs_right_iff [simp]:
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
  1746
  "coprime k \<bar>l\<bar> \<longleftrightarrow> coprime k l" for k l :: int
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
  1747
  using coprime_abs_left_iff [of l k] by (simp add: ac_simps)
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
  1748
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
  1749
lemma coprime_nat_abs_left_iff [simp]:
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
  1750
  "coprime (nat \<bar>k\<bar>) n \<longleftrightarrow> coprime k (int n)"
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
  1751
proof -
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
  1752
  define m where "m = nat \<bar>k\<bar>"
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
  1753
  then have "\<bar>k\<bar> = int m"
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
  1754
    by simp
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
  1755
  moreover have "coprime k (int n) \<longleftrightarrow> coprime \<bar>k\<bar> (int n)"
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
  1756
    by simp
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
  1757
  ultimately show ?thesis
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
  1758
    by simp
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
  1759
qed
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
  1760
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
  1761
lemma coprime_nat_abs_right_iff [simp]:
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
  1762
  "coprime n (nat \<bar>k\<bar>) \<longleftrightarrow> coprime (int n) k"
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
  1763
  using coprime_nat_abs_left_iff [of k n] by (simp add: ac_simps)
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
  1764
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
  1765
lemma coprime_common_divisor_int: "coprime a b \<Longrightarrow> x dvd a \<Longrightarrow> x dvd b \<Longrightarrow> \<bar>x\<bar> = 1"
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
  1766
  for a b :: int
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
  1767
  by (drule coprime_common_divisor [of _ _ x]) simp_all
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
  1768
75876
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  1769
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  1770
subsubsection \<open>Basic conversions\<close>
66816
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1771
75876
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  1772
lemma div_abs_eq_div_nat:
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  1773
  "\<bar>k\<bar> div \<bar>l\<bar> = int (nat \<bar>k\<bar> div nat \<bar>l\<bar>)"
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  1774
  by (auto simp add: divide_int_def)
66816
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1775
75876
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  1776
lemma div_eq_div_abs:
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  1777
  \<open>k div l = sgn k * sgn l * (\<bar>k\<bar> div \<bar>l\<bar>)
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  1778
    - of_bool (l \<noteq> 0 \<and> sgn k \<noteq> sgn l \<and> \<not> l dvd k)\<close>
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  1779
  for k l :: int
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  1780
  by (simp add: divide_int_def [of k l] div_abs_eq_div_nat)
66816
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1781
75876
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  1782
lemma div_abs_eq:
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  1783
  \<open>\<bar>k\<bar> div \<bar>l\<bar> = sgn k * sgn l * (k div l + of_bool (sgn k \<noteq> sgn l \<and> \<not> l dvd k))\<close>
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  1784
  for k l :: int
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  1785
  by (simp add: div_eq_div_abs [of k l] ac_simps)
66838
17989f6bc7b2 clarified uniqueness criterion for euclidean rings
haftmann
parents: 66837
diff changeset
  1786
75875
48d032035744 streamlined primitive definitions for integer division
haftmann
parents: 75669
diff changeset
  1787
lemma mod_abs_eq_div_nat:
48d032035744 streamlined primitive definitions for integer division
haftmann
parents: 75669
diff changeset
  1788
  "\<bar>k\<bar> mod \<bar>l\<bar> = int (nat \<bar>k\<bar> mod nat \<bar>l\<bar>)"
48d032035744 streamlined primitive definitions for integer division
haftmann
parents: 75669
diff changeset
  1789
  by (simp add: modulo_int_def)
48d032035744 streamlined primitive definitions for integer division
haftmann
parents: 75669
diff changeset
  1790
48d032035744 streamlined primitive definitions for integer division
haftmann
parents: 75669
diff changeset
  1791
lemma mod_eq_mod_abs:
48d032035744 streamlined primitive definitions for integer division
haftmann
parents: 75669
diff changeset
  1792
  \<open>k mod l = sgn k * (\<bar>k\<bar> mod \<bar>l\<bar>) + l * of_bool (sgn k \<noteq> sgn l \<and> \<not> l dvd k)\<close>
48d032035744 streamlined primitive definitions for integer division
haftmann
parents: 75669
diff changeset
  1793
  for k l :: int
48d032035744 streamlined primitive definitions for integer division
haftmann
parents: 75669
diff changeset
  1794
  by (simp add: modulo_int_def [of k l] mod_abs_eq_div_nat)
48d032035744 streamlined primitive definitions for integer division
haftmann
parents: 75669
diff changeset
  1795
48d032035744 streamlined primitive definitions for integer division
haftmann
parents: 75669
diff changeset
  1796
lemma mod_abs_eq:
48d032035744 streamlined primitive definitions for integer division
haftmann
parents: 75669
diff changeset
  1797
  \<open>\<bar>k\<bar> mod \<bar>l\<bar> = sgn k * (k mod l - l * of_bool (sgn k \<noteq> sgn l \<and> \<not> l dvd k))\<close>
48d032035744 streamlined primitive definitions for integer division
haftmann
parents: 75669
diff changeset
  1798
  for k l :: int
48d032035744 streamlined primitive definitions for integer division
haftmann
parents: 75669
diff changeset
  1799
  by (auto simp: mod_eq_mod_abs [of k l])
48d032035744 streamlined primitive definitions for integer division
haftmann
parents: 75669
diff changeset
  1800
75876
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  1801
lemma div_sgn_abs_cancel:
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  1802
  fixes k l v :: int
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  1803
  assumes "v \<noteq> 0"
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  1804
  shows "(sgn v * \<bar>k\<bar>) div (sgn v * \<bar>l\<bar>) = \<bar>k\<bar> div \<bar>l\<bar>"
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  1805
  using assms by (simp add: sgn_mult abs_mult sgn_0_0
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  1806
    divide_int_def [of "sgn v * \<bar>k\<bar>" "sgn v * \<bar>l\<bar>"] flip: div_abs_eq_div_nat)
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  1807
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  1808
lemma div_eq_sgn_abs:
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  1809
  fixes k l v :: int
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  1810
  assumes "sgn k = sgn l"
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  1811
  shows "k div l = \<bar>k\<bar> div \<bar>l\<bar>"
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  1812
  using assms by (auto simp add: div_abs_eq)
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  1813
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  1814
lemma div_dvd_sgn_abs:
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  1815
  fixes k l :: int
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  1816
  assumes "l dvd k"
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  1817
  shows "k div l = (sgn k * sgn l) * (\<bar>k\<bar> div \<bar>l\<bar>)"
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  1818
  using assms by (auto simp add: div_abs_eq ac_simps)
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  1819
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  1820
lemma div_noneq_sgn_abs:
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  1821
  fixes k l :: int
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  1822
  assumes "l \<noteq> 0"
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  1823
  assumes "sgn k \<noteq> sgn l"
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  1824
  shows "k div l = - (\<bar>k\<bar> div \<bar>l\<bar>) - of_bool (\<not> l dvd k)"
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  1825
  using assms by (auto simp add: div_abs_eq ac_simps sgn_0_0 dest!: sgn_not_eq_imp)
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  1826
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  1827
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  1828
subsubsection \<open>Euclidean division\<close>
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  1829
66838
17989f6bc7b2 clarified uniqueness criterion for euclidean rings
haftmann
parents: 66837
diff changeset
  1830
instantiation int :: unique_euclidean_ring
17989f6bc7b2 clarified uniqueness criterion for euclidean rings
haftmann
parents: 66837
diff changeset
  1831
begin
17989f6bc7b2 clarified uniqueness criterion for euclidean rings
haftmann
parents: 66837
diff changeset
  1832
17989f6bc7b2 clarified uniqueness criterion for euclidean rings
haftmann
parents: 66837
diff changeset
  1833
definition euclidean_size_int :: "int \<Rightarrow> nat"
17989f6bc7b2 clarified uniqueness criterion for euclidean rings
haftmann
parents: 66837
diff changeset
  1834
  where [simp]: "euclidean_size_int = (nat \<circ> abs :: int \<Rightarrow> nat)"
17989f6bc7b2 clarified uniqueness criterion for euclidean rings
haftmann
parents: 66837
diff changeset
  1835
17989f6bc7b2 clarified uniqueness criterion for euclidean rings
haftmann
parents: 66837
diff changeset
  1836
definition division_segment_int :: "int \<Rightarrow> int"
17989f6bc7b2 clarified uniqueness criterion for euclidean rings
haftmann
parents: 66837
diff changeset
  1837
  where "division_segment_int k = (if k \<ge> 0 then 1 else - 1)"
17989f6bc7b2 clarified uniqueness criterion for euclidean rings
haftmann
parents: 66837
diff changeset
  1838
17989f6bc7b2 clarified uniqueness criterion for euclidean rings
haftmann
parents: 66837
diff changeset
  1839
lemma division_segment_eq_sgn:
17989f6bc7b2 clarified uniqueness criterion for euclidean rings
haftmann
parents: 66837
diff changeset
  1840
  "division_segment k = sgn k" if "k \<noteq> 0" for k :: int
17989f6bc7b2 clarified uniqueness criterion for euclidean rings
haftmann
parents: 66837
diff changeset
  1841
  using that by (simp add: division_segment_int_def)
17989f6bc7b2 clarified uniqueness criterion for euclidean rings
haftmann
parents: 66837
diff changeset
  1842
17989f6bc7b2 clarified uniqueness criterion for euclidean rings
haftmann
parents: 66837
diff changeset
  1843
lemma abs_division_segment [simp]:
17989f6bc7b2 clarified uniqueness criterion for euclidean rings
haftmann
parents: 66837
diff changeset
  1844
  "\<bar>division_segment k\<bar> = 1" for k :: int
17989f6bc7b2 clarified uniqueness criterion for euclidean rings
haftmann
parents: 66837
diff changeset
  1845
  by (simp add: division_segment_int_def)
17989f6bc7b2 clarified uniqueness criterion for euclidean rings
haftmann
parents: 66837
diff changeset
  1846
66816
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1847
lemma abs_mod_less:
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1848
  "\<bar>k mod l\<bar> < \<bar>l\<bar>" if "l \<noteq> 0" for k l :: int
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1849
proof -
76224
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  1850
  obtain n m and s t where "k = sgn s * int n" and "l = sgn t * int m"
66816
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1851
    by (blast intro: int_sgnE elim: that)
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1852
  with that show ?thesis
75875
48d032035744 streamlined primitive definitions for integer division
haftmann
parents: 75669
diff changeset
  1853
    by (auto simp add: modulo_int_unfold abs_mult mod_greater_zero_iff_not_dvd
48d032035744 streamlined primitive definitions for integer division
haftmann
parents: 75669
diff changeset
  1854
        simp flip: right_diff_distrib dest!: sgn_not_eq_imp)
48d032035744 streamlined primitive definitions for integer division
haftmann
parents: 75669
diff changeset
  1855
      (simp add: sgn_0_0)
66816
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1856
qed
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1857
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1858
lemma sgn_mod:
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1859
  "sgn (k mod l) = sgn l" if "l \<noteq> 0" "\<not> l dvd k" for k l :: int
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1860
proof -
76224
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  1861
  obtain n m and s t where "k = sgn s * int n" and "l = sgn t * int m"
66816
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1862
    by (blast intro: int_sgnE elim: that)
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1863
  with that show ?thesis
75875
48d032035744 streamlined primitive definitions for integer division
haftmann
parents: 75669
diff changeset
  1864
    by (auto simp add: modulo_int_unfold sgn_mult mod_greater_zero_iff_not_dvd
48d032035744 streamlined primitive definitions for integer division
haftmann
parents: 75669
diff changeset
  1865
      simp flip: right_diff_distrib dest!: sgn_not_eq_imp)
66816
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1866
qed
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1867
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1868
instance proof
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1869
  fix k l :: int
66838
17989f6bc7b2 clarified uniqueness criterion for euclidean rings
haftmann
parents: 66837
diff changeset
  1870
  show "division_segment (k mod l) = division_segment l" if
17989f6bc7b2 clarified uniqueness criterion for euclidean rings
haftmann
parents: 66837
diff changeset
  1871
    "l \<noteq> 0" and "\<not> l dvd k"
17989f6bc7b2 clarified uniqueness criterion for euclidean rings
haftmann
parents: 66837
diff changeset
  1872
    using that by (simp add: division_segment_eq_sgn dvd_eq_mod_eq_0 sgn_mod)
66816
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1873
next
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1874
  fix l q r :: int
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1875
  obtain n m and s t
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1876
     where l: "l = sgn s * int n" and q: "q = sgn t * int m"
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1877
    by (blast intro: int_sgnE elim: that)
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1878
  assume \<open>l \<noteq> 0\<close>
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1879
  with l have "s \<noteq> 0" and "n > 0"
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1880
    by (simp_all add: sgn_0_0)
66838
17989f6bc7b2 clarified uniqueness criterion for euclidean rings
haftmann
parents: 66837
diff changeset
  1881
  assume "division_segment r = division_segment l"
66816
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1882
  moreover have "r = sgn r * \<bar>r\<bar>"
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1883
    by (simp add: sgn_mult_abs)
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1884
  moreover define u where "u = nat \<bar>r\<bar>"
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1885
  ultimately have "r = sgn l * int u"
66838
17989f6bc7b2 clarified uniqueness criterion for euclidean rings
haftmann
parents: 66837
diff changeset
  1886
    using division_segment_eq_sgn \<open>l \<noteq> 0\<close> by (cases "r = 0") simp_all
66816
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1887
  with l \<open>n > 0\<close> have r: "r = sgn s * int u"
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1888
    by (simp add: sgn_mult)
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1889
  assume "euclidean_size r < euclidean_size l"
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1890
  with l r \<open>s \<noteq> 0\<close> have "u < n"
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1891
    by (simp add: abs_mult)
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1892
  show "(q * l + r) div l = q"
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1893
  proof (cases "q = 0 \<or> r = 0")
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1894
    case True
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1895
    then show ?thesis
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1896
    proof
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1897
      assume "q = 0"
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1898
      then show ?thesis
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1899
        using l r \<open>u < n\<close> by (simp add: divide_int_unfold)
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1900
    next
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1901
      assume "r = 0"
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1902
      from \<open>r = 0\<close> have *: "q * l + r = sgn (t * s) * int (n * m)"
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1903
        using q l by (simp add: ac_simps sgn_mult)
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1904
      from \<open>s \<noteq> 0\<close> \<open>n > 0\<close> show ?thesis
75875
48d032035744 streamlined primitive definitions for integer division
haftmann
parents: 75669
diff changeset
  1905
        by (simp only: *, simp only: * q l divide_int_unfold)
48d032035744 streamlined primitive definitions for integer division
haftmann
parents: 75669
diff changeset
  1906
          (auto simp add: sgn_mult ac_simps)
66816
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1907
    qed
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1908
  next
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1909
    case False
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1910
    with q r have "t \<noteq> 0" and "m > 0" and "s \<noteq> 0" and "u > 0"
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1911
      by (simp_all add: sgn_0_0)
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1912
    moreover from \<open>0 < m\<close> \<open>u < n\<close> have "u \<le> m * n"
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1913
      using mult_le_less_imp_less [of 1 m u n] by simp
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1914
    ultimately have *: "q * l + r = sgn (s * t)
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1915
      * int (if t < 0 then m * n - u else m * n + u)"
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1916
      using l q r
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1917
      by (simp add: sgn_mult algebra_simps of_nat_diff)
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1918
    have "(m * n - u) div n = m - 1" if "u > 0"
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1919
      using \<open>0 < m\<close> \<open>u < n\<close> that
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1920
      by (auto intro: div_nat_eqI simp add: algebra_simps)
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1921
    moreover have "n dvd m * n - u \<longleftrightarrow> n dvd u"
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1922
      using \<open>u \<le> m * n\<close> dvd_diffD1 [of n "m * n" u]
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1923
      by auto
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1924
    ultimately show ?thesis
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1925
      using \<open>s \<noteq> 0\<close> \<open>m > 0\<close> \<open>u > 0\<close> \<open>u < n\<close> \<open>u \<le> m * n\<close>
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1926
      by (simp only: *, simp only: l q divide_int_unfold)
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1927
        (auto simp add: sgn_mult sgn_0_0 sgn_1_pos algebra_simps dest: dvd_imp_le)
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1928
  qed
68536
e14848001c4c avoid pending shyps in global theory facts;
wenzelm
parents: 67118
diff changeset
  1929
qed (use mult_le_mono2 [of 1] in \<open>auto simp add: division_segment_int_def not_le zero_less_mult_iff mult_less_0_iff abs_mult sgn_mult abs_mod_less sgn_mod nat_mult_distrib\<close>)
66816
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1930
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1931
end
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1932
76141
e7497a1de8b9 more concise instance-specific rules on euclidean relation
haftmann
parents: 76053
diff changeset
  1933
lemma euclidean_relation_intI [case_names by0 divides euclidean_relation]:
e7497a1de8b9 more concise instance-specific rules on euclidean relation
haftmann
parents: 76053
diff changeset
  1934
  \<open>(k div l, k mod l) = (q, r)\<close>
e7497a1de8b9 more concise instance-specific rules on euclidean relation
haftmann
parents: 76053
diff changeset
  1935
    if by0': \<open>l = 0 \<Longrightarrow> q = 0 \<and> r = k\<close>
e7497a1de8b9 more concise instance-specific rules on euclidean relation
haftmann
parents: 76053
diff changeset
  1936
    and divides': \<open>l \<noteq> 0 \<Longrightarrow> l dvd k \<Longrightarrow> r = 0 \<and> k = q * l\<close>
e7497a1de8b9 more concise instance-specific rules on euclidean relation
haftmann
parents: 76053
diff changeset
  1937
    and euclidean_relation': \<open>l \<noteq> 0 \<Longrightarrow> \<not> l dvd k \<Longrightarrow> sgn r = sgn l
e7497a1de8b9 more concise instance-specific rules on euclidean relation
haftmann
parents: 76053
diff changeset
  1938
      \<and> \<bar>r\<bar> < \<bar>l\<bar> \<and> k = q * l + r\<close> for k l :: int
76245
4111c94657b4 slightly less abusive proof pattern
haftmann
parents: 76231
diff changeset
  1939
proof (induction rule: euclidean_relationI)
76141
e7497a1de8b9 more concise instance-specific rules on euclidean relation
haftmann
parents: 76053
diff changeset
  1940
  case by0
e7497a1de8b9 more concise instance-specific rules on euclidean relation
haftmann
parents: 76053
diff changeset
  1941
  then show ?case
e7497a1de8b9 more concise instance-specific rules on euclidean relation
haftmann
parents: 76053
diff changeset
  1942
    by (rule by0')
e7497a1de8b9 more concise instance-specific rules on euclidean relation
haftmann
parents: 76053
diff changeset
  1943
next
e7497a1de8b9 more concise instance-specific rules on euclidean relation
haftmann
parents: 76053
diff changeset
  1944
  case divides
e7497a1de8b9 more concise instance-specific rules on euclidean relation
haftmann
parents: 76053
diff changeset
  1945
  then show ?case
e7497a1de8b9 more concise instance-specific rules on euclidean relation
haftmann
parents: 76053
diff changeset
  1946
    by (rule divides')
e7497a1de8b9 more concise instance-specific rules on euclidean relation
haftmann
parents: 76053
diff changeset
  1947
next
e7497a1de8b9 more concise instance-specific rules on euclidean relation
haftmann
parents: 76053
diff changeset
  1948
  case euclidean_relation
e7497a1de8b9 more concise instance-specific rules on euclidean relation
haftmann
parents: 76053
diff changeset
  1949
  with euclidean_relation' have \<open>sgn r = sgn l\<close> \<open>\<bar>r\<bar> < \<bar>l\<bar>\<close> \<open>k = q * l + r\<close>
e7497a1de8b9 more concise instance-specific rules on euclidean relation
haftmann
parents: 76053
diff changeset
  1950
    by simp_all
e7497a1de8b9 more concise instance-specific rules on euclidean relation
haftmann
parents: 76053
diff changeset
  1951
  from \<open>sgn r = sgn l\<close> \<open>l \<noteq> 0\<close> have \<open>division_segment r = division_segment l\<close>
e7497a1de8b9 more concise instance-specific rules on euclidean relation
haftmann
parents: 76053
diff changeset
  1952
    by (simp add: division_segment_int_def sgn_if split: if_splits)
e7497a1de8b9 more concise instance-specific rules on euclidean relation
haftmann
parents: 76053
diff changeset
  1953
  with \<open>\<bar>r\<bar> < \<bar>l\<bar>\<close> \<open>k = q * l + r\<close>
e7497a1de8b9 more concise instance-specific rules on euclidean relation
haftmann
parents: 76053
diff changeset
  1954
  show ?case
e7497a1de8b9 more concise instance-specific rules on euclidean relation
haftmann
parents: 76053
diff changeset
  1955
    by simp
e7497a1de8b9 more concise instance-specific rules on euclidean relation
haftmann
parents: 76053
diff changeset
  1956
qed
e7497a1de8b9 more concise instance-specific rules on euclidean relation
haftmann
parents: 76053
diff changeset
  1957
66816
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1958
71157
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1959
subsection \<open>Special case: euclidean rings containing the natural numbers\<close>
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1960
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1961
class unique_euclidean_semiring_with_nat = semidom + semiring_char_0 + unique_euclidean_semiring +
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1962
  assumes of_nat_div: "of_nat (m div n) = of_nat m div of_nat n"
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1963
    and division_segment_of_nat [simp]: "division_segment (of_nat n) = 1"
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1964
    and division_segment_euclidean_size [simp]: "division_segment a * of_nat (euclidean_size a) = a"
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1965
begin
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1966
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1967
lemma division_segment_eq_iff:
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1968
  "a = b" if "division_segment a = division_segment b"
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1969
    and "euclidean_size a = euclidean_size b"
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1970
  using that division_segment_euclidean_size [of a] by simp
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1971
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1972
lemma euclidean_size_of_nat [simp]:
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1973
  "euclidean_size (of_nat n) = n"
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1974
proof -
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1975
  have "division_segment (of_nat n) * of_nat (euclidean_size (of_nat n)) = of_nat n"
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1976
    by (fact division_segment_euclidean_size)
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1977
  then show ?thesis by simp
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1978
qed
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1979
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1980
lemma of_nat_euclidean_size:
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1981
  "of_nat (euclidean_size a) = a div division_segment a"
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1982
proof -
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1983
  have "of_nat (euclidean_size a) = division_segment a * of_nat (euclidean_size a) div division_segment a"
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1984
    by (subst nonzero_mult_div_cancel_left) simp_all
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1985
  also have "\<dots> = a div division_segment a"
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1986
    by simp
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1987
  finally show ?thesis .
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1988
qed
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1989
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1990
lemma division_segment_1 [simp]:
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1991
  "division_segment 1 = 1"
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1992
  using division_segment_of_nat [of 1] by simp
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1993
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1994
lemma division_segment_numeral [simp]:
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1995
  "division_segment (numeral k) = 1"
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1996
  using division_segment_of_nat [of "numeral k"] by simp
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1997
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1998
lemma euclidean_size_1 [simp]:
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1999
  "euclidean_size 1 = 1"
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2000
  using euclidean_size_of_nat [of 1] by simp
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2001
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2002
lemma euclidean_size_numeral [simp]:
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2003
  "euclidean_size (numeral k) = numeral k"
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2004
  using euclidean_size_of_nat [of "numeral k"] by simp
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2005
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2006
lemma of_nat_dvd_iff:
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2007
  "of_nat m dvd of_nat n \<longleftrightarrow> m dvd n" (is "?P \<longleftrightarrow> ?Q")
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2008
proof (cases "m = 0")
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2009
  case True
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2010
  then show ?thesis
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2011
    by simp
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2012
next
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2013
  case False
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2014
  show ?thesis
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2015
  proof
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2016
    assume ?Q
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2017
    then show ?P
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2018
      by auto
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2019
  next
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2020
    assume ?P
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2021
    with False have "of_nat n = of_nat n div of_nat m * of_nat m"
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2022
      by simp
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2023
    then have "of_nat n = of_nat (n div m * m)"
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2024
      by (simp add: of_nat_div)
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2025
    then have "n = n div m * m"
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2026
      by (simp only: of_nat_eq_iff)
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2027
    then have "n = m * (n div m)"
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2028
      by (simp add: ac_simps)
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2029
    then show ?Q ..
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2030
  qed
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2031
qed
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2032
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2033
lemma of_nat_mod:
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2034
  "of_nat (m mod n) = of_nat m mod of_nat n"
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2035
proof -
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2036
  have "of_nat m div of_nat n * of_nat n + of_nat m mod of_nat n = of_nat m"
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2037
    by (simp add: div_mult_mod_eq)
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2038
  also have "of_nat m = of_nat (m div n * n + m mod n)"
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2039
    by simp
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2040
  finally show ?thesis
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2041
    by (simp only: of_nat_div of_nat_mult of_nat_add) simp
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2042
qed
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2043
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2044
lemma one_div_two_eq_zero [simp]:
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2045
  "1 div 2 = 0"
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2046
proof -
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2047
  from of_nat_div [symmetric] have "of_nat 1 div of_nat 2 = of_nat 0"
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2048
    by (simp only:) simp
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2049
  then show ?thesis
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2050
    by simp
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2051
qed
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2052
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2053
lemma one_mod_two_eq_one [simp]:
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2054
  "1 mod 2 = 1"
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2055
proof -
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2056
  from of_nat_mod [symmetric] have "of_nat 1 mod of_nat 2 = of_nat 1"
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2057
    by (simp only:) simp
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2058
  then show ?thesis
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2059
    by simp
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2060
qed
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2061
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2062
lemma one_mod_2_pow_eq [simp]:
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2063
  "1 mod (2 ^ n) = of_bool (n > 0)"
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2064
proof -
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2065
  have "1 mod (2 ^ n) = of_nat (1 mod (2 ^ n))"
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2066
    using of_nat_mod [of 1 "2 ^ n"] by simp
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2067
  also have "\<dots> = of_bool (n > 0)"
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2068
    by simp
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2069
  finally show ?thesis .
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2070
qed
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2071
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2072
lemma one_div_2_pow_eq [simp]:
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2073
  "1 div (2 ^ n) = of_bool (n = 0)"
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2074
  using div_mult_mod_eq [of 1 "2 ^ n"] by auto
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2075
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2076
lemma div_mult2_eq':
76053
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
  2077
  \<open>a div (of_nat m * of_nat n) = a div of_nat m div of_nat n\<close>
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
  2078
proof (cases \<open>m = 0 \<or> n = 0\<close>)
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
  2079
  case True
71157
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2080
  then show ?thesis
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2081
    by auto
76053
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
  2082
next
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
  2083
  case False
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
  2084
  then have \<open>m > 0\<close> \<open>n > 0\<close>
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
  2085
    by simp_all
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
  2086
  show ?thesis
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
  2087
  proof (cases \<open>of_nat m * of_nat n dvd a\<close>)
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
  2088
    case True
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
  2089
    then obtain b where \<open>a = (of_nat m * of_nat n) * b\<close> ..
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
  2090
    then have \<open>a = of_nat m * (of_nat n * b)\<close>
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
  2091
      by (simp add: ac_simps)
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
  2092
    then show ?thesis
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
  2093
      by simp
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
  2094
  next
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
  2095
    case False
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
  2096
    define q where \<open>q = a div (of_nat m * of_nat n)\<close>
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
  2097
    define r where \<open>r = a mod (of_nat m * of_nat n)\<close>
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
  2098
    from \<open>m > 0\<close> \<open>n > 0\<close> \<open>\<not> of_nat m * of_nat n dvd a\<close> r_def have "division_segment r = 1"
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
  2099
      using division_segment_of_nat [of "m * n"] by (simp add: division_segment_mod)
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
  2100
    with division_segment_euclidean_size [of r]
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
  2101
    have "of_nat (euclidean_size r) = r"
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
  2102
      by simp
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
  2103
    have "a mod (of_nat m * of_nat n) div (of_nat m * of_nat n) = 0"
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
  2104
      by simp
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
  2105
    with \<open>m > 0\<close> \<open>n > 0\<close> r_def have "r div (of_nat m * of_nat n) = 0"
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
  2106
      by simp
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
  2107
    with \<open>of_nat (euclidean_size r) = r\<close>
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
  2108
    have "of_nat (euclidean_size r) div (of_nat m * of_nat n) = 0"
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
  2109
      by simp
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
  2110
    then have "of_nat (euclidean_size r div (m * n)) = 0"
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
  2111
      by (simp add: of_nat_div)
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
  2112
    then have "of_nat (euclidean_size r div m div n) = 0"
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
  2113
      by (simp add: div_mult2_eq)
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
  2114
    with \<open>of_nat (euclidean_size r) = r\<close> have "r div of_nat m div of_nat n = 0"
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
  2115
      by (simp add: of_nat_div)
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
  2116
    with \<open>m > 0\<close> \<open>n > 0\<close> q_def
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
  2117
    have "q = (r div of_nat m + q * of_nat n * of_nat m div of_nat m) div of_nat n"
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
  2118
      by simp
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
  2119
    moreover have \<open>a = q * (of_nat m * of_nat n) + r\<close>
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
  2120
      by (simp add: q_def r_def div_mult_mod_eq)
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
  2121
    ultimately show \<open>a div (of_nat m * of_nat n) = a div of_nat m div of_nat n\<close>
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
  2122
      using q_def [symmetric] div_plus_div_distrib_dvd_right [of \<open>of_nat m\<close> \<open>q * (of_nat m * of_nat n)\<close> r]
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
  2123
      by (simp add: ac_simps)
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
  2124
  qed
71157
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2125
qed
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2126
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2127
lemma mod_mult2_eq':
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2128
  "a mod (of_nat m * of_nat n) = of_nat m * (a div of_nat m mod of_nat n) + a mod of_nat m"
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2129
proof -
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2130
  have "a div (of_nat m * of_nat n) * (of_nat m * of_nat n) + a mod (of_nat m * of_nat n) = a div of_nat m div of_nat n * of_nat n * of_nat m + (a div of_nat m mod of_nat n * of_nat m + a mod of_nat m)"
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2131
    by (simp add: combine_common_factor div_mult_mod_eq)
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2132
  moreover have "a div of_nat m div of_nat n * of_nat n * of_nat m = of_nat n * of_nat m * (a div of_nat m div of_nat n)"
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2133
    by (simp add: ac_simps)
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2134
  ultimately show ?thesis
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2135
    by (simp add: div_mult2_eq' mult_commute)
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2136
qed
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2137
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2138
lemma div_mult2_numeral_eq:
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2139
  "a div numeral k div numeral l = a div numeral (k * l)" (is "?A = ?B")
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2140
proof -
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2141
  have "?A = a div of_nat (numeral k) div of_nat (numeral l)"
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2142
    by simp
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2143
  also have "\<dots> = a div (of_nat (numeral k) * of_nat (numeral l))"
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2144
    by (fact div_mult2_eq' [symmetric])
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2145
  also have "\<dots> = ?B"
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2146
    by simp
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2147
  finally show ?thesis .
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2148
qed
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2149
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2150
lemma numeral_Bit0_div_2:
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2151
  "numeral (num.Bit0 n) div 2 = numeral n"
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2152
proof -
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2153
  have "numeral (num.Bit0 n) = numeral n + numeral n"
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2154
    by (simp only: numeral.simps)
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2155
  also have "\<dots> = numeral n * 2"
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2156
    by (simp add: mult_2_right)
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2157
  finally have "numeral (num.Bit0 n) div 2 = numeral n * 2 div 2"
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2158
    by simp
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2159
  also have "\<dots> = numeral n"
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2160
    by (rule nonzero_mult_div_cancel_right) simp
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2161
  finally show ?thesis .
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2162
qed
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2163
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2164
lemma numeral_Bit1_div_2:
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2165
  "numeral (num.Bit1 n) div 2 = numeral n"
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2166
proof -
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2167
  have "numeral (num.Bit1 n) = numeral n + numeral n + 1"
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2168
    by (simp only: numeral.simps)
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2169
  also have "\<dots> = numeral n * 2 + 1"
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2170
    by (simp add: mult_2_right)
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2171
  finally have "numeral (num.Bit1 n) div 2 = (numeral n * 2 + 1) div 2"
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2172
    by simp
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2173
  also have "\<dots> = numeral n * 2 div 2 + 1 div 2"
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2174
    using dvd_triv_right by (rule div_plus_div_distrib_dvd_left)
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2175
  also have "\<dots> = numeral n * 2 div 2"
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2176
    by simp
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2177
  also have "\<dots> = numeral n"
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2178
    by (rule nonzero_mult_div_cancel_right) simp
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2179
  finally show ?thesis .
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2180
qed
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2181
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2182
lemma exp_mod_exp:
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2183
  \<open>2 ^ m mod 2 ^ n = of_bool (m < n) * 2 ^ m\<close>
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2184
proof -
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2185
  have \<open>(2::nat) ^ m mod 2 ^ n = of_bool (m < n) * 2 ^ m\<close> (is \<open>?lhs = ?rhs\<close>)
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2186
    by (auto simp add: not_less monoid_mult_class.power_add dest!: le_Suc_ex)
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2187
  then have \<open>of_nat ?lhs = of_nat ?rhs\<close>
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2188
    by simp
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2189
  then show ?thesis
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2190
    by (simp add: of_nat_mod)
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2191
qed
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2192
71412
96d126844adc more theorems
haftmann
parents: 71408
diff changeset
  2193
lemma mask_mod_exp:
71408
554385d4cf59 more theorems
haftmann
parents: 71157
diff changeset
  2194
  \<open>(2 ^ n - 1) mod 2 ^ m = 2 ^ min m n - 1\<close>
554385d4cf59 more theorems
haftmann
parents: 71157
diff changeset
  2195
proof -
554385d4cf59 more theorems
haftmann
parents: 71157
diff changeset
  2196
  have \<open>(2 ^ n - 1) mod 2 ^ m = 2 ^ min m n - (1::nat)\<close> (is \<open>?lhs = ?rhs\<close>)
554385d4cf59 more theorems
haftmann
parents: 71157
diff changeset
  2197
  proof (cases \<open>n \<le> m\<close>)
554385d4cf59 more theorems
haftmann
parents: 71157
diff changeset
  2198
    case True
554385d4cf59 more theorems
haftmann
parents: 71157
diff changeset
  2199
    then show ?thesis
75875
48d032035744 streamlined primitive definitions for integer division
haftmann
parents: 75669
diff changeset
  2200
      by (simp add: Suc_le_lessD)
71408
554385d4cf59 more theorems
haftmann
parents: 71157
diff changeset
  2201
  next
554385d4cf59 more theorems
haftmann
parents: 71157
diff changeset
  2202
    case False
554385d4cf59 more theorems
haftmann
parents: 71157
diff changeset
  2203
    then have \<open>m < n\<close>
554385d4cf59 more theorems
haftmann
parents: 71157
diff changeset
  2204
      by simp
554385d4cf59 more theorems
haftmann
parents: 71157
diff changeset
  2205
    then obtain q where n: \<open>n = Suc q + m\<close>
554385d4cf59 more theorems
haftmann
parents: 71157
diff changeset
  2206
      by (auto dest: less_imp_Suc_add)
554385d4cf59 more theorems
haftmann
parents: 71157
diff changeset
  2207
    then have \<open>min m n = m\<close>
554385d4cf59 more theorems
haftmann
parents: 71157
diff changeset
  2208
      by simp
554385d4cf59 more theorems
haftmann
parents: 71157
diff changeset
  2209
    moreover have \<open>(2::nat) ^ m \<le> 2 * 2 ^ q * 2 ^ m\<close>
554385d4cf59 more theorems
haftmann
parents: 71157
diff changeset
  2210
      using mult_le_mono1 [of 1 \<open>2 * 2 ^ q\<close> \<open>2 ^ m\<close>] by simp
554385d4cf59 more theorems
haftmann
parents: 71157
diff changeset
  2211
    with n have \<open>2 ^ n - 1 = (2 ^ Suc q - 1) * 2 ^ m + (2 ^ m - (1::nat))\<close>
554385d4cf59 more theorems
haftmann
parents: 71157
diff changeset
  2212
      by (simp add: monoid_mult_class.power_add algebra_simps)
554385d4cf59 more theorems
haftmann
parents: 71157
diff changeset
  2213
    ultimately show ?thesis
554385d4cf59 more theorems
haftmann
parents: 71157
diff changeset
  2214
      by (simp only: euclidean_semiring_cancel_class.mod_mult_self3) simp
554385d4cf59 more theorems
haftmann
parents: 71157
diff changeset
  2215
  qed
554385d4cf59 more theorems
haftmann
parents: 71157
diff changeset
  2216
  then have \<open>of_nat ?lhs = of_nat ?rhs\<close>
554385d4cf59 more theorems
haftmann
parents: 71157
diff changeset
  2217
    by simp
554385d4cf59 more theorems
haftmann
parents: 71157
diff changeset
  2218
  then show ?thesis
554385d4cf59 more theorems
haftmann
parents: 71157
diff changeset
  2219
    by (simp add: of_nat_mod of_nat_diff)
554385d4cf59 more theorems
haftmann
parents: 71157
diff changeset
  2220
qed
554385d4cf59 more theorems
haftmann
parents: 71157
diff changeset
  2221
71535
b612edee9b0c more frugal simp rules for bit operations; more pervasive use of bit selector
haftmann
parents: 71413
diff changeset
  2222
lemma of_bool_half_eq_0 [simp]:
b612edee9b0c more frugal simp rules for bit operations; more pervasive use of bit selector
haftmann
parents: 71413
diff changeset
  2223
  \<open>of_bool b div 2 = 0\<close>
b612edee9b0c more frugal simp rules for bit operations; more pervasive use of bit selector
haftmann
parents: 71413
diff changeset
  2224
  by simp
b612edee9b0c more frugal simp rules for bit operations; more pervasive use of bit selector
haftmann
parents: 71413
diff changeset
  2225
71157
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2226
end
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2227
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2228
class unique_euclidean_ring_with_nat = ring + unique_euclidean_semiring_with_nat
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2229
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2230
instance nat :: unique_euclidean_semiring_with_nat
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2231
  by standard (simp_all add: dvd_eq_mod_eq_0)
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2232
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2233
instance int :: unique_euclidean_ring_with_nat
75875
48d032035744 streamlined primitive definitions for integer division
haftmann
parents: 75669
diff changeset
  2234
  by standard (auto simp add: divide_int_def division_segment_int_def elim: contrapos_np)
48d032035744 streamlined primitive definitions for integer division
haftmann
parents: 75669
diff changeset
  2235
75876
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  2236
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  2237
subsection \<open>More on euclidean division on \<^typ>\<open>int\<close>\<close>
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  2238
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  2239
subsubsection \<open>Trivial reduction steps\<close>
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  2240
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  2241
lemma div_pos_pos_trivial [simp]:
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  2242
  "k div l = 0" if "k \<ge> 0" and "k < l" for k l :: int
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  2243
  using that by (simp add: unique_euclidean_semiring_class.div_eq_0_iff division_segment_int_def)
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  2244
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  2245
lemma mod_pos_pos_trivial [simp]:
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  2246
  "k mod l = k" if "k \<ge> 0" and "k < l" for k l :: int
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  2247
  using that by (simp add: mod_eq_self_iff_div_eq_0)
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  2248
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  2249
lemma div_neg_neg_trivial [simp]:
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  2250
  "k div l = 0" if "k \<le> 0" and "l < k" for k l :: int
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  2251
  using that by (cases "k = 0") (simp, simp add: unique_euclidean_semiring_class.div_eq_0_iff division_segment_int_def)
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  2252
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  2253
lemma mod_neg_neg_trivial [simp]:
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  2254
  "k mod l = k" if "k \<le> 0" and "l < k" for k l :: int
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  2255
  using that by (simp add: mod_eq_self_iff_div_eq_0)
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  2256
76224
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2257
lemma
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2258
  div_pos_neg_trivial: \<open>k div l = - 1\<close>  (is ?Q)
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2259
  and mod_pos_neg_trivial: \<open>k mod l = k + l\<close>  (is ?R)
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2260
    if \<open>0 < k\<close> and \<open>k + l \<le> 0\<close> for k l :: int
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2261
proof -
76053
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
  2262
  from that have \<open>l < 0\<close>
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
  2263
    by simp
76224
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2264
  have \<open>(k div l, k mod l) = (- 1, k + l)\<close>
76245
4111c94657b4 slightly less abusive proof pattern
haftmann
parents: 76231
diff changeset
  2265
  proof (induction rule: euclidean_relation_intI)
76224
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2266
    case by0
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2267
    with \<open>l < 0\<close> show ?case
76053
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
  2268
      by simp
76224
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2269
  next
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2270
    case divides
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2271
    from \<open>l dvd k\<close> obtain j where \<open>k = l * j\<close> ..
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2272
    with \<open>l < 0\<close> \<open>0 < k\<close> have \<open>j < 0\<close>
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2273
      by (simp add: zero_less_mult_iff)
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2274
    moreover from \<open>k + l \<le> 0\<close> \<open>k = l * j\<close> have \<open>l * (j + 1) \<le> 0\<close>
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2275
      by (simp add: algebra_simps)
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2276
    with \<open>l < 0\<close> have \<open>j + 1 \<ge> 0\<close>
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2277
      by (simp add: mult_le_0_iff)
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2278
    with \<open>j < 0\<close> have \<open>j = - 1\<close>
76053
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
  2279
      by simp
76224
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2280
    with \<open>k = l * j\<close> show ?case
76053
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
  2281
      by simp
76224
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2282
  next
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2283
    case euclidean_relation
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2284
    with \<open>k + l \<le> 0\<close> have \<open>k + l < 0\<close>
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2285
      by (auto simp add: less_le add_eq_0_iff)
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2286
    with \<open>0 < k\<close> show ?case
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2287
      by simp
76053
3310317cc484 clarified generic euclidean relation
haftmann
parents: 76029
diff changeset
  2288
  qed
76224
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2289
  then show ?Q and ?R
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2290
    by simp_all
75876
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  2291
qed
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  2292
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  2293
text \<open>There is neither \<open>div_neg_pos_trivial\<close> nor \<open>mod_neg_pos_trivial\<close>
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  2294
  because \<^term>\<open>0 div l = 0\<close> would supersede it.\<close>
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  2295
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  2296
76224
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2297
subsubsection \<open>More uniqueness rules\<close>
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2298
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2299
lemma
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2300
  fixes a b q r :: int
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2301
  assumes \<open>a = b * q + r\<close> \<open>0 \<le> r\<close> \<open>r < b\<close>
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2302
  shows int_div_pos_eq:
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2303
      \<open>a div b = q\<close> (is ?Q)
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2304
    and int_mod_pos_eq:
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2305
      \<open>a mod b = r\<close> (is ?R)
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2306
proof -
76245
4111c94657b4 slightly less abusive proof pattern
haftmann
parents: 76231
diff changeset
  2307
  have \<open>(a div b, a mod b) = (q, r)\<close>
4111c94657b4 slightly less abusive proof pattern
haftmann
parents: 76231
diff changeset
  2308
    by (induction rule: euclidean_relation_intI)
4111c94657b4 slightly less abusive proof pattern
haftmann
parents: 76231
diff changeset
  2309
      (use assms in \<open>auto simp add: ac_simps dvd_add_left_iff sgn_1_pos le_less dest: zdvd_imp_le\<close>)
76224
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2310
  then show ?Q and ?R
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2311
    by simp_all
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2312
qed
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2313
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2314
lemma int_div_neg_eq:
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2315
  \<open>a div b = q\<close> if \<open>a = b * q + r\<close> \<open>r \<le> 0\<close> \<open>b < r\<close> for a b q r :: int
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2316
  using that int_div_pos_eq [of a \<open>- b\<close> \<open>- q\<close> \<open>- r\<close>] by simp_all
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2317
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2318
lemma int_mod_neg_eq:
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2319
  \<open>a mod b = r\<close> if \<open>a = b * q + r\<close> \<open>r \<le> 0\<close> \<open>b < r\<close> for a b q r :: int
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2320
  using that int_div_neg_eq [of a b q r] by simp
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2321
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2322
75877
dc758531077b streamlined theorems
haftmann
parents: 75876
diff changeset
  2323
subsubsection \<open>Laws for unary minus\<close>
dc758531077b streamlined theorems
haftmann
parents: 75876
diff changeset
  2324
dc758531077b streamlined theorems
haftmann
parents: 75876
diff changeset
  2325
lemma zmod_zminus1_not_zero:
dc758531077b streamlined theorems
haftmann
parents: 75876
diff changeset
  2326
  fixes k l :: int
dc758531077b streamlined theorems
haftmann
parents: 75876
diff changeset
  2327
  shows "- k mod l \<noteq> 0 \<Longrightarrow> k mod l \<noteq> 0"
dc758531077b streamlined theorems
haftmann
parents: 75876
diff changeset
  2328
  by (simp add: mod_eq_0_iff_dvd)
dc758531077b streamlined theorems
haftmann
parents: 75876
diff changeset
  2329
dc758531077b streamlined theorems
haftmann
parents: 75876
diff changeset
  2330
lemma zmod_zminus2_not_zero:
dc758531077b streamlined theorems
haftmann
parents: 75876
diff changeset
  2331
  fixes k l :: int
dc758531077b streamlined theorems
haftmann
parents: 75876
diff changeset
  2332
  shows "k mod - l \<noteq> 0 \<Longrightarrow> k mod l \<noteq> 0"
dc758531077b streamlined theorems
haftmann
parents: 75876
diff changeset
  2333
  by (simp add: mod_eq_0_iff_dvd)
dc758531077b streamlined theorems
haftmann
parents: 75876
diff changeset
  2334
dc758531077b streamlined theorems
haftmann
parents: 75876
diff changeset
  2335
lemma zdiv_zminus1_eq_if:
dc758531077b streamlined theorems
haftmann
parents: 75876
diff changeset
  2336
  \<open>(- a) div b = (if a mod b = 0 then - (a div b) else - (a div b) - 1)\<close>
dc758531077b streamlined theorems
haftmann
parents: 75876
diff changeset
  2337
  if \<open>b \<noteq> 0\<close> for a b :: int
dc758531077b streamlined theorems
haftmann
parents: 75876
diff changeset
  2338
  using that sgn_not_eq_imp [of b \<open>- a\<close>]
dc758531077b streamlined theorems
haftmann
parents: 75876
diff changeset
  2339
  by (cases \<open>a = 0\<close>) (auto simp add: div_eq_div_abs [of \<open>- a\<close> b] div_eq_div_abs [of a b] sgn_eq_0_iff)
dc758531077b streamlined theorems
haftmann
parents: 75876
diff changeset
  2340
dc758531077b streamlined theorems
haftmann
parents: 75876
diff changeset
  2341
lemma zdiv_zminus2_eq_if:
dc758531077b streamlined theorems
haftmann
parents: 75876
diff changeset
  2342
  \<open>a div (- b) = (if a mod b = 0 then - (a div b) else - (a div b) - 1)\<close>
dc758531077b streamlined theorems
haftmann
parents: 75876
diff changeset
  2343
  if \<open>b \<noteq> 0\<close> for a b :: int
dc758531077b streamlined theorems
haftmann
parents: 75876
diff changeset
  2344
  using that by (auto simp add: zdiv_zminus1_eq_if div_minus_right)
dc758531077b streamlined theorems
haftmann
parents: 75876
diff changeset
  2345
dc758531077b streamlined theorems
haftmann
parents: 75876
diff changeset
  2346
lemma zmod_zminus1_eq_if:
dc758531077b streamlined theorems
haftmann
parents: 75876
diff changeset
  2347
  \<open>(- a) mod b = (if a mod b = 0 then 0 else b - (a mod b))\<close>
dc758531077b streamlined theorems
haftmann
parents: 75876
diff changeset
  2348
  for a b :: int
dc758531077b streamlined theorems
haftmann
parents: 75876
diff changeset
  2349
  by (cases \<open>b = 0\<close>)
dc758531077b streamlined theorems
haftmann
parents: 75876
diff changeset
  2350
    (auto simp flip: minus_div_mult_eq_mod simp add: zdiv_zminus1_eq_if algebra_simps)
dc758531077b streamlined theorems
haftmann
parents: 75876
diff changeset
  2351
dc758531077b streamlined theorems
haftmann
parents: 75876
diff changeset
  2352
lemma zmod_zminus2_eq_if:
dc758531077b streamlined theorems
haftmann
parents: 75876
diff changeset
  2353
  \<open>a mod (- b) = (if a mod b = 0 then 0 else (a mod b) - b)\<close>
dc758531077b streamlined theorems
haftmann
parents: 75876
diff changeset
  2354
  for a b :: int
dc758531077b streamlined theorems
haftmann
parents: 75876
diff changeset
  2355
  by (auto simp add: zmod_zminus1_eq_if mod_minus_right)
dc758531077b streamlined theorems
haftmann
parents: 75876
diff changeset
  2356
dc758531077b streamlined theorems
haftmann
parents: 75876
diff changeset
  2357
75876
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  2358
subsubsection \<open>Borders\<close>
75875
48d032035744 streamlined primitive definitions for integer division
haftmann
parents: 75669
diff changeset
  2359
75876
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  2360
lemma pos_mod_bound [simp]:
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  2361
  "k mod l < l" if "l > 0" for k l :: int
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  2362
proof -
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  2363
  obtain m and s where "k = sgn s * int m"
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  2364
    by (rule int_sgnE)
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  2365
  moreover from that obtain n where "l = sgn 1 * int n"
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  2366
    by (cases l) simp_all
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  2367
  moreover from this that have "n > 0"
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  2368
    by simp
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  2369
  ultimately show ?thesis
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  2370
    by (simp only: modulo_int_unfold)
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  2371
      (auto simp add: mod_greater_zero_iff_not_dvd sgn_1_pos)
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  2372
qed
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  2373
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  2374
lemma neg_mod_bound [simp]:
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  2375
  "l < k mod l" if "l < 0" for k l :: int
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  2376
proof -
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  2377
  obtain m and s where "k = sgn s * int m"
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  2378
    by (rule int_sgnE)
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  2379
  moreover from that obtain q where "l = sgn (- 1) * int (Suc q)"
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  2380
    by (cases l) simp_all
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  2381
  moreover define n where "n = Suc q"
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  2382
  then have "Suc q = n"
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  2383
    by simp
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  2384
  ultimately show ?thesis
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  2385
    by (simp only: modulo_int_unfold)
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  2386
      (auto simp add: mod_greater_zero_iff_not_dvd sgn_1_neg)
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  2387
qed
75875
48d032035744 streamlined primitive definitions for integer division
haftmann
parents: 75669
diff changeset
  2388
75876
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  2389
lemma pos_mod_sign [simp]:
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  2390
  "0 \<le> k mod l" if "l > 0" for k l :: int
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  2391
proof -
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  2392
  obtain m and s where "k = sgn s * int m"
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  2393
    by (rule int_sgnE)
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  2394
  moreover from that obtain n where "l = sgn 1 * int n"
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  2395
    by (cases l) auto
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  2396
  moreover from this that have "n > 0"
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  2397
    by simp
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  2398
  ultimately show ?thesis
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  2399
    by (simp only: modulo_int_unfold) (auto simp add: sgn_1_pos)
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  2400
qed
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  2401
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  2402
lemma neg_mod_sign [simp]:
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  2403
  "k mod l \<le> 0" if "l < 0" for k l :: int
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  2404
proof -
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  2405
  obtain m and s where "k = sgn s * int m"
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  2406
    by (rule int_sgnE)
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  2407
  moreover from that obtain q where "l = sgn (- 1) * int (Suc q)"
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  2408
    by (cases l) simp_all
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  2409
  moreover define n where "n = Suc q"
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  2410
  then have "Suc q = n"
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  2411
    by simp
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  2412
  moreover have \<open>int (m mod n) \<le> int n\<close>
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  2413
    using \<open>Suc q = n\<close> by simp
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  2414
  then have \<open>sgn s * int (m mod n) \<le> int n\<close>
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  2415
    by (cases s \<open>0::int\<close> rule: linorder_cases) simp_all
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  2416
  ultimately show ?thesis
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  2417
    by (simp only: modulo_int_unfold) auto
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  2418
qed
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  2419
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  2420
75881
83e4b6a5e7de streamlined theorems
haftmann
parents: 75880
diff changeset
  2421
subsubsection \<open>Splitting Rules for div and mod\<close>
83e4b6a5e7de streamlined theorems
haftmann
parents: 75880
diff changeset
  2422
83e4b6a5e7de streamlined theorems
haftmann
parents: 75880
diff changeset
  2423
lemma split_zdiv:
83e4b6a5e7de streamlined theorems
haftmann
parents: 75880
diff changeset
  2424
  \<open>P (n div k) \<longleftrightarrow>
83e4b6a5e7de streamlined theorems
haftmann
parents: 75880
diff changeset
  2425
    (k = 0 \<longrightarrow> P 0) \<and>
83e4b6a5e7de streamlined theorems
haftmann
parents: 75880
diff changeset
  2426
    (0 < k \<longrightarrow> (\<forall>i j. 0 \<le> j \<and> j < k \<and> n = k * i + j \<longrightarrow> P i)) \<and>
83e4b6a5e7de streamlined theorems
haftmann
parents: 75880
diff changeset
  2427
    (k < 0 \<longrightarrow> (\<forall>i j. k < j \<and> j \<le> 0 \<and> n = k * i + j \<longrightarrow> P i))\<close> (is ?div)
83e4b6a5e7de streamlined theorems
haftmann
parents: 75880
diff changeset
  2428
  and split_zmod:
83e4b6a5e7de streamlined theorems
haftmann
parents: 75880
diff changeset
  2429
  \<open>Q (n mod k) \<longleftrightarrow>
83e4b6a5e7de streamlined theorems
haftmann
parents: 75880
diff changeset
  2430
    (k = 0 \<longrightarrow> Q n) \<and>
83e4b6a5e7de streamlined theorems
haftmann
parents: 75880
diff changeset
  2431
    (0 < k \<longrightarrow> (\<forall>i j. 0 \<le> j \<and> j < k \<and> n = k * i + j \<longrightarrow> Q j)) \<and>
83e4b6a5e7de streamlined theorems
haftmann
parents: 75880
diff changeset
  2432
    (k < 0 \<longrightarrow> (\<forall>i j. k < j \<and> j \<le> 0 \<and> n = k * i + j \<longrightarrow> Q j))\<close> (is ?mod)
83e4b6a5e7de streamlined theorems
haftmann
parents: 75880
diff changeset
  2433
  for n k :: int
83e4b6a5e7de streamlined theorems
haftmann
parents: 75880
diff changeset
  2434
proof -
83e4b6a5e7de streamlined theorems
haftmann
parents: 75880
diff changeset
  2435
  have *: \<open>R (n div k) (n mod k) \<longleftrightarrow>
83e4b6a5e7de streamlined theorems
haftmann
parents: 75880
diff changeset
  2436
    (k = 0 \<longrightarrow> R 0 n) \<and>
83e4b6a5e7de streamlined theorems
haftmann
parents: 75880
diff changeset
  2437
    (0 < k \<longrightarrow> (\<forall>i j. 0 \<le> j \<and> j < k \<and> n = k * i + j \<longrightarrow> R i j)) \<and>
83e4b6a5e7de streamlined theorems
haftmann
parents: 75880
diff changeset
  2438
    (k < 0 \<longrightarrow> (\<forall>i j. k < j \<and> j \<le> 0 \<and> n = k * i + j \<longrightarrow> R i j))\<close> for R
83e4b6a5e7de streamlined theorems
haftmann
parents: 75880
diff changeset
  2439
    by (cases \<open>k = 0\<close>)
83e4b6a5e7de streamlined theorems
haftmann
parents: 75880
diff changeset
  2440
      (auto simp add: linorder_class.neq_iff)
83e4b6a5e7de streamlined theorems
haftmann
parents: 75880
diff changeset
  2441
  from * [of \<open>\<lambda>q _. P q\<close>] show ?div .
83e4b6a5e7de streamlined theorems
haftmann
parents: 75880
diff changeset
  2442
  from * [of \<open>\<lambda>_ r. Q r\<close>] show ?mod .
83e4b6a5e7de streamlined theorems
haftmann
parents: 75880
diff changeset
  2443
qed
76224
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2444
75881
83e4b6a5e7de streamlined theorems
haftmann
parents: 75880
diff changeset
  2445
text \<open>Enable (lin)arith to deal with \<^const>\<open>divide\<close> and \<^const>\<open>modulo\<close>
83e4b6a5e7de streamlined theorems
haftmann
parents: 75880
diff changeset
  2446
  when these are applied to some constant that is of the form
83e4b6a5e7de streamlined theorems
haftmann
parents: 75880
diff changeset
  2447
  \<^term>\<open>numeral k\<close>:\<close>
83e4b6a5e7de streamlined theorems
haftmann
parents: 75880
diff changeset
  2448
declare split_zdiv [of _ _ \<open>numeral n\<close>, linarith_split] for n
83e4b6a5e7de streamlined theorems
haftmann
parents: 75880
diff changeset
  2449
declare split_zdiv [of _ _ \<open>- numeral n\<close>, linarith_split] for n
83e4b6a5e7de streamlined theorems
haftmann
parents: 75880
diff changeset
  2450
declare split_zmod [of _ _ \<open>numeral n\<close>, linarith_split] for n
83e4b6a5e7de streamlined theorems
haftmann
parents: 75880
diff changeset
  2451
declare split_zmod [of _ _ \<open>- numeral n\<close>, linarith_split] for n
83e4b6a5e7de streamlined theorems
haftmann
parents: 75880
diff changeset
  2452
83e4b6a5e7de streamlined theorems
haftmann
parents: 75880
diff changeset
  2453
lemma zdiv_eq_0_iff:
83e4b6a5e7de streamlined theorems
haftmann
parents: 75880
diff changeset
  2454
  "i div k = 0 \<longleftrightarrow> k = 0 \<or> 0 \<le> i \<and> i < k \<or> i \<le> 0 \<and> k < i" (is "?L = ?R")
83e4b6a5e7de streamlined theorems
haftmann
parents: 75880
diff changeset
  2455
  for i k :: int
83e4b6a5e7de streamlined theorems
haftmann
parents: 75880
diff changeset
  2456
proof
83e4b6a5e7de streamlined theorems
haftmann
parents: 75880
diff changeset
  2457
  assume ?L
83e4b6a5e7de streamlined theorems
haftmann
parents: 75880
diff changeset
  2458
  moreover have "?L \<longrightarrow> ?R"
83e4b6a5e7de streamlined theorems
haftmann
parents: 75880
diff changeset
  2459
    by (rule split_zdiv [THEN iffD2]) simp
83e4b6a5e7de streamlined theorems
haftmann
parents: 75880
diff changeset
  2460
  ultimately show ?R
83e4b6a5e7de streamlined theorems
haftmann
parents: 75880
diff changeset
  2461
    by blast
83e4b6a5e7de streamlined theorems
haftmann
parents: 75880
diff changeset
  2462
next
83e4b6a5e7de streamlined theorems
haftmann
parents: 75880
diff changeset
  2463
  assume ?R then show ?L
83e4b6a5e7de streamlined theorems
haftmann
parents: 75880
diff changeset
  2464
    by auto
83e4b6a5e7de streamlined theorems
haftmann
parents: 75880
diff changeset
  2465
qed
83e4b6a5e7de streamlined theorems
haftmann
parents: 75880
diff changeset
  2466
83e4b6a5e7de streamlined theorems
haftmann
parents: 75880
diff changeset
  2467
lemma zmod_trivial_iff:
83e4b6a5e7de streamlined theorems
haftmann
parents: 75880
diff changeset
  2468
  fixes i k :: int
83e4b6a5e7de streamlined theorems
haftmann
parents: 75880
diff changeset
  2469
  shows "i mod k = i \<longleftrightarrow> k = 0 \<or> 0 \<le> i \<and> i < k \<or> i \<le> 0 \<and> k < i"
83e4b6a5e7de streamlined theorems
haftmann
parents: 75880
diff changeset
  2470
proof -
83e4b6a5e7de streamlined theorems
haftmann
parents: 75880
diff changeset
  2471
  have "i mod k = i \<longleftrightarrow> i div k = 0"
83e4b6a5e7de streamlined theorems
haftmann
parents: 75880
diff changeset
  2472
    using div_mult_mod_eq [of i k] by safe auto
83e4b6a5e7de streamlined theorems
haftmann
parents: 75880
diff changeset
  2473
  with zdiv_eq_0_iff
83e4b6a5e7de streamlined theorems
haftmann
parents: 75880
diff changeset
  2474
  show ?thesis
83e4b6a5e7de streamlined theorems
haftmann
parents: 75880
diff changeset
  2475
    by simp
83e4b6a5e7de streamlined theorems
haftmann
parents: 75880
diff changeset
  2476
qed
83e4b6a5e7de streamlined theorems
haftmann
parents: 75880
diff changeset
  2477
83e4b6a5e7de streamlined theorems
haftmann
parents: 75880
diff changeset
  2478
75876
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  2479
subsubsection \<open>Algebraic rewrites\<close>
71157
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2480
74592
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 73853
diff changeset
  2481
lemma zdiv_zmult2_eq:
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 73853
diff changeset
  2482
  \<open>a div (b * c) = (a div b) div c\<close> if \<open>c \<ge> 0\<close> for a b c :: int
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 73853
diff changeset
  2483
proof (cases \<open>b \<ge> 0\<close>)
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 73853
diff changeset
  2484
  case True
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 73853
diff changeset
  2485
  with that show ?thesis
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 73853
diff changeset
  2486
    using div_mult2_eq' [of a \<open>nat b\<close> \<open>nat c\<close>] by simp
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 73853
diff changeset
  2487
next
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 73853
diff changeset
  2488
  case False
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 73853
diff changeset
  2489
  with that show ?thesis
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 73853
diff changeset
  2490
    using div_mult2_eq' [of \<open>- a\<close> \<open>nat (- b)\<close> \<open>nat c\<close>] by simp
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 73853
diff changeset
  2491
qed
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 73853
diff changeset
  2492
75876
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  2493
lemma zdiv_zmult2_eq':
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  2494
  \<open>k div (l * j) = ((sgn j * k) div l) div \<bar>j\<bar>\<close> for k l j :: int
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  2495
proof -
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  2496
  have \<open>k div (l * j) = (sgn j * k) div (sgn j * (l * j))\<close>
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  2497
    by (simp add: sgn_0_0)
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  2498
  also have \<open>sgn j * (l * j) = l * \<bar>j\<bar>\<close>
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  2499
    by (simp add: mult.left_commute [of _ l] abs_sgn) (simp add: ac_simps)
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  2500
  also have \<open>(sgn j * k) div (l * \<bar>j\<bar>) = ((sgn j * k) div l) div \<bar>j\<bar>\<close>
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  2501
    by (simp add: zdiv_zmult2_eq)
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  2502
  finally show ?thesis .
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  2503
qed
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  2504
74592
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 73853
diff changeset
  2505
lemma zmod_zmult2_eq:
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 73853
diff changeset
  2506
  \<open>a mod (b * c) = b * (a div b mod c) + a mod b\<close> if \<open>c \<ge> 0\<close> for a b c :: int
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 73853
diff changeset
  2507
proof (cases \<open>b \<ge> 0\<close>)
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 73853
diff changeset
  2508
  case True
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 73853
diff changeset
  2509
  with that show ?thesis
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 73853
diff changeset
  2510
    using mod_mult2_eq' [of a \<open>nat b\<close> \<open>nat c\<close>] by simp
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 73853
diff changeset
  2511
next
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 73853
diff changeset
  2512
  case False
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 73853
diff changeset
  2513
  with that show ?thesis
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 73853
diff changeset
  2514
    using mod_mult2_eq' [of \<open>- a\<close> \<open>nat (- b)\<close> \<open>nat c\<close>] by simp
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 73853
diff changeset
  2515
qed
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 73853
diff changeset
  2516
75881
83e4b6a5e7de streamlined theorems
haftmann
parents: 75880
diff changeset
  2517
lemma half_nonnegative_int_iff [simp]:
83e4b6a5e7de streamlined theorems
haftmann
parents: 75880
diff changeset
  2518
  \<open>k div 2 \<ge> 0 \<longleftrightarrow> k \<ge> 0\<close> for k :: int
83e4b6a5e7de streamlined theorems
haftmann
parents: 75880
diff changeset
  2519
  by auto
83e4b6a5e7de streamlined theorems
haftmann
parents: 75880
diff changeset
  2520
83e4b6a5e7de streamlined theorems
haftmann
parents: 75880
diff changeset
  2521
lemma half_negative_int_iff [simp]:
83e4b6a5e7de streamlined theorems
haftmann
parents: 75880
diff changeset
  2522
  \<open>k div 2 < 0 \<longleftrightarrow> k < 0\<close> for k :: int
83e4b6a5e7de streamlined theorems
haftmann
parents: 75880
diff changeset
  2523
  by auto
83e4b6a5e7de streamlined theorems
haftmann
parents: 75880
diff changeset
  2524
75876
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  2525
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  2526
subsubsection \<open>Distributive laws for conversions.\<close>
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  2527
75875
48d032035744 streamlined primitive definitions for integer division
haftmann
parents: 75669
diff changeset
  2528
lemma zdiv_int:
48d032035744 streamlined primitive definitions for integer division
haftmann
parents: 75669
diff changeset
  2529
  "int (a div b) = int a div int b"
48d032035744 streamlined primitive definitions for integer division
haftmann
parents: 75669
diff changeset
  2530
  by (fact of_nat_div)
48d032035744 streamlined primitive definitions for integer division
haftmann
parents: 75669
diff changeset
  2531
48d032035744 streamlined primitive definitions for integer division
haftmann
parents: 75669
diff changeset
  2532
lemma zmod_int:
48d032035744 streamlined primitive definitions for integer division
haftmann
parents: 75669
diff changeset
  2533
  "int (a mod b) = int a mod int b"
48d032035744 streamlined primitive definitions for integer division
haftmann
parents: 75669
diff changeset
  2534
  by (fact of_nat_mod)
48d032035744 streamlined primitive definitions for integer division
haftmann
parents: 75669
diff changeset
  2535
75876
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  2536
lemma nat_div_distrib:
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  2537
  \<open>nat (x div y) = nat x div nat y\<close> if \<open>0 \<le> x\<close>
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  2538
  using that by (simp add: divide_int_def sgn_if)
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  2539
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  2540
lemma nat_div_distrib':
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  2541
  \<open>nat (x div y) = nat x div nat y\<close> if \<open>0 \<le> y\<close>
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  2542
  using that by (simp add: divide_int_def sgn_if)
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  2543
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  2544
lemma nat_mod_distrib: \<comment> \<open>Fails if y<0: the LHS collapses to (nat z) but the RHS doesn't\<close>
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  2545
  \<open>nat (x mod y) = nat x mod nat y\<close> if \<open>0 \<le> x\<close> \<open>0 \<le> y\<close>
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  2546
  using that by (simp add: modulo_int_def sgn_if)
647879691c1c streamlined theorems and sections
haftmann
parents: 75875
diff changeset
  2547
71157
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2548
76224
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2549
subsubsection \<open>Monotonicity in the First Argument (Dividend)\<close>
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2550
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2551
lemma zdiv_mono1:
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2552
  \<open>a div b \<le> a' div b\<close>
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2553
    if \<open>a \<le> a'\<close> \<open>0 < b\<close>
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2554
    for a b b' :: int
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2555
proof -
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2556
  from \<open>a \<le> a'\<close> have \<open>b * (a div b) + a mod b \<le> b * (a' div b) + a' mod b\<close>
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2557
    by simp
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2558
  then have \<open>b * (a div b) \<le> (a' mod b - a mod b) + b * (a' div b)\<close>
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2559
    by (simp add: algebra_simps)
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2560
  moreover have \<open>a' mod b < b + a mod b\<close>
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2561
    by (rule less_le_trans [of _ b]) (use \<open>0 < b\<close> in simp_all)
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2562
  ultimately have \<open>b * (a div b) < b * (1 + a' div b)\<close>
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2563
    by (simp add: distrib_left)
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2564
  with \<open>0 < b\<close> have \<open>a div b < 1 + a' div b\<close>
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2565
    by (simp add: mult_less_cancel_left)
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2566
  then show ?thesis
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2567
    by simp
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2568
qed
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2569
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2570
lemma zdiv_mono1_neg:
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2571
  \<open>a' div b \<le> a div b\<close>
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2572
    if \<open>a \<le> a'\<close> \<open>b < 0\<close>
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2573
    for a a' b :: int
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2574
  using that zdiv_mono1 [of \<open>- a'\<close> \<open>- a\<close> \<open>- b\<close>] by simp
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2575
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2576
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2577
subsubsection \<open>Monotonicity in the Second Argument (Divisor)\<close>
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2578
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2579
lemma zdiv_mono2:
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2580
  \<open>a div b \<le> a div b'\<close> if \<open>0 \<le> a\<close> \<open>0 < b'\<close> \<open>b' \<le> b\<close> for a b b' :: int
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2581
proof -
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2582
  define q q' r r' where **: \<open>q = a div b\<close> \<open>q' = a div b'\<close> \<open>r = a mod b\<close> \<open>r' = a mod b'\<close>
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2583
  then have *: \<open>b * q + r = b' * q' + r'\<close> \<open>0 \<le> b' * q' + r'\<close>
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2584
    \<open>r' < b'\<close> \<open>0 \<le> r\<close> \<open>0 < b'\<close> \<open>b' \<le> b\<close>
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2585
    using that by simp_all
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2586
  have \<open>0 < b' * (q' + 1)\<close>
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2587
    using * by (simp add: distrib_left)
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2588
  with * have \<open>0 \<le> q'\<close>
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2589
    by (simp add: zero_less_mult_iff)
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2590
  moreover have \<open>b * q = r' - r + b' * q'\<close>
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2591
    using * by linarith
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2592
  ultimately have \<open>b * q < b * (q' + 1)\<close>
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2593
    using mult_right_mono * unfolding distrib_left by fastforce
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2594
  with * have \<open>q \<le> q'\<close>
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2595
    by (simp add: mult_less_cancel_left_pos)
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2596
  with ** show ?thesis
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2597
    by simp
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2598
qed
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2599
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2600
lemma zdiv_mono2_neg:
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2601
  \<open>a div b' \<le> a div b\<close> if \<open>a < 0\<close> \<open>0 < b'\<close> \<open>b' \<le> b\<close> for a b b' :: int
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2602
proof -
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2603
  define q q' r r' where **: \<open>q = a div b\<close> \<open>q' = a div b'\<close> \<open>r = a mod b\<close> \<open>r' = a mod b'\<close>
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2604
  then have *: \<open>b * q + r = b' * q' + r'\<close> \<open>b' * q' + r' < 0\<close>
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2605
    \<open>r < b\<close> \<open>0 \<le> r'\<close> \<open>0 < b'\<close> \<open>b' \<le> b\<close>
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2606
    using that by simp_all
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2607
  have \<open>b' * q' < 0\<close>
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2608
    using * by linarith
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2609
  with * have \<open>q' \<le> 0\<close>
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2610
    by (simp add: mult_less_0_iff)
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2611
  have \<open>b * q' \<le> b' * q'\<close>
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2612
    by (simp add: \<open>q' \<le> 0\<close> * mult_right_mono_neg)
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2613
  then have "b * q' < b * (q + 1)"
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2614
    using * by (simp add: distrib_left)
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2615
  then have \<open>q' \<le> q\<close>
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2616
    using * by (simp add: mult_less_cancel_left)
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2617
  then show ?thesis
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2618
    by (simp add: **)
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2619
qed
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2620
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2621
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2622
subsubsection \<open>Quotients of Signs\<close>
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2623
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2624
lemma div_eq_minus1:
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2625
  \<open>0 < b \<Longrightarrow> - 1 div b = - 1\<close> for b :: int
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2626
  by (simp add: divide_int_def)
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2627
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2628
lemma zmod_minus1:
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2629
  \<open>0 < b \<Longrightarrow> - 1 mod b = b - 1\<close> for b :: int
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2630
  by (auto simp add: modulo_int_def)
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2631
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2632
lemma minus_mod_int_eq:
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2633
  \<open>- k mod l = l - 1 - (k - 1) mod l\<close> if \<open>l \<ge> 0\<close> for k l :: int
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2634
proof (cases \<open>l = 0\<close>)
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2635
  case True
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2636
  then show ?thesis
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2637
    by simp
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2638
next
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2639
  case False
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2640
  with that have \<open>l > 0\<close>
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2641
    by simp
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2642
  then show ?thesis
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2643
  proof (cases \<open>l dvd k\<close>)
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2644
    case True
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2645
    then obtain j where \<open>k = l * j\<close> ..
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2646
    moreover have \<open>(l * j mod l - 1) mod l = l - 1\<close>
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2647
      using \<open>l > 0\<close> by (simp add: zmod_minus1)
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2648
    then have \<open>(l * j - 1) mod l = l - 1\<close>
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2649
      by (simp only: mod_simps)
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2650
    ultimately show ?thesis
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2651
      by simp
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2652
  next
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2653
    case False
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2654
    moreover have 1: \<open>0 < k mod l\<close>
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2655
      using \<open>0 < l\<close> False le_less by fastforce
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2656
    moreover have 2: \<open>k mod l < 1 + l\<close>
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2657
      using \<open>0 < l\<close> pos_mod_bound[of l k] by linarith
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2658
    from 1 2 \<open>l > 0\<close> have \<open>(k mod l - 1) mod l = k mod l - 1\<close>
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2659
      by (simp add: zmod_trivial_iff)
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2660
    ultimately show ?thesis
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2661
      by (simp only: zmod_zminus1_eq_if)
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2662
         (simp add: mod_eq_0_iff_dvd algebra_simps mod_simps)
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2663
  qed
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2664
qed
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2665
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2666
lemma div_neg_pos_less0:
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2667
  \<open>a div b < 0\<close> if \<open>a < 0\<close> \<open>0 < b\<close> for a b :: int
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2668
proof -
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2669
  have "a div b \<le> - 1 div b"
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2670
    using zdiv_mono1 that by auto
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2671
  also have "... \<le> -1"
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2672
    by (simp add: that(2) div_eq_minus1)
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2673
  finally show ?thesis
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2674
    by force
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2675
qed
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2676
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2677
lemma div_nonneg_neg_le0:
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2678
  \<open>a div b \<le> 0\<close> if \<open>0 \<le> a\<close> \<open>b < 0\<close> for a b :: int
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2679
  using that by (auto dest: zdiv_mono1_neg)
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2680
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2681
lemma div_nonpos_pos_le0:
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2682
  \<open>a div b \<le> 0\<close> if \<open>a \<le> 0\<close> \<open>0 < b\<close> for a b :: int
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2683
  using that by (auto dest: zdiv_mono1)
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2684
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2685
text\<open>Now for some equivalences of the form \<open>a div b >=< 0 \<longleftrightarrow> \<dots>\<close>
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2686
conditional upon the sign of \<open>a\<close> or \<open>b\<close>. There are many more.
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2687
They should all be simp rules unless that causes too much search.\<close>
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2688
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2689
lemma pos_imp_zdiv_nonneg_iff:
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2690
  \<open>0 \<le> a div b \<longleftrightarrow> 0 \<le> a\<close>
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2691
  if \<open>0 < b\<close> for a b :: int
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2692
proof
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2693
  assume \<open>0 \<le> a div b\<close>
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2694
  show \<open>0 \<le> a\<close>
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2695
  proof (rule ccontr)
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2696
    assume \<open>\<not> 0 \<le> a\<close>
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2697
    then have \<open>a < 0\<close>
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2698
      by simp
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2699
    then have \<open>a div b < 0\<close>
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2700
      using that by (rule div_neg_pos_less0)
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2701
    with \<open>0 \<le> a div b\<close> show False
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2702
      by simp
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2703
  qed
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2704
next
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2705
  assume "0 \<le> a"
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2706
  then have "0 div b \<le> a div b"
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2707
    using zdiv_mono1 that by blast
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2708
  then show "0 \<le> a div b"
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2709
    by auto
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2710
qed
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2711
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2712
lemma neg_imp_zdiv_nonneg_iff:
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2713
  \<open>0 \<le> a div b \<longleftrightarrow> a \<le> 0\<close> if \<open>b < 0\<close> for a b :: int
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2714
  using that pos_imp_zdiv_nonneg_iff [of \<open>- b\<close> \<open>- a\<close>] by simp
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2715
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2716
lemma pos_imp_zdiv_pos_iff:
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2717
  \<open>0 < (i::int) div k \<longleftrightarrow> k \<le> i\<close> if \<open>0 < k\<close> for i k :: int
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2718
  using that pos_imp_zdiv_nonneg_iff [of k i] zdiv_eq_0_iff [of i k] by arith
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2719
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2720
lemma pos_imp_zdiv_neg_iff:
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2721
  \<open>a div b < 0 \<longleftrightarrow> a < 0\<close> if \<open>0 < b\<close> for a b :: int
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2722
    \<comment> \<open>But not \<^prop>\<open>a div b \<le> 0 \<longleftrightarrow> a \<le> 0\<close>; consider \<^prop>\<open>a = 1\<close>, \<^prop>\<open>b = 2\<close> when \<^prop>\<open>a div b = 0\<close>.\<close>
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2723
  using that by (simp add: pos_imp_zdiv_nonneg_iff flip: linorder_not_le)
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2724
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2725
lemma neg_imp_zdiv_neg_iff:
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2726
    \<comment> \<open>But not \<^prop>\<open>a div b \<le> 0 \<longleftrightarrow> 0 \<le> a\<close>; consider \<^prop>\<open>a = - 1\<close>, \<^prop>\<open>b = - 2\<close> when \<^prop>\<open>a div b = 0\<close>.\<close>
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2727
  \<open>a div b < 0 \<longleftrightarrow> 0 < a\<close> if \<open>b < 0\<close> for a b :: int
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2728
  using that by (simp add: neg_imp_zdiv_nonneg_iff flip: linorder_not_le)
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2729
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2730
lemma nonneg1_imp_zdiv_pos_iff:
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2731
  \<open>a div b > 0 \<longleftrightarrow> a \<ge> b \<and> b > 0\<close> if \<open>0 \<le> a\<close> for a b :: int
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2732
proof -
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2733
  have "0 < a div b \<Longrightarrow> b \<le> a"
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2734
    using div_pos_pos_trivial[of a b] that by arith
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2735
  moreover have "0 < a div b \<Longrightarrow> b > 0"
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2736
    using that div_nonneg_neg_le0[of a b] by (cases "b=0"; force)
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2737
  moreover have "b \<le> a \<and> 0 < b \<Longrightarrow> 0 < a div b"
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2738
    using int_one_le_iff_zero_less[of "a div b"] zdiv_mono1[of b a b] by simp
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2739
  ultimately show ?thesis
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2740
    by blast
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2741
qed
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2742
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2743
lemma zmod_le_nonneg_dividend:
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2744
  \<open>m mod k \<le> m\<close> if \<open>(m::int) \<ge> 0\<close> for m k :: int
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2745
proof -
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2746
  from that have \<open>m > 0 \<or> m = 0\<close>
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2747
    by auto
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2748
  then show ?thesis proof
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2749
    assume \<open>m = 0\<close> then show ?thesis
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2750
      by simp
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2751
  next
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2752
    assume \<open>m > 0\<close> then show ?thesis
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2753
    proof (cases k \<open>0::int\<close> rule: linorder_cases)
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2754
      case less
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2755
      moreover define l where \<open>l = - k\<close>
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2756
      ultimately have \<open>l > 0\<close>
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2757
        by simp
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2758
      with \<open>m > 0\<close> have \<open>int (nat m mod nat l) \<le> m\<close>
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2759
        by (simp flip: le_nat_iff)
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2760
      then have \<open>int (nat m mod nat l) - l \<le> m\<close>
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2761
        using \<open>l > 0\<close> by simp
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2762
      with \<open>m > 0\<close> \<open>l > 0\<close> show ?thesis
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2763
        by (simp add: modulo_int_def l_def flip: le_nat_iff)
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2764
    qed (simp_all add: modulo_int_def flip: le_nat_iff)
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2765
  qed
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2766
qed
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2767
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2768
lemma sgn_div_eq_sgn_mult:
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2769
  \<open>sgn (k div l) = of_bool (k div l \<noteq> 0) * sgn (k * l)\<close>
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2770
  for k l :: int
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2771
proof (cases \<open>k div l = 0\<close>)
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2772
  case True
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2773
  then show ?thesis
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2774
    by simp
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2775
next
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2776
  case False
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2777
  have \<open>0 \<le> \<bar>k\<bar> div \<bar>l\<bar>\<close>
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2778
    by (cases \<open>l = 0\<close>) (simp_all add: pos_imp_zdiv_nonneg_iff)
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2779
  then have \<open>\<bar>k\<bar> div \<bar>l\<bar> \<noteq> 0 \<longleftrightarrow> 0 < \<bar>k\<bar> div \<bar>l\<bar>\<close>
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2780
    by (simp add: less_le)
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2781
  also have \<open>\<dots> \<longleftrightarrow> \<bar>k\<bar> \<ge> \<bar>l\<bar>\<close>
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2782
    using False nonneg1_imp_zdiv_pos_iff by auto
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2783
  finally have *: \<open>\<bar>k\<bar> div \<bar>l\<bar> \<noteq> 0 \<longleftrightarrow> \<bar>l\<bar> \<le> \<bar>k\<bar>\<close> .
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2784
  show ?thesis
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2785
    using \<open>0 \<le> \<bar>k\<bar> div \<bar>l\<bar>\<close> False
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2786
  by (auto simp add: div_eq_div_abs [of k l] div_eq_sgn_abs [of k l]
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2787
    sgn_mult sgn_1_pos sgn_1_neg sgn_eq_0_iff nonneg1_imp_zdiv_pos_iff * dest: sgn_not_eq_imp)
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2788
qed
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2789
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2790
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2791
subsubsection \<open>Further properties\<close>
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2792
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2793
lemma div_int_pos_iff:
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2794
  "k div l \<ge> 0 \<longleftrightarrow> k = 0 \<or> l = 0 \<or> k \<ge> 0 \<and> l \<ge> 0
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2795
    \<or> k < 0 \<and> l < 0"
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2796
  for k l :: int
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2797
proof (cases "k = 0 \<or> l = 0")
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2798
  case False
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2799
  then have *: "k \<noteq> 0" "l \<noteq> 0"
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2800
    by auto
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2801
  then have "0 \<le> k div l \<Longrightarrow> \<not> k < 0 \<Longrightarrow> 0 \<le> l"
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2802
    by (meson neg_imp_zdiv_neg_iff not_le not_less_iff_gr_or_eq)
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2803
  then show ?thesis
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2804
   using * by (auto simp add: pos_imp_zdiv_nonneg_iff neg_imp_zdiv_nonneg_iff)
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2805
qed auto
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2806
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2807
lemma mod_int_pos_iff:
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2808
  \<open>k mod l \<ge> 0 \<longleftrightarrow> l dvd k \<or> l = 0 \<and> k \<ge> 0 \<or> l > 0\<close>
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2809
  for k l :: int
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2810
proof (cases "l > 0")
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2811
  case False
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2812
  then show ?thesis
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2813
    by (simp add: dvd_eq_mod_eq_0) (use neg_mod_sign [of l k] in \<open>auto simp add: le_less not_less\<close>)
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2814
qed auto
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2815
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2816
lemma abs_div:
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2817
  \<open>\<bar>x div y\<bar> = \<bar>x\<bar> div \<bar>y\<bar>\<close> if \<open>y dvd x\<close> for x y :: int
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2818
  using that by (cases \<open>y = 0\<close>) (auto simp add: abs_mult)
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2819
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2820
lemma int_power_div_base: \<^marker>\<open>contributor \<open>Matthias Daum\<close>\<close>
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2821
  \<open>k ^ m div k = k ^ (m - Suc 0)\<close> if \<open>0 < m\<close> \<open>0 < k\<close> for k :: int
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2822
  using that by (cases m) simp_all
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2823
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2824
lemma int_div_less_self: \<^marker>\<open>contributor \<open>Matthias Daum\<close>\<close>
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2825
  \<open>x div k < x\<close> if \<open>0 < x\<close> \<open>1 < k\<close> for x k :: int
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2826
proof -
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2827
  from that have \<open>nat (x div k) = nat x div nat k\<close>
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2828
    by (simp add: nat_div_distrib)
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2829
  also from that have \<open>nat x div nat k < nat x\<close>
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2830
    by simp
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2831
  finally show ?thesis
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2832
    by simp
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2833
qed
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2834
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2835
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2836
subsubsection \<open>Computing \<open>div\<close> and \<open>mod\<close> by shifting\<close>
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2837
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2838
lemma div_pos_geq:
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2839
  \<open>k div l = (k - l) div l + 1\<close> if \<open>0 < l\<close> \<open>l \<le> k\<close> for k l :: int
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2840
proof -
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2841
  have "k = (k - l) + l" by simp
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2842
  then obtain j where k: "k = j + l" ..
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2843
  with that show ?thesis by (simp add: div_add_self2)
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2844
qed
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2845
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2846
lemma mod_pos_geq:
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2847
  \<open>k mod l = (k - l) mod l\<close>  if \<open>0 < l\<close> \<open>l \<le> k\<close> for k l :: int
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2848
proof -
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2849
  have "k = (k - l) + l" by simp
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2850
  then obtain j where k: "k = j + l" ..
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2851
  with that show ?thesis by simp
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2852
qed
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2853
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2854
lemma pos_zdiv_mult_2: \<open>(1 + 2 * b) div (2 * a) = b div a\<close> (is ?Q)
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2855
  and pos_zmod_mult_2: \<open>(1 + 2 * b) mod (2 * a) = 1 + 2 * (b mod a)\<close> (is ?R)
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2856
  if \<open>0 \<le> a\<close> for a b :: int
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2857
proof -
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2858
  have \<open>((1 + 2 * b) div (2 * a), (1 + 2 * b) mod (2 * a)) = (b div a, 1 + 2 * (b mod a))\<close>
76245
4111c94657b4 slightly less abusive proof pattern
haftmann
parents: 76231
diff changeset
  2859
  proof (induction rule: euclidean_relation_intI)
76224
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2860
    case by0
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2861
    then show ?case
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2862
      by simp
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2863
  next
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2864
    case divides
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2865
    have \<open>2 dvd (2 * a)\<close>
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2866
      by simp
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2867
    then have \<open>2 dvd (1 + 2 * b)\<close>
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2868
      using \<open>2 * a dvd 1 + 2 * b\<close> by (rule dvd_trans)
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2869
    then have \<open>2 dvd (1 + b * 2)\<close>
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2870
      by (simp add: ac_simps)
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2871
    then have \<open>is_unit (2 :: int)\<close>
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2872
      by simp
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2873
    then show ?case
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2874
      by simp
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2875
  next
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2876
    case euclidean_relation
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2877
    with that have \<open>a > 0\<close>
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2878
      by simp
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2879
    moreover have \<open>b mod a < a\<close>
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2880
      using \<open>a > 0\<close> by simp
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2881
    then have \<open>1 + 2 * (b mod a) < 2 * a\<close>
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2882
      by simp
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2883
    moreover have \<open>2 * (b mod a) + a * (2 * (b div a)) = 2 * (b div a * a + b mod a)\<close>
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2884
      by (simp only: algebra_simps)
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2885
    moreover have \<open>0 \<le> 2 * (b mod a)\<close>
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2886
      using \<open>a > 0\<close> by simp
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2887
    ultimately show ?case
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2888
      by (simp add: algebra_simps)
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2889
  qed
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2890
  then show ?Q and ?R
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2891
    by simp_all
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2892
qed
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2893
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2894
lemma neg_zdiv_mult_2: \<open>(1 + 2 * b) div (2 * a) = (b + 1) div a\<close> (is ?Q)
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2895
  and neg_zmod_mult_2: \<open>(1 + 2 * b) mod (2 * a) = 2 * ((b + 1) mod a) - 1\<close> (is ?R)
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2896
  if \<open>a \<le> 0\<close> for a b :: int
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2897
proof -
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2898
  have \<open>((1 + 2 * b) div (2 * a), (1 + 2 * b) mod (2 * a)) = ((b + 1) div a, 2 * ((b + 1) mod a) - 1)\<close>
76245
4111c94657b4 slightly less abusive proof pattern
haftmann
parents: 76231
diff changeset
  2899
  proof (induction rule: euclidean_relation_intI)
76224
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2900
    case by0
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2901
    then show ?case
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2902
      by simp
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2903
  next
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2904
    case divides
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2905
    have \<open>2 dvd (2 * a)\<close>
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2906
      by simp
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2907
    then have \<open>2 dvd (1 + 2 * b)\<close>
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2908
      using \<open>2 * a dvd 1 + 2 * b\<close> by (rule dvd_trans)
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2909
    then have \<open>2 dvd (1 + b * 2)\<close>
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2910
      by (simp add: ac_simps)
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2911
    then have \<open>is_unit (2 :: int)\<close>
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2912
      by simp
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2913
    then show ?case
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2914
      by simp
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2915
  next
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2916
    case euclidean_relation
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2917
    with that have \<open>a < 0\<close>
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2918
      by simp
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2919
    moreover have \<open>(b + 1) mod a > a\<close>
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2920
      using \<open>a < 0\<close> by simp
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2921
    then have \<open>2 * ((b + 1) mod a) > 1 + 2 * a\<close>
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2922
      by simp
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2923
    moreover have \<open>((1 + b) mod a) \<le> 0\<close>
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2924
      using \<open>a < 0\<close> by simp
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2925
    then have \<open>2 * ((1 + b) mod a) \<le> 0\<close>
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2926
      by simp
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2927
    moreover have \<open>2 * ((1 + b) mod a) + a * (2 * ((1 + b) div a)) =
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2928
      2 * ((1 + b) div a * a + (1 + b) mod a)\<close>
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2929
      by (simp only: algebra_simps)
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2930
    ultimately show ?case
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2931
      by (simp add: algebra_simps sgn_mult abs_mult)
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2932
  qed
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2933
  then show ?Q and ?R
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2934
    by simp_all
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2935
qed
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2936
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2937
lemma zdiv_numeral_Bit0 [simp]:
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2938
  \<open>numeral (Num.Bit0 v) div numeral (Num.Bit0 w) =
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2939
    numeral v div (numeral w :: int)\<close>
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2940
  unfolding numeral.simps unfolding mult_2 [symmetric]
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2941
  by (rule div_mult_mult1) simp
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2942
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2943
lemma zdiv_numeral_Bit1 [simp]:
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2944
  \<open>numeral (Num.Bit1 v) div numeral (Num.Bit0 w) =
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2945
    (numeral v div (numeral w :: int))\<close>
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2946
  unfolding numeral.simps
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2947
  unfolding mult_2 [symmetric] add.commute [of _ 1]
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2948
  by (rule pos_zdiv_mult_2) simp
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2949
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2950
lemma zmod_numeral_Bit0 [simp]:
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2951
  \<open>numeral (Num.Bit0 v) mod numeral (Num.Bit0 w) =
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2952
    (2::int) * (numeral v mod numeral w)\<close>
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2953
  unfolding numeral_Bit0 [of v] numeral_Bit0 [of w]
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2954
  unfolding mult_2 [symmetric] by (rule mod_mult_mult1)
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2955
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2956
lemma zmod_numeral_Bit1 [simp]:
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2957
  \<open>numeral (Num.Bit1 v) mod numeral (Num.Bit0 w) =
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2958
    2 * (numeral v mod numeral w) + (1::int)\<close>
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2959
  unfolding numeral_Bit1 [of v] numeral_Bit0 [of w]
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2960
  unfolding mult_2 [symmetric] add.commute [of _ 1]
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2961
  by (rule pos_zmod_mult_2) simp
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2962
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  2963
75937
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  2964
subsection \<open>Generic symbolic computations\<close>
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  2965
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  2966
text \<open>
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  2967
  The following type class contains everything necessary to formulate
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  2968
  a division algorithm in ring structures with numerals, restricted
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  2969
  to its positive segments.
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  2970
\<close>
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  2971
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  2972
class unique_euclidean_semiring_with_nat_division = unique_euclidean_semiring_with_nat +
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  2973
  fixes divmod :: \<open>num \<Rightarrow> num \<Rightarrow> 'a \<times> 'a\<close>
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  2974
    and divmod_step :: \<open>'a \<Rightarrow> 'a \<times> 'a \<Rightarrow> 'a \<times> 'a\<close> \<comment> \<open>
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  2975
      These are conceptually definitions but force generated code
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  2976
      to be monomorphic wrt. particular instances of this class which
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  2977
      yields a significant speedup.\<close>
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  2978
  assumes divmod_def: \<open>divmod m n = (numeral m div numeral n, numeral m mod numeral n)\<close>
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  2979
    and divmod_step_def [simp]: \<open>divmod_step l (q, r) =
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  2980
      (if euclidean_size l \<le> euclidean_size r then (2 * q + 1, r - l)
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  2981
       else (2 * q, r))\<close> \<comment> \<open>
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  2982
         This is a formulation of one step (referring to one digit position)
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  2983
         in school-method division: compare the dividend at the current
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  2984
         digit position with the remainder from previous division steps
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  2985
         and evaluate accordingly.\<close>
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  2986
begin
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  2987
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  2988
lemma fst_divmod:
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  2989
  \<open>fst (divmod m n) = numeral m div numeral n\<close>
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  2990
  by (simp add: divmod_def)
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  2991
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  2992
lemma snd_divmod:
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  2993
  \<open>snd (divmod m n) = numeral m mod numeral n\<close>
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  2994
  by (simp add: divmod_def)
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  2995
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  2996
text \<open>
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  2997
  Following a formulation of school-method division.
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  2998
  If the divisor is smaller than the dividend, terminate.
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  2999
  If not, shift the dividend to the right until termination
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3000
  occurs and then reiterate single division steps in the
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3001
  opposite direction.
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3002
\<close>
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3003
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3004
lemma divmod_divmod_step:
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3005
  \<open>divmod m n = (if m < n then (0, numeral m)
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3006
    else divmod_step (numeral n) (divmod m (Num.Bit0 n)))\<close>
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3007
proof (cases \<open>m < n\<close>)
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3008
  case True
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3009
  then show ?thesis
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3010
    by (simp add: prod_eq_iff fst_divmod snd_divmod flip: of_nat_numeral of_nat_div of_nat_mod)
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3011
next
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3012
  case False
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3013
  define r s t where \<open>r = (numeral m :: nat)\<close> \<open>s = (numeral n :: nat)\<close> \<open>t = 2 * s\<close>
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3014
  then have *: \<open>numeral m = of_nat r\<close> \<open>numeral n = of_nat s\<close> \<open>numeral (num.Bit0 n) = of_nat t\<close>
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3015
    and \<open>\<not> s \<le> r mod s\<close>
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3016
    by (simp_all add: not_le)
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3017
  have t: \<open>2 * (r div t) = r div s - r div s mod 2\<close>
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3018
    \<open>r mod t = s * (r div s mod 2) + r mod s\<close>
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3019
    by (simp add: Rings.minus_mod_eq_mult_div Groups.mult.commute [of 2] Euclidean_Division.div_mult2_eq \<open>t = 2 * s\<close>)
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3020
      (use mod_mult2_eq [of r s 2] in \<open>simp add: ac_simps \<open>t = 2 * s\<close>\<close>)
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3021
  have rs: \<open>r div s mod 2 = 0 \<or> r div s mod 2 = Suc 0\<close>
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3022
    by auto
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3023
  from \<open>\<not> s \<le> r mod s\<close> have \<open>s \<le> r mod t \<Longrightarrow>
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3024
     r div s = Suc (2 * (r div t)) \<and>
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3025
     r mod s = r mod t - s\<close>
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3026
    using rs
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3027
    by (auto simp add: t)
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3028
  moreover have \<open>r mod t < s \<Longrightarrow>
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3029
     r div s = 2 * (r div t) \<and>
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3030
     r mod s = r mod t\<close>
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3031
    using rs
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3032
    by (auto simp add: t)
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3033
  ultimately show ?thesis
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3034
    by (simp add: divmod_def prod_eq_iff split_def Let_def
76029
5f94db3a7e25 eliminated tabs, assuming tab-width=4;
wenzelm
parents: 75937
diff changeset
  3035
        not_less mod_eq_0_iff_dvd Rings.mod_eq_0_iff_dvd False not_le *)
75937
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3036
    (simp add: flip: of_nat_numeral of_nat_mult add.commute [of 1] of_nat_div of_nat_mod of_nat_Suc of_nat_diff)
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3037
qed
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3038
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3039
text \<open>The division rewrite proper -- first, trivial results involving \<open>1\<close>\<close>
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3040
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3041
lemma divmod_trivial [simp]:
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3042
  "divmod m Num.One = (numeral m, 0)"
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3043
  "divmod num.One (num.Bit0 n) = (0, Numeral1)"
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3044
  "divmod num.One (num.Bit1 n) = (0, Numeral1)"
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3045
  using divmod_divmod_step [of "Num.One"] by (simp_all add: divmod_def)
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3046
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3047
text \<open>Division by an even number is a right-shift\<close>
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3048
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3049
lemma divmod_cancel [simp]:
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3050
  \<open>divmod (Num.Bit0 m) (Num.Bit0 n) = (case divmod m n of (q, r) \<Rightarrow> (q, 2 * r))\<close> (is ?P)
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3051
  \<open>divmod (Num.Bit1 m) (Num.Bit0 n) = (case divmod m n of (q, r) \<Rightarrow> (q, 2 * r + 1))\<close> (is ?Q)
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3052
proof -
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3053
  define r s where \<open>r = (numeral m :: nat)\<close> \<open>s = (numeral n :: nat)\<close>
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3054
  then have *: \<open>numeral m = of_nat r\<close> \<open>numeral n = of_nat s\<close>
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3055
    \<open>numeral (num.Bit0 m) = of_nat (2 * r)\<close> \<open>numeral (num.Bit0 n) = of_nat (2 * s)\<close>
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3056
    \<open>numeral (num.Bit1 m) = of_nat (Suc (2 * r))\<close>
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3057
    by simp_all
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3058
  have **: \<open>Suc (2 * r) div 2 = r\<close>
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3059
    by simp
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3060
  show ?P and ?Q
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3061
    by (simp_all add: divmod_def *)
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3062
      (simp_all flip: of_nat_numeral of_nat_div of_nat_mod of_nat_mult add.commute [of 1] of_nat_Suc
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3063
       add: Euclidean_Division.mod_mult_mult1 div_mult2_eq [of _ 2] mod_mult2_eq [of _ 2] **)
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3064
qed
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3065
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3066
text \<open>The really hard work\<close>
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3067
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3068
lemma divmod_steps [simp]:
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3069
  "divmod (num.Bit0 m) (num.Bit1 n) =
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3070
      (if m \<le> n then (0, numeral (num.Bit0 m))
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3071
       else divmod_step (numeral (num.Bit1 n))
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3072
             (divmod (num.Bit0 m)
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3073
               (num.Bit0 (num.Bit1 n))))"
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3074
  "divmod (num.Bit1 m) (num.Bit1 n) =
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3075
      (if m < n then (0, numeral (num.Bit1 m))
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3076
       else divmod_step (numeral (num.Bit1 n))
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3077
             (divmod (num.Bit1 m)
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3078
               (num.Bit0 (num.Bit1 n))))"
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3079
  by (simp_all add: divmod_divmod_step)
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3080
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3081
lemmas divmod_algorithm_code = divmod_trivial divmod_cancel divmod_steps
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3082
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3083
text \<open>Special case: divisibility\<close>
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3084
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3085
definition divides_aux :: "'a \<times> 'a \<Rightarrow> bool"
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3086
where
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3087
  "divides_aux qr \<longleftrightarrow> snd qr = 0"
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3088
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3089
lemma divides_aux_eq [simp]:
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3090
  "divides_aux (q, r) \<longleftrightarrow> r = 0"
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3091
  by (simp add: divides_aux_def)
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3092
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3093
lemma dvd_numeral_simp [simp]:
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3094
  "numeral m dvd numeral n \<longleftrightarrow> divides_aux (divmod n m)"
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3095
  by (simp add: divmod_def mod_eq_0_iff_dvd)
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3096
76224
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  3097
text \<open>Generic computation of quotient and remainder\<close>
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  3098
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  3099
lemma numeral_div_numeral [simp]:
75937
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3100
  "numeral k div numeral l = fst (divmod k l)"
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3101
  by (simp add: fst_divmod)
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3102
76224
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  3103
lemma numeral_mod_numeral [simp]:
75937
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3104
  "numeral k mod numeral l = snd (divmod k l)"
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3105
  by (simp add: snd_divmod)
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3106
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3107
lemma one_div_numeral [simp]:
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3108
  "1 div numeral n = fst (divmod num.One n)"
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3109
  by (simp add: fst_divmod)
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3110
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3111
lemma one_mod_numeral [simp]:
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3112
  "1 mod numeral n = snd (divmod num.One n)"
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3113
  by (simp add: snd_divmod)
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3114
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3115
end
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3116
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3117
instantiation nat :: unique_euclidean_semiring_with_nat_division
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3118
begin
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3119
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3120
definition divmod_nat :: "num \<Rightarrow> num \<Rightarrow> nat \<times> nat"
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3121
where
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3122
  divmod'_nat_def: "divmod_nat m n = (numeral m div numeral n, numeral m mod numeral n)"
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3123
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3124
definition divmod_step_nat :: "nat \<Rightarrow> nat \<times> nat \<Rightarrow> nat \<times> nat"
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3125
where
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3126
  "divmod_step_nat l qr = (let (q, r) = qr
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3127
    in if r \<ge> l then (2 * q + 1, r - l)
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3128
    else (2 * q, r))"
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3129
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3130
instance
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3131
  by standard (simp_all add: divmod'_nat_def divmod_step_nat_def)
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3132
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3133
end
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3134
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3135
declare divmod_algorithm_code [where ?'a = nat, code]
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3136
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3137
lemma Suc_0_div_numeral [simp]:
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3138
  \<open>Suc 0 div numeral Num.One = 1\<close>
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3139
  \<open>Suc 0 div numeral (Num.Bit0 n) = 0\<close>
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3140
  \<open>Suc 0 div numeral (Num.Bit1 n) = 0\<close>
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3141
  by simp_all
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3142
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3143
lemma Suc_0_mod_numeral [simp]:
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3144
  \<open>Suc 0 mod numeral Num.One = 0\<close>
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3145
  \<open>Suc 0 mod numeral (Num.Bit0 n) = 1\<close>
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3146
  \<open>Suc 0 mod numeral (Num.Bit1 n) = 1\<close>
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3147
  by simp_all
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3148
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3149
instantiation int :: unique_euclidean_semiring_with_nat_division
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3150
begin
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3151
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3152
definition divmod_int :: "num \<Rightarrow> num \<Rightarrow> int \<times> int"
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3153
where
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3154
  "divmod_int m n = (numeral m div numeral n, numeral m mod numeral n)"
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3155
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3156
definition divmod_step_int :: "int \<Rightarrow> int \<times> int \<Rightarrow> int \<times> int"
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3157
where
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3158
  "divmod_step_int l qr = (let (q, r) = qr
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3159
    in if \<bar>l\<bar> \<le> \<bar>r\<bar> then (2 * q + 1, r - l)
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3160
    else (2 * q, r))"
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3161
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3162
instance
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3163
  by standard (auto simp add: divmod_int_def divmod_step_int_def)
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3164
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3165
end
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3166
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3167
declare divmod_algorithm_code [where ?'a = int, code]
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3168
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3169
context
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3170
begin
76224
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  3171
75937
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3172
qualified definition adjust_div :: "int \<times> int \<Rightarrow> int"
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3173
where
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3174
  "adjust_div qr = (let (q, r) = qr in q + of_bool (r \<noteq> 0))"
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3175
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3176
qualified lemma adjust_div_eq [simp, code]:
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3177
  "adjust_div (q, r) = q + of_bool (r \<noteq> 0)"
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3178
  by (simp add: adjust_div_def)
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3179
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3180
qualified definition adjust_mod :: "num \<Rightarrow> int \<Rightarrow> int"
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3181
where
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3182
  [simp]: "adjust_mod l r = (if r = 0 then 0 else numeral l - r)"
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3183
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3184
lemma minus_numeral_div_numeral [simp]:
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3185
  "- numeral m div numeral n = - (adjust_div (divmod m n) :: int)"
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3186
proof -
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3187
  have "int (fst (divmod m n)) = fst (divmod m n)"
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3188
    by (simp only: fst_divmod divide_int_def) auto
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3189
  then show ?thesis
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3190
    by (auto simp add: split_def Let_def adjust_div_def divides_aux_def divide_int_def)
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3191
qed
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3192
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3193
lemma minus_numeral_mod_numeral [simp]:
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3194
  "- numeral m mod numeral n = adjust_mod n (snd (divmod m n) :: int)"
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3195
proof (cases "snd (divmod m n) = (0::int)")
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3196
  case True
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3197
  then show ?thesis
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3198
    by (simp add: mod_eq_0_iff_dvd divides_aux_def)
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3199
next
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3200
  case False
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3201
  then have "int (snd (divmod m n)) = snd (divmod m n)" if "snd (divmod m n) \<noteq> (0::int)"
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3202
    by (simp only: snd_divmod modulo_int_def) auto
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3203
  then show ?thesis
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3204
    by (simp add: divides_aux_def adjust_div_def)
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3205
      (simp add: divides_aux_def modulo_int_def)
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3206
qed
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3207
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3208
lemma numeral_div_minus_numeral [simp]:
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3209
  "numeral m div - numeral n = - (adjust_div (divmod m n) :: int)"
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3210
proof -
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3211
  have "int (fst (divmod m n)) = fst (divmod m n)"
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3212
    by (simp only: fst_divmod divide_int_def) auto
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3213
  then show ?thesis
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3214
    by (auto simp add: split_def Let_def adjust_div_def divides_aux_def divide_int_def)
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3215
qed
76224
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  3216
75937
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3217
lemma numeral_mod_minus_numeral [simp]:
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3218
  "numeral m mod - numeral n = - adjust_mod n (snd (divmod m n) :: int)"
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3219
proof (cases "snd (divmod m n) = (0::int)")
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3220
  case True
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3221
  then show ?thesis
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3222
    by (simp add: mod_eq_0_iff_dvd divides_aux_def)
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3223
next
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3224
  case False
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3225
  then have "int (snd (divmod m n)) = snd (divmod m n)" if "snd (divmod m n) \<noteq> (0::int)"
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3226
    by (simp only: snd_divmod modulo_int_def) auto
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3227
  then show ?thesis
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3228
    by (simp add: divides_aux_def adjust_div_def)
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3229
      (simp add: divides_aux_def modulo_int_def)
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3230
qed
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3231
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3232
lemma minus_one_div_numeral [simp]:
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3233
  "- 1 div numeral n = - (adjust_div (divmod Num.One n) :: int)"
76224
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  3234
  using minus_numeral_div_numeral [of Num.One n] by simp
75937
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3235
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3236
lemma minus_one_mod_numeral [simp]:
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3237
  "- 1 mod numeral n = adjust_mod n (snd (divmod Num.One n) :: int)"
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3238
  using minus_numeral_mod_numeral [of Num.One n] by simp
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3239
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3240
lemma one_div_minus_numeral [simp]:
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3241
  "1 div - numeral n = - (adjust_div (divmod Num.One n) :: int)"
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3242
  using numeral_div_minus_numeral [of Num.One n] by simp
76224
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  3243
75937
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3244
lemma one_mod_minus_numeral [simp]:
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3245
  "1 mod - numeral n = - adjust_mod n (snd (divmod Num.One n) :: int)"
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3246
  using numeral_mod_minus_numeral [of Num.One n] by simp
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3247
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3248
lemma [code]:
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3249
  fixes k :: int
76224
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  3250
  shows
75937
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3251
    "k div 0 = 0"
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3252
    "k mod 0 = k"
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3253
    "0 div k = 0"
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3254
    "0 mod k = 0"
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3255
    "k div Int.Pos Num.One = k"
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3256
    "k mod Int.Pos Num.One = 0"
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3257
    "k div Int.Neg Num.One = - k"
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3258
    "k mod Int.Neg Num.One = 0"
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3259
    "Int.Pos m div Int.Pos n = (fst (divmod m n) :: int)"
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3260
    "Int.Pos m mod Int.Pos n = (snd (divmod m n) :: int)"
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3261
    "Int.Neg m div Int.Pos n = - (adjust_div (divmod m n) :: int)"
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3262
    "Int.Neg m mod Int.Pos n = adjust_mod n (snd (divmod m n) :: int)"
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3263
    "Int.Pos m div Int.Neg n = - (adjust_div (divmod m n) :: int)"
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3264
    "Int.Pos m mod Int.Neg n = - adjust_mod n (snd (divmod m n) :: int)"
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3265
    "Int.Neg m div Int.Neg n = (fst (divmod m n) :: int)"
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3266
    "Int.Neg m mod Int.Neg n = - (snd (divmod m n) :: int)"
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3267
  by simp_all
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3268
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3269
end
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3270
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3271
lemma divmod_BitM_2_eq [simp]:
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3272
  \<open>divmod (Num.BitM m) (Num.Bit0 Num.One) = (numeral m - 1, (1 :: int))\<close>
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3273
  by (cases m) simp_all
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3274
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3275
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3276
subsubsection \<open>Computation by simplification\<close>
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3277
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3278
lemma euclidean_size_nat_less_eq_iff:
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3279
  \<open>euclidean_size m \<le> euclidean_size n \<longleftrightarrow> m \<le> n\<close> for m n :: nat
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3280
  by simp
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3281
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3282
lemma euclidean_size_int_less_eq_iff:
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3283
  \<open>euclidean_size k \<le> euclidean_size l \<longleftrightarrow> \<bar>k\<bar> \<le> \<bar>l\<bar>\<close> for k l :: int
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3284
  by auto
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3285
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3286
simproc_setup numeral_divmod
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3287
  ("0 div 0 :: 'a :: unique_euclidean_semiring_with_nat_division" | "0 mod 0 :: 'a :: unique_euclidean_semiring_with_nat_division" |
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3288
   "0 div 1 :: 'a :: unique_euclidean_semiring_with_nat_division" | "0 mod 1 :: 'a :: unique_euclidean_semiring_with_nat_division" |
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3289
   "0 div - 1 :: int" | "0 mod - 1 :: int" |
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3290
   "0 div numeral b :: 'a :: unique_euclidean_semiring_with_nat_division" | "0 mod numeral b :: 'a :: unique_euclidean_semiring_with_nat_division" |
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3291
   "0 div - numeral b :: int" | "0 mod - numeral b :: int" |
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3292
   "1 div 0 :: 'a :: unique_euclidean_semiring_with_nat_division" | "1 mod 0 :: 'a :: unique_euclidean_semiring_with_nat_division" |
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3293
   "1 div 1 :: 'a :: unique_euclidean_semiring_with_nat_division" | "1 mod 1 :: 'a :: unique_euclidean_semiring_with_nat_division" |
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3294
   "1 div - 1 :: int" | "1 mod - 1 :: int" |
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3295
   "1 div numeral b :: 'a :: unique_euclidean_semiring_with_nat_division" | "1 mod numeral b :: 'a :: unique_euclidean_semiring_with_nat_division" |
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3296
   "1 div - numeral b :: int" |"1 mod - numeral b :: int" |
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3297
   "- 1 div 0 :: int" | "- 1 mod 0 :: int" | "- 1 div 1 :: int" | "- 1 mod 1 :: int" |
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3298
   "- 1 div - 1 :: int" | "- 1 mod - 1 :: int" | "- 1 div numeral b :: int" | "- 1 mod numeral b :: int" |
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3299
   "- 1 div - numeral b :: int" | "- 1 mod - numeral b :: int" |
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3300
   "numeral a div 0 :: 'a :: unique_euclidean_semiring_with_nat_division" | "numeral a mod 0 :: 'a :: unique_euclidean_semiring_with_nat_division" |
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3301
   "numeral a div 1 :: 'a :: unique_euclidean_semiring_with_nat_division" | "numeral a mod 1 :: 'a :: unique_euclidean_semiring_with_nat_division" |
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3302
   "numeral a div - 1 :: int" | "numeral a mod - 1 :: int" |
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3303
   "numeral a div numeral b :: 'a :: unique_euclidean_semiring_with_nat_division" | "numeral a mod numeral b :: 'a :: unique_euclidean_semiring_with_nat_division" |
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3304
   "numeral a div - numeral b :: int" | "numeral a mod - numeral b :: int" |
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3305
   "- numeral a div 0 :: int" | "- numeral a mod 0 :: int" |
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3306
   "- numeral a div 1 :: int" | "- numeral a mod 1 :: int" |
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3307
   "- numeral a div - 1 :: int" | "- numeral a mod - 1 :: int" |
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3308
   "- numeral a div numeral b :: int" | "- numeral a mod numeral b :: int" |
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3309
   "- numeral a div - numeral b :: int" | "- numeral a mod - numeral b :: int") = \<open>
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3310
  let
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3311
    val if_cong = the (Code.get_case_cong \<^theory> \<^const_name>\<open>If\<close>);
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3312
    fun successful_rewrite ctxt ct =
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3313
      let
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3314
        val thm = Simplifier.rewrite ctxt ct
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3315
      in if Thm.is_reflexive thm then NONE else SOME thm end;
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3316
  in fn phi =>
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3317
    let
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3318
      val simps = Morphism.fact phi (@{thms div_0 mod_0 div_by_0 mod_by_0 div_by_1 mod_by_1
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3319
        one_div_numeral one_mod_numeral minus_one_div_numeral minus_one_mod_numeral
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3320
        one_div_minus_numeral one_mod_minus_numeral
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3321
        numeral_div_numeral numeral_mod_numeral minus_numeral_div_numeral minus_numeral_mod_numeral
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3322
        numeral_div_minus_numeral numeral_mod_minus_numeral
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3323
        div_minus_minus mod_minus_minus Euclidean_Division.adjust_div_eq of_bool_eq one_neq_zero
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3324
        numeral_neq_zero neg_equal_0_iff_equal arith_simps arith_special divmod_trivial
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3325
        divmod_cancel divmod_steps divmod_step_def fst_conv snd_conv numeral_One
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3326
        case_prod_beta rel_simps Euclidean_Division.adjust_mod_def div_minus1_right mod_minus1_right
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3327
        minus_minus numeral_times_numeral mult_zero_right mult_1_right
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3328
        euclidean_size_nat_less_eq_iff euclidean_size_int_less_eq_iff diff_nat_numeral nat_numeral}
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3329
        @ [@{lemma "0 = 0 \<longleftrightarrow> True" by simp}]);
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3330
      fun prepare_simpset ctxt = HOL_ss |> Simplifier.simpset_map ctxt
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3331
        (Simplifier.add_cong if_cong #> fold Simplifier.add_simp simps)
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3332
    in fn ctxt => successful_rewrite (Simplifier.put_simpset (prepare_simpset ctxt) ctxt) end
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3333
  end
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3334
\<close> \<comment> \<open>
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3335
  There is space for improvement here: the calculation itself
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3336
  could be carried out outside the logic, and a generic simproc
76224
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 76141
diff changeset
  3337
  (simplifier setup) for generic calculation would be helpful.
75937
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3338
\<close>
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3339
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3340
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3341
subsubsection \<open>Code generation\<close>
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3342
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3343
context
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3344
begin
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3345
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3346
qualified definition divmod_nat :: "nat \<Rightarrow> nat \<Rightarrow> nat \<times> nat"
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3347
  where "divmod_nat m n = (m div n, m mod n)"
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3348
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3349
qualified lemma divmod_nat_if [code]:
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3350
  "divmod_nat m n = (if n = 0 \<or> m < n then (0, m) else
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3351
    let (q, r) = divmod_nat (m - n) n in (Suc q, r))"
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3352
  by (simp add: divmod_nat_def prod_eq_iff case_prod_beta not_less le_div_geq le_mod_geq)
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3353
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3354
qualified lemma [code]:
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3355
  "m div n = fst (divmod_nat m n)"
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3356
  "m mod n = snd (divmod_nat m n)"
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3357
  by (simp_all add: divmod_nat_def)
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3358
02b18f59f903 streamlined
haftmann
parents: 75881
diff changeset
  3359
end
66808
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  3360
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  3361
code_identifier
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  3362
  code_module Euclidean_Division \<rightharpoonup> (SML) Arith and (OCaml) Arith and (Haskell) Arith
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  3363
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  3364
end