src/HOL/Euclidean_Division.thy
author paulson <lp15@cam.ac.uk>
Mon, 30 Nov 2020 11:06:01 +0000
changeset 72794 3757e64e75bb
parent 72187 e4aecb0c7296
child 73535 0f33c7031ec9
permissions -rw-r--r--
tweaked
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>
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
    13
  
66817
0b12755ccbb2 euclidean rings need no normalization
haftmann
parents: 66816
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"
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
    17
  assumes mod_size_less: 
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"
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
    49
    and "b dvd a" 
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
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
    86
lemma unit_iff_euclidean_size: 
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)"
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
    99
  with size_mult_mono'[OF assms(1), of b] 
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"
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   102
    by (rule dvd_euclidean_size_eq_imp_dvd [OF _ eq]) (insert assms, simp_all)
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   103
  hence "a * b dvd 1 * b" by simp
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   104
  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
   105
  with assms(3) show False by contradiction
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   106
qed
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   107
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   108
lemma dvd_imp_size_le:
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   109
  assumes "a dvd b" "b \<noteq> 0" 
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   110
  shows   "euclidean_size a \<le> euclidean_size b"
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   111
  using assms by (auto elim!: dvdE simp: size_mult_mono)
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   112
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   113
lemma dvd_proper_imp_size_less:
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   114
  assumes "a dvd b" "\<not> b dvd a" "b \<noteq> 0" 
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   115
  shows   "euclidean_size a < euclidean_size b"
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   116
proof -
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   117
  from assms(1) obtain c where "b = a * c" by (erule dvdE)
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   118
  hence z: "b = c * a" by (simp add: mult.commute)
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   119
  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
   120
  with z assms show ?thesis
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   121
    by (auto intro!: euclidean_size_times_nonunit)
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   122
qed
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   123
66798
39bb2462e681 fundamental property of division by units
haftmann
parents: 64785
diff changeset
   124
lemma unit_imp_mod_eq_0:
39bb2462e681 fundamental property of division by units
haftmann
parents: 64785
diff changeset
   125
  "a mod b = 0" if "is_unit b"
39bb2462e681 fundamental property of division by units
haftmann
parents: 64785
diff changeset
   126
  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
   127
69695
753ae9e9773d algebraized more material from theory Divides
haftmann
parents: 69593
diff changeset
   128
lemma mod_eq_self_iff_div_eq_0:
753ae9e9773d algebraized more material from theory Divides
haftmann
parents: 69593
diff changeset
   129
  "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
   130
proof
753ae9e9773d algebraized more material from theory Divides
haftmann
parents: 69593
diff changeset
   131
  assume ?P
753ae9e9773d algebraized more material from theory Divides
haftmann
parents: 69593
diff changeset
   132
  with div_mult_mod_eq [of a b] show ?Q
753ae9e9773d algebraized more material from theory Divides
haftmann
parents: 69593
diff changeset
   133
    by auto
753ae9e9773d algebraized more material from theory Divides
haftmann
parents: 69593
diff changeset
   134
next
753ae9e9773d algebraized more material from theory Divides
haftmann
parents: 69593
diff changeset
   135
  assume ?Q
753ae9e9773d algebraized more material from theory Divides
haftmann
parents: 69593
diff changeset
   136
  with div_mult_mod_eq [of a b] show ?P
753ae9e9773d algebraized more material from theory Divides
haftmann
parents: 69593
diff changeset
   137
    by simp
753ae9e9773d algebraized more material from theory Divides
haftmann
parents: 69593
diff changeset
   138
qed
753ae9e9773d algebraized more material from theory Divides
haftmann
parents: 69593
diff changeset
   139
67051
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
   140
lemma coprime_mod_left_iff [simp]:
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
   141
  "coprime (a mod b) b \<longleftrightarrow> coprime a b" if "b \<noteq> 0"
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
   142
  by (rule; rule coprimeI)
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
   143
    (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
   144
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
   145
lemma coprime_mod_right_iff [simp]:
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
   146
  "coprime a (b mod a) \<longleftrightarrow> coprime a b" if "a \<noteq> 0"
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
   147
  using that coprime_mod_left_iff [of a b] by (simp add: ac_simps)
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
   148
64785
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   149
end
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   150
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   151
class euclidean_ring = idom_modulo + euclidean_semiring
66886
960509bfd47e added lemmas and tuned proofs
haftmann
parents: 66840
diff changeset
   152
begin
960509bfd47e added lemmas and tuned proofs
haftmann
parents: 66840
diff changeset
   153
67087
733017b19de9 generalized more lemmas
haftmann
parents: 67083
diff changeset
   154
lemma dvd_diff_commute [ac_simps]:
66886
960509bfd47e added lemmas and tuned proofs
haftmann
parents: 66840
diff changeset
   155
  "a dvd c - b \<longleftrightarrow> a dvd b - c"
960509bfd47e added lemmas and tuned proofs
haftmann
parents: 66840
diff changeset
   156
proof -
960509bfd47e added lemmas and tuned proofs
haftmann
parents: 66840
diff changeset
   157
  have "a dvd c - b \<longleftrightarrow> a dvd (c - b) * - 1"
960509bfd47e added lemmas and tuned proofs
haftmann
parents: 66840
diff changeset
   158
    by (subst dvd_mult_unit_iff) simp_all
960509bfd47e added lemmas and tuned proofs
haftmann
parents: 66840
diff changeset
   159
  then show ?thesis
960509bfd47e added lemmas and tuned proofs
haftmann
parents: 66840
diff changeset
   160
    by simp
960509bfd47e added lemmas and tuned proofs
haftmann
parents: 66840
diff changeset
   161
qed
960509bfd47e added lemmas and tuned proofs
haftmann
parents: 66840
diff changeset
   162
 
960509bfd47e added lemmas and tuned proofs
haftmann
parents: 66840
diff changeset
   163
end
64785
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   164
66840
0d689d71dbdc canonical multiplicative euclidean size
haftmann
parents: 66839
diff changeset
   165
66806
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   166
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
   167
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
   168
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
   169
  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
   170
  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
   171
begin
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   172
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   173
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
   174
  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
   175
  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
   176
  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
   177
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   178
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
   179
  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
   180
  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
   181
  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
   182
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   183
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
   184
  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
   185
  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
   186
  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
   187
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   188
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
   189
proof (cases "b = 0")
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   190
  case True then show ?thesis by simp
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   191
next
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   192
  case False
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   193
  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
   194
    by (simp add: div_mult_mod_eq)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   195
  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
   196
    "\<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
   197
      by (simp add: algebra_simps)
66806
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   198
  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
   199
    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
   200
  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
   201
    by (simp add: div_mult_mod_eq)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   202
  then show ?thesis by simp
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   203
qed
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   204
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   205
lemma mod_mult_self2 [simp]:
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   206
  "(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
   207
  by (simp add: mult.commute [of b])
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   208
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   209
lemma mod_mult_self3 [simp]:
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   210
  "(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
   211
  by (simp add: add.commute)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   212
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   213
lemma mod_mult_self4 [simp]:
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   214
  "(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
   215
  by (simp add: add.commute)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   216
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   217
lemma mod_mult_self1_is_0 [simp]:
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   218
  "b * a mod b = 0"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   219
  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
   220
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   221
lemma mod_mult_self2_is_0 [simp]:
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   222
  "a * b mod b = 0"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   223
  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
   224
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
   225
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
   226
  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
   227
  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
   228
  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
   229
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
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
   231
  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
   232
  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
   233
  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
   234
66806
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   235
lemma mod_add_self1 [simp]:
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   236
  "(b + a) mod b = a mod b"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   237
  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
   238
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   239
lemma mod_add_self2 [simp]:
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   240
  "(a + b) mod b = a mod b"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   241
  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
   242
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   243
lemma mod_div_trivial [simp]:
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   244
  "a mod b div b = 0"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   245
proof (cases "b = 0")
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   246
  assume "b = 0"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   247
  thus ?thesis by simp
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   248
next
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   249
  assume "b \<noteq> 0"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   250
  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
   251
    by (rule div_mult_self1 [symmetric])
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   252
  also have "\<dots> = a div b"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   253
    by (simp only: mod_div_mult_eq)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   254
  also have "\<dots> = a div b + 0"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   255
    by simp
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   256
  finally show ?thesis
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   257
    by (rule add_left_imp_eq)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   258
qed
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   259
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   260
lemma mod_mod_trivial [simp]:
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   261
  "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
   262
proof -
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   263
  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
   264
    by (simp only: mod_mult_self1)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   265
  also have "\<dots> = a mod b"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   266
    by (simp only: mod_div_mult_eq)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   267
  finally show ?thesis .
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   268
qed
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   269
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   270
lemma mod_mod_cancel:
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   271
  assumes "c dvd b"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   272
  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
   273
proof -
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   274
  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
   275
    by (rule dvdE)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   276
  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
   277
    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
   278
  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
   279
    by (simp only: mod_mult_self1)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   280
  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
   281
    by (simp only: ac_simps)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   282
  also have "\<dots> = a mod c"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   283
    by (simp only: div_mult_mod_eq)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   284
  finally show ?thesis .
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   285
qed
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   286
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
   287
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
   288
  "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
   289
  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
   290
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
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
   292
  "(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
   293
  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
   294
66806
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   295
lemma mod_mult_mult1:
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   296
  "(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
   297
proof (cases "c = 0")
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   298
  case True then show ?thesis by simp
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   299
next
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   300
  case False
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   301
  from div_mult_mod_eq
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   302
  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
   303
  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
   304
    = 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
   305
  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
   306
qed
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   307
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   308
lemma mod_mult_mult2:
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   309
  "(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
   310
  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
   311
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   312
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
   313
  by (fact mod_mult_mult2 [symmetric])
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   314
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   315
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
   316
  by (fact mod_mult_mult1 [symmetric])
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   317
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   318
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
   319
  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
   320
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   321
lemma div_plus_div_distrib_dvd_left:
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   322
  "c dvd a \<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
   323
  by (cases "c = 0") (auto elim: dvdE)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   324
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   325
lemma div_plus_div_distrib_dvd_right:
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   326
  "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
   327
  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
   328
  by (simp add: ac_simps)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   329
71413
65ffe9e910d4 more specific class assumptions
haftmann
parents: 71412
diff changeset
   330
lemma sum_div_partition:
65ffe9e910d4 more specific class assumptions
haftmann
parents: 71412
diff changeset
   331
  \<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
   332
    if \<open>finite A\<close>
65ffe9e910d4 more specific class assumptions
haftmann
parents: 71412
diff changeset
   333
proof -
65ffe9e910d4 more specific class assumptions
haftmann
parents: 71412
diff changeset
   334
  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
   335
    by auto
65ffe9e910d4 more specific class assumptions
haftmann
parents: 71412
diff changeset
   336
  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
   337
    by simp
65ffe9e910d4 more specific class assumptions
haftmann
parents: 71412
diff changeset
   338
  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
   339
    using \<open>finite A\<close> by (auto intro: sum.union_inter_neutral)
65ffe9e910d4 more specific class assumptions
haftmann
parents: 71412
diff changeset
   340
  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
   341
  define B where B: \<open>B = A \<inter> {a. b dvd f a}\<close>
65ffe9e910d4 more specific class assumptions
haftmann
parents: 71412
diff changeset
   342
  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
   343
    by simp_all
65ffe9e910d4 more specific class assumptions
haftmann
parents: 71412
diff changeset
   344
  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
   345
    by induction (simp_all add: div_plus_div_distrib_dvd_left)
65ffe9e910d4 more specific class assumptions
haftmann
parents: 71412
diff changeset
   346
  then show ?thesis using *
65ffe9e910d4 more specific class assumptions
haftmann
parents: 71412
diff changeset
   347
    by (simp add: B div_plus_div_distrib_dvd_left)
65ffe9e910d4 more specific class assumptions
haftmann
parents: 71412
diff changeset
   348
qed
65ffe9e910d4 more specific class assumptions
haftmann
parents: 71412
diff changeset
   349
66806
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   350
named_theorems mod_simps
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   351
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   352
text \<open>Addition respects modular equivalence.\<close>
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   353
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   354
lemma mod_add_left_eq [mod_simps]:
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   355
  "(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
   356
proof -
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   357
  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
   358
    by (simp only: div_mult_mod_eq)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   359
  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
   360
    by (simp only: ac_simps)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   361
  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
   362
    by (rule mod_mult_self1)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   363
  finally show ?thesis
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   364
    by (rule sym)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   365
qed
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   366
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   367
lemma mod_add_right_eq [mod_simps]:
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   368
  "(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
   369
  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
   370
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   371
lemma mod_add_eq:
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   372
  "(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
   373
  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
   374
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   375
lemma mod_sum_eq [mod_simps]:
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   376
  "(\<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
   377
proof (induct A rule: infinite_finite_induct)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   378
  case (insert i A)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   379
  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
   380
    = (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
   381
    by simp
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   382
  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
   383
    by (simp add: mod_simps)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   384
  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
   385
    by (simp add: insert.hyps)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   386
  finally show ?case
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   387
    by (simp add: insert.hyps mod_simps)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   388
qed simp_all
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   389
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   390
lemma mod_add_cong:
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   391
  assumes "a mod c = a' mod c"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   392
  assumes "b mod c = b' mod c"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   393
  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
   394
proof -
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   395
  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
   396
    unfolding assms ..
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   397
  then show ?thesis
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   398
    by (simp add: mod_add_eq)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   399
qed
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   400
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   401
text \<open>Multiplication respects modular equivalence.\<close>
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   402
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   403
lemma mod_mult_left_eq [mod_simps]:
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   404
  "((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
   405
proof -
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   406
  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
   407
    by (simp only: div_mult_mod_eq)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   408
  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
   409
    by (simp only: algebra_simps)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   410
  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
   411
    by (rule mod_mult_self1)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   412
  finally show ?thesis
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   413
    by (rule sym)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   414
qed
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   415
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   416
lemma mod_mult_right_eq [mod_simps]:
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   417
  "(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
   418
  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
   419
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   420
lemma mod_mult_eq:
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   421
  "((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
   422
  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
   423
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   424
lemma mod_prod_eq [mod_simps]:
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   425
  "(\<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
   426
proof (induct A rule: infinite_finite_induct)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   427
  case (insert i A)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   428
  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
   429
    = (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
   430
    by simp
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   431
  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
   432
    by (simp add: mod_simps)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   433
  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
   434
    by (simp add: insert.hyps)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   435
  finally show ?case
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   436
    by (simp add: insert.hyps mod_simps)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   437
qed simp_all
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   438
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   439
lemma mod_mult_cong:
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   440
  assumes "a mod c = a' mod c"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   441
  assumes "b mod c = b' mod c"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   442
  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
   443
proof -
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   444
  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
   445
    unfolding assms ..
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   446
  then show ?thesis
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   447
    by (simp add: mod_mult_eq)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   448
qed
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   449
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   450
text \<open>Exponentiation respects modular equivalence.\<close>
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   451
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   452
lemma power_mod [mod_simps]: 
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   453
  "((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
   454
proof (induct n)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   455
  case 0
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   456
  then show ?case by simp
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   457
next
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   458
  case (Suc n)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   459
  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
   460
    by (simp add: mod_mult_right_eq)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   461
  with Suc show ?case
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   462
    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
   463
qed
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   464
71413
65ffe9e910d4 more specific class assumptions
haftmann
parents: 71412
diff changeset
   465
lemma power_diff_power_eq:
65ffe9e910d4 more specific class assumptions
haftmann
parents: 71412
diff changeset
   466
  \<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
   467
    if \<open>a \<noteq> 0\<close>
65ffe9e910d4 more specific class assumptions
haftmann
parents: 71412
diff changeset
   468
proof (cases \<open>n \<le> m\<close>)
65ffe9e910d4 more specific class assumptions
haftmann
parents: 71412
diff changeset
   469
  case True
65ffe9e910d4 more specific class assumptions
haftmann
parents: 71412
diff changeset
   470
  with that power_diff [symmetric, of a n m] show ?thesis by simp
65ffe9e910d4 more specific class assumptions
haftmann
parents: 71412
diff changeset
   471
next
65ffe9e910d4 more specific class assumptions
haftmann
parents: 71412
diff changeset
   472
  case False
65ffe9e910d4 more specific class assumptions
haftmann
parents: 71412
diff changeset
   473
  then obtain q where n: \<open>n = m + Suc q\<close>
65ffe9e910d4 more specific class assumptions
haftmann
parents: 71412
diff changeset
   474
    by (auto simp add: not_le dest: less_imp_Suc_add)
65ffe9e910d4 more specific class assumptions
haftmann
parents: 71412
diff changeset
   475
  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
   476
    by (simp add: power_add ac_simps)
65ffe9e910d4 more specific class assumptions
haftmann
parents: 71412
diff changeset
   477
  moreover from that have \<open>a ^ m \<noteq> 0\<close>
65ffe9e910d4 more specific class assumptions
haftmann
parents: 71412
diff changeset
   478
    by simp
65ffe9e910d4 more specific class assumptions
haftmann
parents: 71412
diff changeset
   479
  ultimately have \<open>a ^ m div a ^ n = 1 div a ^ Suc q\<close>
65ffe9e910d4 more specific class assumptions
haftmann
parents: 71412
diff changeset
   480
    by (subst (asm) div_mult_mult1) simp
65ffe9e910d4 more specific class assumptions
haftmann
parents: 71412
diff changeset
   481
  with False n show ?thesis
65ffe9e910d4 more specific class assumptions
haftmann
parents: 71412
diff changeset
   482
    by simp
65ffe9e910d4 more specific class assumptions
haftmann
parents: 71412
diff changeset
   483
qed
65ffe9e910d4 more specific class assumptions
haftmann
parents: 71412
diff changeset
   484
66806
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   485
end
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   486
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
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
   489
begin
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   490
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
   491
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
   492
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
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
   494
  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
   495
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   496
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
   497
  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
   498
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
   499
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
   500
  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
   501
66806
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   502
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
   503
  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
   504
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
   505
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
   506
  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
   507
66806
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   508
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
   509
  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
   510
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   511
text \<open>Negation respects modular equivalence.\<close>
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   512
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   513
lemma mod_minus_eq [mod_simps]:
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   514
  "(- (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
   515
proof -
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   516
  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
   517
    by (simp only: div_mult_mod_eq)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   518
  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
   519
    by (simp add: ac_simps)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   520
  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
   521
    by (rule mod_mult_self1)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   522
  finally show ?thesis
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   523
    by (rule sym)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   524
qed
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   525
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   526
lemma mod_minus_cong:
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   527
  assumes "a mod b = a' mod b"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   528
  shows "(- a) mod b = (- a') mod b"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   529
proof -
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   530
  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
   531
    unfolding assms ..
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   532
  then show ?thesis
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   533
    by (simp add: mod_minus_eq)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   534
qed
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   535
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   536
text \<open>Subtraction respects modular equivalence.\<close>
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   537
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   538
lemma mod_diff_left_eq [mod_simps]:
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   539
  "(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
   540
  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
   541
  by simp
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   542
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   543
lemma mod_diff_right_eq [mod_simps]:
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   544
  "(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
   545
  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
   546
  by simp
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   547
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   548
lemma mod_diff_eq:
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   549
  "(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
   550
  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
   551
  by simp
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   552
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   553
lemma mod_diff_cong:
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   554
  assumes "a mod c = a' mod c"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   555
  assumes "b mod c = b' mod c"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   556
  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
   557
  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
   558
  by simp
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   559
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   560
lemma minus_mod_self2 [simp]:
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   561
  "(a - b) mod b = a mod b"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   562
  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
   563
  by (simp add: mod_diff_right_eq)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   564
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   565
lemma minus_mod_self1 [simp]:
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   566
  "(b - a) mod b = - a mod b"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   567
  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
   568
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   569
lemma mod_eq_dvd_iff:
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   570
  "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
   571
proof
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   572
  assume ?P
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   573
  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
   574
    by simp
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   575
  then show ?Q
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   576
    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
   577
next
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   578
  assume ?Q
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   579
  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
   580
  then have "a = c * d + b"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   581
    by (simp add: algebra_simps)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   582
  then show ?P by simp
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   583
qed
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   584
66837
6ba663ff2b1c tuned proofs
haftmann
parents: 66817
diff changeset
   585
lemma mod_eqE:
6ba663ff2b1c tuned proofs
haftmann
parents: 66817
diff changeset
   586
  assumes "a mod c = b mod c"
6ba663ff2b1c tuned proofs
haftmann
parents: 66817
diff changeset
   587
  obtains d where "b = a + c * d"
6ba663ff2b1c tuned proofs
haftmann
parents: 66817
diff changeset
   588
proof -
6ba663ff2b1c tuned proofs
haftmann
parents: 66817
diff changeset
   589
  from assms have "c dvd a - b"
6ba663ff2b1c tuned proofs
haftmann
parents: 66817
diff changeset
   590
    by (simp add: mod_eq_dvd_iff)
6ba663ff2b1c tuned proofs
haftmann
parents: 66817
diff changeset
   591
  then obtain d where "a - b = c * d" ..
6ba663ff2b1c tuned proofs
haftmann
parents: 66817
diff changeset
   592
  then have "b = a + c * - d"
6ba663ff2b1c tuned proofs
haftmann
parents: 66817
diff changeset
   593
    by (simp add: algebra_simps)
6ba663ff2b1c tuned proofs
haftmann
parents: 66817
diff changeset
   594
  with that show thesis .
6ba663ff2b1c tuned proofs
haftmann
parents: 66817
diff changeset
   595
qed
6ba663ff2b1c tuned proofs
haftmann
parents: 66817
diff changeset
   596
67051
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
   597
lemma invertible_coprime:
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
   598
  "coprime a c" if "a * b mod c = 1"
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
   599
  by (rule coprimeI) (use that dvd_mod_iff [of _ c "a * b"] in auto)
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
   600
66806
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   601
end
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   602
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   603
  
64785
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   604
subsection \<open>Uniquely determined division\<close>
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   605
  
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   606
class unique_euclidean_semiring = euclidean_semiring + 
66840
0d689d71dbdc canonical multiplicative euclidean size
haftmann
parents: 66839
diff changeset
   607
  assumes euclidean_size_mult: "euclidean_size (a * b) = euclidean_size a * euclidean_size b"
66838
17989f6bc7b2 clarified uniqueness criterion for euclidean rings
haftmann
parents: 66837
diff changeset
   608
  fixes division_segment :: "'a \<Rightarrow> 'a"
66839
909ba5ed93dd clarified parity
haftmann
parents: 66838
diff changeset
   609
  assumes is_unit_division_segment [simp]: "is_unit (division_segment a)"
66838
17989f6bc7b2 clarified uniqueness criterion for euclidean rings
haftmann
parents: 66837
diff changeset
   610
    and division_segment_mult:
17989f6bc7b2 clarified uniqueness criterion for euclidean rings
haftmann
parents: 66837
diff changeset
   611
    "a \<noteq> 0 \<Longrightarrow> b \<noteq> 0 \<Longrightarrow> division_segment (a * b) = division_segment a * division_segment b"
17989f6bc7b2 clarified uniqueness criterion for euclidean rings
haftmann
parents: 66837
diff changeset
   612
    and division_segment_mod:
17989f6bc7b2 clarified uniqueness criterion for euclidean rings
haftmann
parents: 66837
diff changeset
   613
    "b \<noteq> 0 \<Longrightarrow> \<not> b dvd a \<Longrightarrow> division_segment (a mod b) = division_segment b"
64785
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   614
  assumes div_bounded:
66838
17989f6bc7b2 clarified uniqueness criterion for euclidean rings
haftmann
parents: 66837
diff changeset
   615
    "b \<noteq> 0 \<Longrightarrow> division_segment r = division_segment b
64785
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   616
    \<Longrightarrow> euclidean_size r < euclidean_size b
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   617
    \<Longrightarrow> (q * b + r) div b = q"
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   618
begin
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   619
66839
909ba5ed93dd clarified parity
haftmann
parents: 66838
diff changeset
   620
lemma division_segment_not_0 [simp]:
909ba5ed93dd clarified parity
haftmann
parents: 66838
diff changeset
   621
  "division_segment a \<noteq> 0"
909ba5ed93dd clarified parity
haftmann
parents: 66838
diff changeset
   622
  using is_unit_division_segment [of a] is_unitE [of "division_segment a"] by blast
909ba5ed93dd clarified parity
haftmann
parents: 66838
diff changeset
   623
64785
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   624
lemma divmod_cases [case_names divides remainder by0]:
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   625
  obtains 
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   626
    (divides) q where "b \<noteq> 0"
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   627
      and "a div b = q"
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   628
      and "a mod b = 0"
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   629
      and "a = q * b"
66814
a24cde9588bb generalized some rules
haftmann
parents: 66813
diff changeset
   630
  | (remainder) q r where "b \<noteq> 0"
66838
17989f6bc7b2 clarified uniqueness criterion for euclidean rings
haftmann
parents: 66837
diff changeset
   631
      and "division_segment r = division_segment b"
64785
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   632
      and "euclidean_size r < euclidean_size b"
66814
a24cde9588bb generalized some rules
haftmann
parents: 66813
diff changeset
   633
      and "r \<noteq> 0"
64785
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   634
      and "a div b = q"
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   635
      and "a mod b = r"
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   636
      and "a = q * b + r"
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   637
  | (by0) "b = 0"
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   638
proof (cases "b = 0")
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   639
  case True
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   640
  then show thesis
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   641
  by (rule by0)
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   642
next
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   643
  case False
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   644
  show thesis
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   645
  proof (cases "b dvd a")
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   646
    case True
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   647
    then obtain q where "a = b * q" ..
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   648
    with \<open>b \<noteq> 0\<close> divides
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   649
    show thesis
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   650
      by (simp add: ac_simps)
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   651
  next
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   652
    case False
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   653
    then have "a mod b \<noteq> 0"
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   654
      by (simp add: mod_eq_0_iff_dvd)
66838
17989f6bc7b2 clarified uniqueness criterion for euclidean rings
haftmann
parents: 66837
diff changeset
   655
    moreover from \<open>b \<noteq> 0\<close> \<open>\<not> b dvd a\<close> have "division_segment (a mod b) = division_segment b"
17989f6bc7b2 clarified uniqueness criterion for euclidean rings
haftmann
parents: 66837
diff changeset
   656
      by (rule division_segment_mod)
64785
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   657
    moreover have "euclidean_size (a mod b) < euclidean_size b"
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   658
      using \<open>b \<noteq> 0\<close> by (rule mod_size_less)
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   659
    moreover have "a = a div b * b + a mod b"
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   660
      by (simp add: div_mult_mod_eq)
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   661
    ultimately show thesis
66838
17989f6bc7b2 clarified uniqueness criterion for euclidean rings
haftmann
parents: 66837
diff changeset
   662
      using \<open>b \<noteq> 0\<close> by (blast intro!: remainder)
64785
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   663
  qed
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   664
qed
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   665
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   666
lemma div_eqI:
66838
17989f6bc7b2 clarified uniqueness criterion for euclidean rings
haftmann
parents: 66837
diff changeset
   667
  "a div b = q" if "b \<noteq> 0" "division_segment r = division_segment b"
64785
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   668
    "euclidean_size r < euclidean_size b" "q * b + r = a"
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   669
proof -
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   670
  from that have "(q * b + r) div b = q"
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   671
    by (auto intro: div_bounded)
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   672
  with that show ?thesis
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   673
    by simp
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   674
qed
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   675
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   676
lemma mod_eqI:
66838
17989f6bc7b2 clarified uniqueness criterion for euclidean rings
haftmann
parents: 66837
diff changeset
   677
  "a mod b = r" if "b \<noteq> 0" "division_segment r = division_segment b"
64785
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   678
    "euclidean_size r < euclidean_size b" "q * b + r = a" 
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   679
proof -
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   680
  from that have "a div b = q"
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   681
    by (rule div_eqI)
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   682
  moreover have "a div b * b + a mod b = a"
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   683
    by (fact div_mult_mod_eq)
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   684
  ultimately have "a div b * b + a mod b = a div b * b + r"
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   685
    using \<open>q * b + r = a\<close> by simp
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   686
  then show ?thesis
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   687
    by simp
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   688
qed
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   689
66806
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   690
subclass euclidean_semiring_cancel
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   691
proof
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   692
  show "(a + c * b) div b = c + a div b" if "b \<noteq> 0" for a b c
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   693
  proof (cases a b rule: divmod_cases)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   694
    case by0
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   695
    with \<open>b \<noteq> 0\<close> show ?thesis
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   696
      by simp
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   697
  next
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   698
    case (divides q)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   699
    then show ?thesis
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   700
      by (simp add: ac_simps)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   701
  next
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   702
    case (remainder q r)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   703
    then show ?thesis
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   704
      by (auto intro: div_eqI simp add: algebra_simps)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   705
  qed
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   706
next
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   707
  show"(c * a) div (c * b) = a div b" if "c \<noteq> 0" for a b c
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   708
  proof (cases a b rule: divmod_cases)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   709
    case by0
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   710
    then show ?thesis
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   711
      by simp
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   712
  next
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   713
    case (divides q)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   714
    with \<open>c \<noteq> 0\<close> show ?thesis
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   715
      by (simp add: mult.left_commute [of c])
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   716
  next
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   717
    case (remainder q r)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   718
    from \<open>b \<noteq> 0\<close> \<open>c \<noteq> 0\<close> have "b * c \<noteq> 0"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   719
      by simp
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   720
    from remainder \<open>c \<noteq> 0\<close>
66838
17989f6bc7b2 clarified uniqueness criterion for euclidean rings
haftmann
parents: 66837
diff changeset
   721
    have "division_segment (r * c) = division_segment (b * c)"
66806
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   722
      and "euclidean_size (r * c) < euclidean_size (b * c)"
66840
0d689d71dbdc canonical multiplicative euclidean size
haftmann
parents: 66839
diff changeset
   723
      by (simp_all add: 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
   724
    with remainder show ?thesis
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   725
      by (auto intro!: div_eqI [of _ "c * (a mod b)"] simp add: algebra_simps)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   726
        (use \<open>b * c \<noteq> 0\<close> in simp)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   727
  qed
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   728
qed
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   729
66814
a24cde9588bb generalized some rules
haftmann
parents: 66813
diff changeset
   730
lemma div_mult1_eq:
a24cde9588bb generalized some rules
haftmann
parents: 66813
diff changeset
   731
  "(a * b) div c = a * (b div c) + a * (b mod c) div c"
a24cde9588bb generalized some rules
haftmann
parents: 66813
diff changeset
   732
proof (cases "a * (b mod c)" c rule: divmod_cases)
a24cde9588bb generalized some rules
haftmann
parents: 66813
diff changeset
   733
  case (divides q)
a24cde9588bb generalized some rules
haftmann
parents: 66813
diff changeset
   734
  have "a * b = a * (b div c * c + b mod c)"
a24cde9588bb generalized some rules
haftmann
parents: 66813
diff changeset
   735
    by (simp add: div_mult_mod_eq)
a24cde9588bb generalized some rules
haftmann
parents: 66813
diff changeset
   736
  also have "\<dots> = (a * (b div c) + q) * c"
a24cde9588bb generalized some rules
haftmann
parents: 66813
diff changeset
   737
    using divides by (simp add: algebra_simps)
a24cde9588bb generalized some rules
haftmann
parents: 66813
diff changeset
   738
  finally have "(a * b) div c = \<dots> div c"
a24cde9588bb generalized some rules
haftmann
parents: 66813
diff changeset
   739
    by simp
a24cde9588bb generalized some rules
haftmann
parents: 66813
diff changeset
   740
  with divides show ?thesis
a24cde9588bb generalized some rules
haftmann
parents: 66813
diff changeset
   741
    by simp
a24cde9588bb generalized some rules
haftmann
parents: 66813
diff changeset
   742
next
a24cde9588bb generalized some rules
haftmann
parents: 66813
diff changeset
   743
  case (remainder q r)
a24cde9588bb generalized some rules
haftmann
parents: 66813
diff changeset
   744
  from remainder(1-3) show ?thesis
a24cde9588bb generalized some rules
haftmann
parents: 66813
diff changeset
   745
  proof (rule div_eqI)
a24cde9588bb generalized some rules
haftmann
parents: 66813
diff changeset
   746
    have "a * b = a * (b div c * c + b mod c)"
a24cde9588bb generalized some rules
haftmann
parents: 66813
diff changeset
   747
      by (simp add: div_mult_mod_eq)
a24cde9588bb generalized some rules
haftmann
parents: 66813
diff changeset
   748
    also have "\<dots> = a * c * (b div c) + q * c + r"
a24cde9588bb generalized some rules
haftmann
parents: 66813
diff changeset
   749
      using remainder by (simp add: algebra_simps)
a24cde9588bb generalized some rules
haftmann
parents: 66813
diff changeset
   750
    finally show "(a * (b div c) + a * (b mod c) div c) * c + r = a * b"
a24cde9588bb generalized some rules
haftmann
parents: 66813
diff changeset
   751
      using remainder(5-7) by (simp add: algebra_simps)
a24cde9588bb generalized some rules
haftmann
parents: 66813
diff changeset
   752
  qed
a24cde9588bb generalized some rules
haftmann
parents: 66813
diff changeset
   753
next
a24cde9588bb generalized some rules
haftmann
parents: 66813
diff changeset
   754
  case by0
a24cde9588bb generalized some rules
haftmann
parents: 66813
diff changeset
   755
  then show ?thesis
a24cde9588bb generalized some rules
haftmann
parents: 66813
diff changeset
   756
    by simp
a24cde9588bb generalized some rules
haftmann
parents: 66813
diff changeset
   757
qed
a24cde9588bb generalized some rules
haftmann
parents: 66813
diff changeset
   758
a24cde9588bb generalized some rules
haftmann
parents: 66813
diff changeset
   759
lemma div_add1_eq:
a24cde9588bb generalized some rules
haftmann
parents: 66813
diff changeset
   760
  "(a + b) div c = a div c + b div c + (a mod c + b mod c) div c"
a24cde9588bb generalized some rules
haftmann
parents: 66813
diff changeset
   761
proof (cases "a mod c + b mod c" c rule: divmod_cases)
a24cde9588bb generalized some rules
haftmann
parents: 66813
diff changeset
   762
  case (divides q)
a24cde9588bb generalized some rules
haftmann
parents: 66813
diff changeset
   763
  have "a + b = (a div c * c + a mod c) + (b div c * c + b mod c)"
a24cde9588bb generalized some rules
haftmann
parents: 66813
diff changeset
   764
    using mod_mult_div_eq [of a c] mod_mult_div_eq [of b c] by (simp add: ac_simps)
a24cde9588bb generalized some rules
haftmann
parents: 66813
diff changeset
   765
  also have "\<dots> = (a div c + b div c) * c + (a mod c + b mod c)"
a24cde9588bb generalized some rules
haftmann
parents: 66813
diff changeset
   766
    by (simp add: algebra_simps)
a24cde9588bb generalized some rules
haftmann
parents: 66813
diff changeset
   767
  also have "\<dots> = (a div c + b div c + q) * c"
a24cde9588bb generalized some rules
haftmann
parents: 66813
diff changeset
   768
    using divides by (simp add: algebra_simps)
a24cde9588bb generalized some rules
haftmann
parents: 66813
diff changeset
   769
  finally have "(a + b) div c = (a div c + b div c + q) * c div c"
a24cde9588bb generalized some rules
haftmann
parents: 66813
diff changeset
   770
    by simp
a24cde9588bb generalized some rules
haftmann
parents: 66813
diff changeset
   771
  with divides show ?thesis
a24cde9588bb generalized some rules
haftmann
parents: 66813
diff changeset
   772
    by simp
a24cde9588bb generalized some rules
haftmann
parents: 66813
diff changeset
   773
next
a24cde9588bb generalized some rules
haftmann
parents: 66813
diff changeset
   774
  case (remainder q r)
a24cde9588bb generalized some rules
haftmann
parents: 66813
diff changeset
   775
  from remainder(1-3) show ?thesis
a24cde9588bb generalized some rules
haftmann
parents: 66813
diff changeset
   776
  proof (rule div_eqI)
a24cde9588bb generalized some rules
haftmann
parents: 66813
diff changeset
   777
    have "(a div c + b div c + q) * c + r + (a mod c + b mod c) =
a24cde9588bb generalized some rules
haftmann
parents: 66813
diff changeset
   778
        (a div c * c + a mod c) + (b div c * c + b mod c) + q * c + r"
a24cde9588bb generalized some rules
haftmann
parents: 66813
diff changeset
   779
      by (simp add: algebra_simps)
a24cde9588bb generalized some rules
haftmann
parents: 66813
diff changeset
   780
    also have "\<dots> = a + b + (a mod c + b mod c)"
a24cde9588bb generalized some rules
haftmann
parents: 66813
diff changeset
   781
      by (simp add: div_mult_mod_eq remainder) (simp add: ac_simps)
a24cde9588bb generalized some rules
haftmann
parents: 66813
diff changeset
   782
    finally show "(a div c + b div c + (a mod c + b mod c) div c) * c + r = a + b"
a24cde9588bb generalized some rules
haftmann
parents: 66813
diff changeset
   783
      using remainder by simp
a24cde9588bb generalized some rules
haftmann
parents: 66813
diff changeset
   784
  qed
a24cde9588bb generalized some rules
haftmann
parents: 66813
diff changeset
   785
next
a24cde9588bb generalized some rules
haftmann
parents: 66813
diff changeset
   786
  case by0
a24cde9588bb generalized some rules
haftmann
parents: 66813
diff changeset
   787
  then show ?thesis
a24cde9588bb generalized some rules
haftmann
parents: 66813
diff changeset
   788
    by simp
a24cde9588bb generalized some rules
haftmann
parents: 66813
diff changeset
   789
qed
a24cde9588bb generalized some rules
haftmann
parents: 66813
diff changeset
   790
66886
960509bfd47e added lemmas and tuned proofs
haftmann
parents: 66840
diff changeset
   791
lemma div_eq_0_iff:
960509bfd47e added lemmas and tuned proofs
haftmann
parents: 66840
diff changeset
   792
  "a div b = 0 \<longleftrightarrow> euclidean_size a < euclidean_size b \<or> b = 0" (is "_ \<longleftrightarrow> ?P")
960509bfd47e added lemmas and tuned proofs
haftmann
parents: 66840
diff changeset
   793
  if "division_segment a = division_segment b"
960509bfd47e added lemmas and tuned proofs
haftmann
parents: 66840
diff changeset
   794
proof
960509bfd47e added lemmas and tuned proofs
haftmann
parents: 66840
diff changeset
   795
  assume ?P
960509bfd47e added lemmas and tuned proofs
haftmann
parents: 66840
diff changeset
   796
  with that show "a div b = 0"
960509bfd47e added lemmas and tuned proofs
haftmann
parents: 66840
diff changeset
   797
    by (cases "b = 0") (auto intro: div_eqI)
960509bfd47e added lemmas and tuned proofs
haftmann
parents: 66840
diff changeset
   798
next
960509bfd47e added lemmas and tuned proofs
haftmann
parents: 66840
diff changeset
   799
  assume "a div b = 0"
960509bfd47e added lemmas and tuned proofs
haftmann
parents: 66840
diff changeset
   800
  then have "a mod b = a"
960509bfd47e added lemmas and tuned proofs
haftmann
parents: 66840
diff changeset
   801
    using div_mult_mod_eq [of a b] by simp
960509bfd47e added lemmas and tuned proofs
haftmann
parents: 66840
diff changeset
   802
  with mod_size_less [of b a] show ?P
960509bfd47e added lemmas and tuned proofs
haftmann
parents: 66840
diff changeset
   803
    by auto
960509bfd47e added lemmas and tuned proofs
haftmann
parents: 66840
diff changeset
   804
qed
960509bfd47e added lemmas and tuned proofs
haftmann
parents: 66840
diff changeset
   805
64785
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   806
end
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   807
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   808
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
   809
begin
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   810
  
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   811
subclass euclidean_ring_cancel ..
64785
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   812
ae0bbc8e45ad moved euclidean ring to HOL
haftmann
parents:
diff changeset
   813
end
66806
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   814
66808
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   815
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 68536
diff changeset
   816
subsection \<open>Euclidean division on \<^typ>\<open>nat\<close>\<close>
66808
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   817
66816
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
   818
instantiation nat :: normalization_semidom
66808
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   819
begin
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   820
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   821
definition normalize_nat :: "nat \<Rightarrow> nat"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   822
  where [simp]: "normalize = (id :: nat \<Rightarrow> nat)"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   823
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   824
definition unit_factor_nat :: "nat \<Rightarrow> nat"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   825
  where "unit_factor n = (if n = 0 then 0 else 1 :: nat)"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   826
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   827
lemma unit_factor_simps [simp]:
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   828
  "unit_factor 0 = (0::nat)"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   829
  "unit_factor (Suc n) = 1"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   830
  by (simp_all add: unit_factor_nat_def)
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   831
66816
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
   832
definition divide_nat :: "nat \<Rightarrow> nat \<Rightarrow> nat"
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
   833
  where "m div n = (if n = 0 then 0 else Max {k::nat. k * n \<le> m})"
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
   834
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
   835
instance
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
   836
  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
   837
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
   838
end
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
   839
67051
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
   840
lemma coprime_Suc_0_left [simp]:
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
   841
  "coprime (Suc 0) n"
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
   842
  using coprime_1_left [of n] by simp
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
   843
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
   844
lemma coprime_Suc_0_right [simp]:
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
   845
  "coprime n (Suc 0)"
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
   846
  using coprime_1_right [of n] by simp
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
   847
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
   848
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
   849
  for a b :: nat
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
   850
  by (drule coprime_common_divisor [of _ _ x]) simp_all
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
   851
66816
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
   852
instantiation nat :: unique_euclidean_semiring
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
   853
begin
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
   854
66808
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   855
definition euclidean_size_nat :: "nat \<Rightarrow> nat"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   856
  where [simp]: "euclidean_size_nat = id"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   857
66838
17989f6bc7b2 clarified uniqueness criterion for euclidean rings
haftmann
parents: 66837
diff changeset
   858
definition division_segment_nat :: "nat \<Rightarrow> nat"
17989f6bc7b2 clarified uniqueness criterion for euclidean rings
haftmann
parents: 66837
diff changeset
   859
  where [simp]: "division_segment_nat n = 1"
66808
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   860
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   861
definition modulo_nat :: "nat \<Rightarrow> nat \<Rightarrow> nat"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   862
  where "m mod n = m - (m div n * (n::nat))"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   863
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   864
instance proof
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   865
  fix m n :: nat
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   866
  have ex: "\<exists>k. k * n \<le> l" for l :: nat
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   867
    by (rule exI [of _ 0]) simp
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   868
  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
   869
  proof -
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   870
    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
   871
      by (cases n) auto
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   872
    then show ?thesis
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   873
      by (rule finite_subset) simp
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   874
  qed
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   875
  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
   876
  proof (cases "n = 0")
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   877
    case True
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   878
    moreover have "{l. l = 0 \<and> l \<le> m} = {0::nat}"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   879
      by auto
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   880
    ultimately show ?thesis
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   881
      by simp
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   882
  next
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   883
    case False
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   884
    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
   885
      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
   886
    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
   887
      by (auto simp add: ac_simps elim!: dvdE)
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   888
    finally show ?thesis
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   889
      using False by (simp add: divide_nat_def ac_simps)
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   890
  qed
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   891
  have less_eq: "m div n * n \<le> m"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   892
    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
   893
  then show "m div n * n + m mod n = m"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   894
    by (simp add: modulo_nat_def)
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   895
  assume "n \<noteq> 0" 
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   896
  show "euclidean_size (m mod n) < euclidean_size n"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   897
  proof -
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   898
    have "m < Suc (m div n) * n"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   899
    proof (rule ccontr)
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   900
      assume "\<not> m < Suc (m div n) * n"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   901
      then have "Suc (m div n) * n \<le> m"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   902
        by (simp add: not_less)
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   903
      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
   904
        by (simp add: divide_nat_def)
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   905
      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
   906
        by auto
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   907
      ultimately have "Suc (m div n) < Suc (m div n)"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   908
        by blast
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   909
      then show False
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   910
        by simp
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   911
    qed
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   912
    with \<open>n \<noteq> 0\<close> show ?thesis
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   913
      by (simp add: modulo_nat_def)
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   914
  qed
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   915
  show "euclidean_size m \<le> euclidean_size (m * n)"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   916
    using \<open>n \<noteq> 0\<close> by (cases n) simp_all
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   917
  fix q r :: nat
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   918
  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
   919
  proof -
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   920
    from that have "r < n"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   921
      by simp
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   922
    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
   923
    proof (rule ccontr)
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   924
      assume "\<not> k \<le> q"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   925
      then have "q < k"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   926
        by simp
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   927
      then obtain l where "k = Suc (q + l)"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   928
        by (auto simp add: less_iff_Suc_add)
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   929
      with \<open>r < n\<close> that show False
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   930
        by (simp add: algebra_simps)
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   931
    qed
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   932
    with \<open>n \<noteq> 0\<close> ex fin show ?thesis
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   933
      by (auto simp add: divide_nat_def Max_eq_iff)
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   934
  qed
66816
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
   935
qed simp_all
66808
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   936
66806
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66798
diff changeset
   937
end
66808
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   938
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   939
text \<open>Tool support\<close>
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   940
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   941
ML \<open>
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   942
structure Cancel_Div_Mod_Nat = Cancel_Div_Mod
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   943
(
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 68536
diff changeset
   944
  val div_name = \<^const_name>\<open>divide\<close>;
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 68536
diff changeset
   945
  val mod_name = \<^const_name>\<open>modulo\<close>;
66808
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   946
  val mk_binop = HOLogic.mk_binop;
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 68536
diff changeset
   947
  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
   948
  val mk_sum = Arith_Data.mk_sum;
66808
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   949
  fun dest_sum tm =
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   950
    if HOLogic.is_zero tm then []
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   951
    else
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   952
      (case try HOLogic.dest_Suc tm of
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   953
        SOME t => HOLogic.Suc_zero :: dest_sum t
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   954
      | NONE =>
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   955
          (case try dest_plus tm of
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   956
            SOME (t, u) => dest_sum t @ dest_sum u
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   957
          | NONE => [tm]));
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   958
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   959
  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
   960
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   961
  val prove_eq_sums = Arith_Data.prove_conv2 all_tac
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   962
    (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
   963
)
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   964
\<close>
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   965
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   966
simproc_setup cancel_div_mod_nat ("(m::nat) + n") =
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   967
  \<open>K Cancel_Div_Mod_Nat.proc\<close>
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   968
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   969
lemma div_nat_eqI:
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   970
  "m div n = q" if "n * q \<le> m" and "m < n * Suc q" for m n q :: nat
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   971
  by (rule div_eqI [of _ "m - n * q"]) (use that in \<open>simp_all add: algebra_simps\<close>)
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   972
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   973
lemma mod_nat_eqI:
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   974
  "m mod n = r" if "r < n" and "r \<le> m" and "n dvd m - r" for m n r :: nat
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   975
  by (rule mod_eqI [of _ _ "(m - r) div n"]) (use that in \<open>simp_all add: algebra_simps\<close>)
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   976
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   977
lemma div_mult_self_is_m [simp]:
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   978
  "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
   979
  using that by simp
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   980
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   981
lemma div_mult_self1_is_m [simp]:
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   982
  "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
   983
  using that by simp
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   984
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   985
lemma mod_less_divisor [simp]:
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   986
  "m mod n < n" if "n > 0" for m n :: nat
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   987
  using mod_size_less [of n m] that by simp
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   988
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   989
lemma mod_le_divisor [simp]:
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   990
  "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
   991
  using that by (auto simp add: le_less)
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   992
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   993
lemma div_times_less_eq_dividend [simp]:
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   994
  "m div n * n \<le> m" for m n :: nat
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   995
  by (simp add: minus_mod_eq_div_mult [symmetric])
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   996
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   997
lemma times_div_less_eq_dividend [simp]:
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   998
  "n * (m div n) \<le> m" for m n :: nat
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
   999
  using div_times_less_eq_dividend [of m n]
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1000
  by (simp add: ac_simps)
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1001
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1002
lemma dividend_less_div_times:
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1003
  "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
  1004
proof -
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1005
  from that have "m mod n < n"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1006
    by simp
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1007
  then show ?thesis
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1008
    by (simp add: minus_mod_eq_div_mult [symmetric])
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1009
qed
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1010
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1011
lemma dividend_less_times_div:
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1012
  "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
  1013
  using dividend_less_div_times [of n m] that
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1014
  by (simp add: ac_simps)
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1015
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1016
lemma mod_Suc_le_divisor [simp]:
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1017
  "m mod Suc n \<le> n"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1018
  using mod_less_divisor [of "Suc n" m] by arith
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1019
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1020
lemma mod_less_eq_dividend [simp]:
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1021
  "m mod n \<le> m" for m n :: nat
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1022
proof (rule add_leD2)
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1023
  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
  1024
  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
  1025
qed
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1026
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1027
lemma
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1028
  div_less [simp]: "m div n = 0"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1029
  and mod_less [simp]: "m mod n = m"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1030
  if "m < n" for m n :: nat
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1031
  using that by (auto intro: div_eqI mod_eqI) 
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1032
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1033
lemma le_div_geq:
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1034
  "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
  1035
proof -
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1036
  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
  1037
    by (auto simp add: le_iff_add)
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1038
  with \<open>0 < n\<close> show ?thesis
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1039
    by (simp add: div_add_self1)
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1040
qed
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1041
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1042
lemma le_mod_geq:
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1043
  "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
  1044
proof -
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1045
  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
  1046
    by (auto simp add: le_iff_add)
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1047
  then show ?thesis
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1048
    by simp
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1049
qed
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1050
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1051
lemma div_if:
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1052
  "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
  1053
  by (simp add: le_div_geq)
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
lemma mod_if:
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1056
  "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
  1057
  by (simp add: le_mod_geq)
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1058
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1059
lemma div_eq_0_iff:
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1060
  "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
  1061
  by (simp add: div_eq_0_iff)
66808
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_greater_zero_iff:
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1064
  "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
  1065
  using div_eq_0_iff [of m n] by auto
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 mod_greater_zero_iff_not_dvd:
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1068
  "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
  1069
  by (simp add: dvd_eq_mod_eq_0)
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 div_by_Suc_0 [simp]:
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1072
  "m div Suc 0 = m"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1073
  using div_by_1 [of m] 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_by_Suc_0 [simp]:
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1076
  "m mod Suc 0 = 0"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1077
  using mod_by_1 [of m] by simp
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 div2_Suc_Suc [simp]:
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1080
  "Suc (Suc m) div 2 = Suc (m div 2)"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1081
  by (simp add: numeral_2_eq_2 le_div_geq)
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 Suc_n_div_2_gt_zero [simp]:
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1084
  "0 < Suc n div 2" if "n > 0" for n :: nat
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1085
  using that by (cases n) simp_all
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1086
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1087
lemma div_2_gt_zero [simp]:
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1088
  "0 < n div 2" if "Suc 0 < n" for n :: nat
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1089
  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
  1090
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1091
lemma mod2_Suc_Suc [simp]:
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1092
  "Suc (Suc m) mod 2 = m mod 2"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1093
  by (simp add: numeral_2_eq_2 le_mod_geq)
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1094
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1095
lemma add_self_div_2 [simp]:
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1096
  "(m + m) div 2 = m" for m :: nat
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1097
  by (simp add: mult_2 [symmetric])
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1098
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1099
lemma add_self_mod_2 [simp]:
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1100
  "(m + m) mod 2 = 0" for m :: nat
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1101
  by (simp add: mult_2 [symmetric])
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1102
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1103
lemma mod2_gr_0 [simp]:
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1104
  "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
  1105
proof -
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1106
  have "m mod 2 < 2"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1107
    by (rule mod_less_divisor) simp
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1108
  then have "m mod 2 = 0 \<or> m mod 2 = 1"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1109
    by arith
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1110
  then show ?thesis
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1111
    by auto     
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1112
qed
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1113
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1114
lemma mod_Suc_eq [mod_simps]:
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1115
  "Suc (m mod n) mod n = Suc m mod n"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1116
proof -
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1117
  have "(m mod n + 1) mod n = (m + 1) mod n"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1118
    by (simp only: mod_simps)
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1119
  then show ?thesis
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1120
    by simp
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1121
qed
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1122
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1123
lemma mod_Suc_Suc_eq [mod_simps]:
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1124
  "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
  1125
proof -
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1126
  have "(m mod n + 2) mod n = (m + 2) mod n"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1127
    by (simp only: mod_simps)
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1128
  then show ?thesis
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1129
    by simp
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1130
qed
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1131
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1132
lemma
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1133
  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
  1134
  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
  1135
  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
  1136
  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
  1137
  by (subst mod_Suc_eq [symmetric], simp add: mod_simps)+
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1138
67083
6b2c0681ef28 new simp rule
haftmann
parents: 67051
diff changeset
  1139
lemma Suc_0_mod_eq [simp]:
6b2c0681ef28 new simp rule
haftmann
parents: 67051
diff changeset
  1140
  "Suc 0 mod n = of_bool (n \<noteq> Suc 0)"
6b2c0681ef28 new simp rule
haftmann
parents: 67051
diff changeset
  1141
  by (cases n) simp_all
6b2c0681ef28 new simp rule
haftmann
parents: 67051
diff changeset
  1142
66808
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1143
context
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1144
  fixes m n q :: nat
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1145
begin
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1146
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1147
private lemma eucl_rel_mult2:
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1148
  "m mod n + n * (m div n mod q) < n * q"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1149
  if "n > 0" and "q > 0"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1150
proof -
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1151
  from \<open>n > 0\<close> have "m mod n < n"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1152
    by (rule mod_less_divisor)
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1153
  from \<open>q > 0\<close> have "m div n mod q < q"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1154
    by (rule mod_less_divisor)
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1155
  then obtain s where "q = Suc (m div n mod q + s)"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1156
    by (blast dest: less_imp_Suc_add)
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1157
  moreover have "m mod n + n * (m div n mod q) < n * Suc (m div n mod q + s)"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1158
    using \<open>m mod n < n\<close> by (simp add: add_mult_distrib2)
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1159
  ultimately show ?thesis
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1160
    by simp
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 div_mult2_eq:
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1164
  "m div (n * q) = (m div n) div q"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1165
proof (cases "n = 0 \<or> q = 0")
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1166
  case True
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1167
  then show ?thesis
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1168
    by auto
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1169
next
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1170
  case False
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1171
  with eucl_rel_mult2 show ?thesis
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1172
    by (auto intro: div_eqI [of _ "n * (m div n mod q) + m mod n"]
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1173
      simp add: algebra_simps add_mult_distrib2 [symmetric])
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1174
qed
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_mult2_eq:
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1177
  "m mod (n * q) = n * (m div n mod q) + m mod n"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1178
proof (cases "n = 0 \<or> q = 0")
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1179
  case True
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1180
  then show ?thesis
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1181
    by auto
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1182
next
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1183
  case False
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1184
  with eucl_rel_mult2 show ?thesis
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1185
    by (auto intro: mod_eqI [of _ _ "(m div n) div q"]
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1186
      simp add: algebra_simps add_mult_distrib2 [symmetric])
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1187
qed
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1188
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1189
end
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1190
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1191
lemma div_le_mono:
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1192
  "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
  1193
proof -
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1194
  from that obtain q where "n = m + q"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1195
    by (auto simp add: le_iff_add)
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1196
  then show ?thesis
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1197
    by (simp add: div_add1_eq [of m q k])
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1198
qed
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1199
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 68536
diff changeset
  1200
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
  1201
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1202
lemma div_le_mono2:
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1203
  "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
  1204
using that proof (induct k arbitrary: m rule: less_induct)
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1205
  case (less k)
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1206
  show ?case
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1207
  proof (cases "n \<le> k")
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1208
    case False
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1209
    then show ?thesis
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1210
      by simp
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1211
  next
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1212
    case True
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1213
    have "(k - n) div n \<le> (k - m) div n"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1214
      using less.prems
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1215
      by (blast intro: div_le_mono diff_le_mono2)
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1216
    also have "\<dots> \<le> (k - m) div m"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1217
      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
  1218
      by simp
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1219
    finally show ?thesis
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1220
      using \<open>n \<le> k\<close> less.prems
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1221
      by (simp add: le_div_geq)
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1222
  qed
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1223
qed
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1224
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1225
lemma div_le_dividend [simp]:
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1226
  "m div n \<le> m" for m n :: nat
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1227
  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
  1228
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1229
lemma div_less_dividend [simp]:
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1230
  "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
  1231
using that proof (induct m rule: less_induct)
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1232
  case (less m)
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1233
  show ?case
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1234
  proof (cases "n < m")
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1235
    case False
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1236
    with less show ?thesis
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1237
      by (cases "n = m") simp_all
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1238
  next
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1239
    case True
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
      using less.hyps [of "m - n"] less.prems
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1242
      by (simp add: le_div_geq)
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1243
  qed
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1244
qed
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1245
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1246
lemma div_eq_dividend_iff:
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1247
  "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
  1248
proof
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1249
  assume "n = 1"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1250
  then show "m div n = m"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1251
    by simp
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1252
next
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1253
  assume P: "m div n = m"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1254
  show "n = 1"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1255
  proof (rule ccontr)
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1256
    have "n \<noteq> 0"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1257
      by (rule ccontr) (use that P in auto)
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1258
    moreover assume "n \<noteq> 1"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1259
    ultimately have "n > 1"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1260
      by simp
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1261
    with that have "m div n < m"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1262
      by simp
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1263
    with P show False
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1264
      by simp
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1265
  qed
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1266
qed
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1267
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1268
lemma less_mult_imp_div_less:
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1269
  "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
  1270
proof -
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1271
  from that have "i * n > 0"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1272
    by (cases "i * n = 0") simp_all
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1273
  then have "i > 0" and "n > 0"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1274
    by simp_all
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1275
  have "m div n * n \<le> m"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1276
    by simp
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1277
  then have "m div n * n < i * n"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1278
    using that by (rule le_less_trans)
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1279
  with \<open>n > 0\<close> show ?thesis
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1280
    by simp
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1281
qed
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1282
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1283
text \<open>A fact for the mutilated chess board\<close>
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1284
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1285
lemma mod_Suc:
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1286
  "Suc m mod n = (if Suc (m mod n) = n then 0 else Suc (m mod n))" (is "_ = ?rhs")
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1287
proof (cases "n = 0")
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1288
  case True
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1289
  then show ?thesis
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1290
    by simp
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1291
next
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1292
  case False
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1293
  have "Suc m mod n = Suc (m mod n) mod n"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1294
    by (simp add: mod_simps)
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1295
  also have "\<dots> = ?rhs"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1296
    using False by (auto intro!: mod_nat_eqI intro: neq_le_trans simp add: Suc_le_eq)
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1297
  finally show ?thesis .
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1298
qed
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1299
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1300
lemma Suc_times_mod_eq:
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1301
  "Suc (m * n) mod m = 1" if "Suc 0 < m"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1302
  using that by (simp add: mod_Suc)
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1303
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1304
lemma Suc_times_numeral_mod_eq [simp]:
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1305
  "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
  1306
  by (rule Suc_times_mod_eq) (use that in simp)
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1307
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1308
lemma Suc_div_le_mono [simp]:
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1309
  "m div n \<le> Suc m div n"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1310
  by (simp add: div_le_mono)
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
text \<open>These lemmas collapse some needless occurrences of Suc:
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1313
  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
  1314
  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
  1315
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1316
lemma div_Suc_eq_div_add3 [simp]:
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1317
  "m div Suc (Suc (Suc n)) = m div (3 + n)"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1318
  by (simp add: Suc3_eq_add_3)
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1319
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1320
lemma mod_Suc_eq_mod_add3 [simp]:
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1321
  "m mod Suc (Suc (Suc n)) = m mod (3 + n)"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1322
  by (simp add: Suc3_eq_add_3)
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1323
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1324
lemma Suc_div_eq_add3_div:
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1325
  "Suc (Suc (Suc m)) div n = (3 + m) div n"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1326
  by (simp add: Suc3_eq_add_3)
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1327
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1328
lemma Suc_mod_eq_add3_mod:
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1329
  "Suc (Suc (Suc m)) mod n = (3 + m) mod n"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1330
  by (simp add: Suc3_eq_add_3)
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1331
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1332
lemmas Suc_div_eq_add3_div_numeral [simp] =
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1333
  Suc_div_eq_add3_div [of _ "numeral v"] for v
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
lemmas Suc_mod_eq_add3_mod_numeral [simp] =
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1336
  Suc_mod_eq_add3_mod [of _ "numeral v"] for v
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1337
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1338
lemma (in field_char_0) of_nat_div:
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1339
  "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
  1340
proof -
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1341
  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
  1342
    unfolding of_nat_add by (cases "n = 0") simp_all
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1343
  then show ?thesis
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1344
    by simp
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1345
qed
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1346
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1347
text \<open>An ``induction'' law for modulus arithmetic.\<close>
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1348
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1349
lemma mod_induct [consumes 3, case_names step]:
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1350
  "P m" if "P n" and "n < p" and "m < p"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1351
    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
  1352
using \<open>m < p\<close> proof (induct m)
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1353
  case 0
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1354
  show ?case
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1355
  proof (rule ccontr)
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1356
    assume "\<not> P 0"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1357
    from \<open>n < p\<close> have "0 < p"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1358
      by simp
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1359
    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
  1360
      by (blast dest: less_imp_add_positive)
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1361
    with \<open>P n\<close> have "P (p - m)"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1362
      by simp
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1363
    moreover have "\<not> P (p - m)"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1364
    using \<open>0 < m\<close> proof (induct m)
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1365
      case 0
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1366
      then show ?case
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1367
        by simp
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1368
    next
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1369
      case (Suc m)
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1370
      show ?case
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1371
      proof
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1372
        assume P: "P (p - Suc m)"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1373
        with \<open>\<not> P 0\<close> have "Suc m < p"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1374
          by (auto intro: ccontr) 
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1375
        then have "Suc (p - Suc m) = p - m"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1376
          by arith
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1377
        moreover from \<open>0 < p\<close> have "p - Suc m < p"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1378
          by arith
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1379
        with P step have "P ((Suc (p - Suc m)) mod p)"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1380
          by blast
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1381
        ultimately show False
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1382
          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
  1383
      qed
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1384
    qed
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1385
    ultimately show False
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1386
      by blast
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1387
  qed
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1388
next
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1389
  case (Suc m)
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1390
  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
  1391
    by simp_all
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1392
  from \<open>m < p\<close> have "P m"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1393
    by (rule Suc.hyps)
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1394
  with \<open>m < p\<close> have "P (Suc m mod p)"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1395
    by (rule step)
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1396
  with mod show ?case
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1397
    by simp
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1398
qed
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1399
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1400
lemma split_div:
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1401
  "P (m div n) \<longleftrightarrow> (n = 0 \<longrightarrow> P 0) \<and> (n \<noteq> 0 \<longrightarrow>
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1402
     (\<forall>i j. j < n \<longrightarrow> m = n * i + j \<longrightarrow> P i))"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1403
     (is "?P = ?Q") for m n :: nat
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1404
proof (cases "n = 0")
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1405
  case True
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1406
  then show ?thesis
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1407
    by simp
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1408
next
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1409
  case False
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1410
  show ?thesis
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1411
  proof
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1412
    assume ?P
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1413
    with False show ?Q
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1414
      by auto
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1415
  next
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1416
    assume ?Q
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1417
    with False have *: "\<And>i j. j < n \<Longrightarrow> m = n * i + j \<Longrightarrow> P i"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1418
      by simp
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1419
    with False show ?P
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1420
      by (auto intro: * [of "m mod n"])
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1421
  qed
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1422
qed
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1423
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1424
lemma split_div':
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1425
  "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)"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1426
proof (cases "n = 0")
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1427
  case True
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1428
  then show ?thesis
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1429
    by simp
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1430
next
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1431
  case False
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1432
  then have "n * q \<le> m \<and> m < n * Suc q \<longleftrightarrow> m div n = q" for q
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1433
    by (auto intro: div_nat_eqI dividend_less_times_div)
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1434
  then show ?thesis
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1435
    by auto
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1436
qed
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1437
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1438
lemma split_mod:
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1439
  "P (m mod n) \<longleftrightarrow> (n = 0 \<longrightarrow> P m) \<and> (n \<noteq> 0 \<longrightarrow>
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1440
     (\<forall>i j. j < n \<longrightarrow> m = n * i + j \<longrightarrow> P j))"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1441
     (is "?P \<longleftrightarrow> ?Q") for m n :: nat
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1442
proof (cases "n = 0")
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1443
  case True
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1444
  then show ?thesis
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1445
    by simp
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1446
next
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1447
  case False
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1448
  show ?thesis
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1449
  proof
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1450
    assume ?P
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1451
    with False show ?Q
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1452
      by auto
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1453
  next
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1454
    assume ?Q
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1455
    with False have *: "\<And>i j. j < n \<Longrightarrow> m = n * i + j \<Longrightarrow> P j"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1456
      by simp
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1457
    with False show ?P
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1458
      by (auto intro: * [of _ "m div n"])
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1459
  qed
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1460
qed
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1461
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1462
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 68536
diff changeset
  1463
subsection \<open>Euclidean division on \<^typ>\<open>int\<close>\<close>
66816
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1464
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1465
instantiation int :: normalization_semidom
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1466
begin
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1467
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1468
definition normalize_int :: "int \<Rightarrow> int"
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1469
  where [simp]: "normalize = (abs :: int \<Rightarrow> int)"
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1470
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1471
definition unit_factor_int :: "int \<Rightarrow> int"
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1472
  where [simp]: "unit_factor = (sgn :: int \<Rightarrow> int)"
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1473
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1474
definition divide_int :: "int \<Rightarrow> int \<Rightarrow> int"
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1475
  where "k div l = (if l = 0 then 0
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1476
    else if sgn k = sgn l
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1477
      then int (nat \<bar>k\<bar> div nat \<bar>l\<bar>)
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1478
      else - int (nat \<bar>k\<bar> div nat \<bar>l\<bar> + of_bool (\<not> l dvd k)))"
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1479
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1480
lemma divide_int_unfold:
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1481
  "(sgn k * int m) div (sgn l * int n) =
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1482
   (if sgn l = 0 \<or> sgn k = 0 \<or> n = 0 then 0
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1483
    else if sgn k = sgn l
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1484
      then int (m div n)
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1485
      else - int (m div n + of_bool (\<not> n dvd m)))"
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1486
  by (auto simp add: divide_int_def sgn_0_0 sgn_1_pos sgn_mult abs_mult
67118
ccab07d1196c more simplification rules
haftmann
parents: 67087
diff changeset
  1487
    nat_mult_distrib)
66816
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1488
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1489
instance proof
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1490
  fix k :: int show "k div 0 = 0"
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1491
  by (simp add: divide_int_def)
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1492
next
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1493
  fix k l :: int
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1494
  assume "l \<noteq> 0"
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1495
  obtain n m and s t where k: "k = sgn s * int n" and l: "l = sgn t * int m" 
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1496
    by (blast intro: int_sgnE elim: that)
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1497
  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
  1498
    by (simp add: ac_simps sgn_mult)
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1499
  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
  1500
    by (simp only: divide_int_unfold)
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1501
      (auto simp add: algebra_simps sgn_mult sgn_1_pos sgn_0_0)
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1502
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
  1503
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1504
end
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1505
67051
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
  1506
lemma coprime_int_iff [simp]:
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
  1507
  "coprime (int m) (int n) \<longleftrightarrow> coprime m n" (is "?P \<longleftrightarrow> ?Q")
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
  1508
proof
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
  1509
  assume ?P
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
  1510
  show ?Q
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
  1511
  proof (rule coprimeI)
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
  1512
    fix q
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
  1513
    assume "q dvd m" "q dvd n"
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
  1514
    then have "int q dvd int m" "int q dvd int n"
67118
ccab07d1196c more simplification rules
haftmann
parents: 67087
diff changeset
  1515
      by simp_all
67051
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
  1516
    with \<open>?P\<close> have "is_unit (int q)"
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
  1517
      by (rule coprime_common_divisor)
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
  1518
    then show "is_unit q"
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
  1519
      by simp
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
  1520
  qed
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
  1521
next
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
  1522
  assume ?Q
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
  1523
  show ?P
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
  1524
  proof (rule coprimeI)
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
  1525
    fix k
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
  1526
    assume "k dvd int m" "k dvd int n"
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
  1527
    then have "nat \<bar>k\<bar> dvd m" "nat \<bar>k\<bar> dvd n"
67118
ccab07d1196c more simplification rules
haftmann
parents: 67087
diff changeset
  1528
      by simp_all
67051
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
  1529
    with \<open>?Q\<close> have "is_unit (nat \<bar>k\<bar>)"
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
  1530
      by (rule coprime_common_divisor)
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
  1531
    then show "is_unit k"
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
  1532
      by simp
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
  1533
  qed
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
  1534
qed
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
  1535
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
  1536
lemma coprime_abs_left_iff [simp]:
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
  1537
  "coprime \<bar>k\<bar> l \<longleftrightarrow> coprime k l" for k l :: int
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
  1538
  using coprime_normalize_left_iff [of k l] by simp
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
  1539
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
  1540
lemma coprime_abs_right_iff [simp]:
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
  1541
  "coprime k \<bar>l\<bar> \<longleftrightarrow> coprime k l" for k l :: int
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
  1542
  using coprime_abs_left_iff [of l k] by (simp add: ac_simps)
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
  1543
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
  1544
lemma coprime_nat_abs_left_iff [simp]:
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
  1545
  "coprime (nat \<bar>k\<bar>) n \<longleftrightarrow> coprime k (int n)"
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
  1546
proof -
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
  1547
  define m where "m = nat \<bar>k\<bar>"
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
  1548
  then have "\<bar>k\<bar> = int m"
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
  1549
    by simp
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
  1550
  moreover have "coprime k (int n) \<longleftrightarrow> coprime \<bar>k\<bar> (int n)"
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
  1551
    by simp
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
  1552
  ultimately show ?thesis
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
  1553
    by simp
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
  1554
qed
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
  1555
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
  1556
lemma coprime_nat_abs_right_iff [simp]:
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
  1557
  "coprime n (nat \<bar>k\<bar>) \<longleftrightarrow> coprime (int n) k"
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
  1558
  using coprime_nat_abs_left_iff [of k n] by (simp add: ac_simps)
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
  1559
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
  1560
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
  1561
  for a b :: int
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
  1562
  by (drule coprime_common_divisor [of _ _ x]) simp_all
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66886
diff changeset
  1563
66838
17989f6bc7b2 clarified uniqueness criterion for euclidean rings
haftmann
parents: 66837
diff changeset
  1564
instantiation int :: idom_modulo
66816
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1565
begin
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1566
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1567
definition modulo_int :: "int \<Rightarrow> int \<Rightarrow> int"
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1568
  where "k mod l = (if l = 0 then k
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1569
    else if sgn k = sgn l
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1570
      then sgn l * int (nat \<bar>k\<bar> mod nat \<bar>l\<bar>)
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1571
      else sgn l * (\<bar>l\<bar> * of_bool (\<not> l dvd k) - int (nat \<bar>k\<bar> mod nat \<bar>l\<bar>)))"
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1572
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1573
lemma modulo_int_unfold:
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1574
  "(sgn k * int m) mod (sgn l * int n) =
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1575
   (if sgn l = 0 \<or> sgn k = 0 \<or> n = 0 then sgn k * int m
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1576
    else if sgn k = sgn l
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1577
      then sgn l * int (m mod n)
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1578
      else sgn l * (int (n * of_bool (\<not> n dvd m)) - int (m mod n)))"
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1579
  by (auto simp add: modulo_int_def sgn_0_0 sgn_1_pos sgn_mult abs_mult
67118
ccab07d1196c more simplification rules
haftmann
parents: 67087
diff changeset
  1580
    nat_mult_distrib)
66816
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1581
66838
17989f6bc7b2 clarified uniqueness criterion for euclidean rings
haftmann
parents: 66837
diff changeset
  1582
instance proof
17989f6bc7b2 clarified uniqueness criterion for euclidean rings
haftmann
parents: 66837
diff changeset
  1583
  fix k l :: int
17989f6bc7b2 clarified uniqueness criterion for euclidean rings
haftmann
parents: 66837
diff changeset
  1584
  obtain n m and s t where "k = sgn s * int n" and "l = sgn t * int m" 
17989f6bc7b2 clarified uniqueness criterion for euclidean rings
haftmann
parents: 66837
diff changeset
  1585
    by (blast intro: int_sgnE elim: that)
17989f6bc7b2 clarified uniqueness criterion for euclidean rings
haftmann
parents: 66837
diff changeset
  1586
  then show "k div l * l + k mod l = k"
17989f6bc7b2 clarified uniqueness criterion for euclidean rings
haftmann
parents: 66837
diff changeset
  1587
    by (auto simp add: divide_int_unfold modulo_int_unfold algebra_simps dest!: sgn_not_eq_imp)
17989f6bc7b2 clarified uniqueness criterion for euclidean rings
haftmann
parents: 66837
diff changeset
  1588
       (simp_all add: of_nat_mult [symmetric] of_nat_add [symmetric]
17989f6bc7b2 clarified uniqueness criterion for euclidean rings
haftmann
parents: 66837
diff changeset
  1589
         distrib_left [symmetric] minus_mult_right
17989f6bc7b2 clarified uniqueness criterion for euclidean rings
haftmann
parents: 66837
diff changeset
  1590
         del: of_nat_mult minus_mult_right [symmetric])
17989f6bc7b2 clarified uniqueness criterion for euclidean rings
haftmann
parents: 66837
diff changeset
  1591
qed
17989f6bc7b2 clarified uniqueness criterion for euclidean rings
haftmann
parents: 66837
diff changeset
  1592
17989f6bc7b2 clarified uniqueness criterion for euclidean rings
haftmann
parents: 66837
diff changeset
  1593
end
17989f6bc7b2 clarified uniqueness criterion for euclidean rings
haftmann
parents: 66837
diff changeset
  1594
17989f6bc7b2 clarified uniqueness criterion for euclidean rings
haftmann
parents: 66837
diff changeset
  1595
instantiation int :: unique_euclidean_ring
17989f6bc7b2 clarified uniqueness criterion for euclidean rings
haftmann
parents: 66837
diff changeset
  1596
begin
17989f6bc7b2 clarified uniqueness criterion for euclidean rings
haftmann
parents: 66837
diff changeset
  1597
17989f6bc7b2 clarified uniqueness criterion for euclidean rings
haftmann
parents: 66837
diff changeset
  1598
definition euclidean_size_int :: "int \<Rightarrow> nat"
17989f6bc7b2 clarified uniqueness criterion for euclidean rings
haftmann
parents: 66837
diff changeset
  1599
  where [simp]: "euclidean_size_int = (nat \<circ> abs :: int \<Rightarrow> nat)"
17989f6bc7b2 clarified uniqueness criterion for euclidean rings
haftmann
parents: 66837
diff changeset
  1600
17989f6bc7b2 clarified uniqueness criterion for euclidean rings
haftmann
parents: 66837
diff changeset
  1601
definition division_segment_int :: "int \<Rightarrow> int"
17989f6bc7b2 clarified uniqueness criterion for euclidean rings
haftmann
parents: 66837
diff changeset
  1602
  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
  1603
17989f6bc7b2 clarified uniqueness criterion for euclidean rings
haftmann
parents: 66837
diff changeset
  1604
lemma division_segment_eq_sgn:
17989f6bc7b2 clarified uniqueness criterion for euclidean rings
haftmann
parents: 66837
diff changeset
  1605
  "division_segment k = sgn k" if "k \<noteq> 0" for k :: int
17989f6bc7b2 clarified uniqueness criterion for euclidean rings
haftmann
parents: 66837
diff changeset
  1606
  using that by (simp add: division_segment_int_def)
17989f6bc7b2 clarified uniqueness criterion for euclidean rings
haftmann
parents: 66837
diff changeset
  1607
17989f6bc7b2 clarified uniqueness criterion for euclidean rings
haftmann
parents: 66837
diff changeset
  1608
lemma abs_division_segment [simp]:
17989f6bc7b2 clarified uniqueness criterion for euclidean rings
haftmann
parents: 66837
diff changeset
  1609
  "\<bar>division_segment k\<bar> = 1" for k :: int
17989f6bc7b2 clarified uniqueness criterion for euclidean rings
haftmann
parents: 66837
diff changeset
  1610
  by (simp add: division_segment_int_def)
17989f6bc7b2 clarified uniqueness criterion for euclidean rings
haftmann
parents: 66837
diff changeset
  1611
66816
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1612
lemma abs_mod_less:
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1613
  "\<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
  1614
proof -
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1615
  obtain n m and s t where "k = sgn s * int n" and "l = sgn t * int m" 
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1616
    by (blast intro: int_sgnE elim: that)
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1617
  with that show ?thesis
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1618
    by (simp add: modulo_int_unfold sgn_0_0 sgn_1_pos sgn_1_neg
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1619
      abs_mult mod_greater_zero_iff_not_dvd)
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1620
qed
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1621
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1622
lemma sgn_mod:
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1623
  "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
  1624
proof -
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1625
  obtain n m and s t where "k = sgn s * int n" and "l = sgn t * int m" 
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1626
    by (blast intro: int_sgnE elim: that)
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1627
  with that show ?thesis
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1628
    by (simp add: modulo_int_unfold sgn_0_0 sgn_1_pos sgn_1_neg
67118
ccab07d1196c more simplification rules
haftmann
parents: 67087
diff changeset
  1629
      sgn_mult mod_eq_0_iff_dvd)
66816
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1630
qed
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1631
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1632
instance proof
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1633
  fix k l :: int
66838
17989f6bc7b2 clarified uniqueness criterion for euclidean rings
haftmann
parents: 66837
diff changeset
  1634
  show "division_segment (k mod l) = division_segment l" if
17989f6bc7b2 clarified uniqueness criterion for euclidean rings
haftmann
parents: 66837
diff changeset
  1635
    "l \<noteq> 0" and "\<not> l dvd k"
17989f6bc7b2 clarified uniqueness criterion for euclidean rings
haftmann
parents: 66837
diff changeset
  1636
    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
  1637
next
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1638
  fix l q r :: int
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1639
  obtain n m and s t
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1640
     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
  1641
    by (blast intro: int_sgnE elim: that)
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1642
  assume \<open>l \<noteq> 0\<close>
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1643
  with l have "s \<noteq> 0" and "n > 0"
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1644
    by (simp_all add: sgn_0_0)
66838
17989f6bc7b2 clarified uniqueness criterion for euclidean rings
haftmann
parents: 66837
diff changeset
  1645
  assume "division_segment r = division_segment l"
66816
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1646
  moreover have "r = sgn r * \<bar>r\<bar>"
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1647
    by (simp add: sgn_mult_abs)
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1648
  moreover define u where "u = nat \<bar>r\<bar>"
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1649
  ultimately have "r = sgn l * int u"
66838
17989f6bc7b2 clarified uniqueness criterion for euclidean rings
haftmann
parents: 66837
diff changeset
  1650
    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
  1651
  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
  1652
    by (simp add: sgn_mult)
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1653
  assume "euclidean_size r < euclidean_size l"
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1654
  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
  1655
    by (simp add: abs_mult)
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1656
  show "(q * l + r) div l = q"
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1657
  proof (cases "q = 0 \<or> r = 0")
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1658
    case True
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1659
    then show ?thesis
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1660
    proof
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1661
      assume "q = 0"
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1662
      then show ?thesis
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1663
        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
  1664
    next
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1665
      assume "r = 0"
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1666
      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
  1667
        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
  1668
      from \<open>s \<noteq> 0\<close> \<open>n > 0\<close> show ?thesis
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1669
        by (simp only: *, simp only: q l divide_int_unfold)
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1670
          (auto simp add: sgn_mult sgn_0_0 sgn_1_pos)
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1671
    qed
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1672
  next
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1673
    case False
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1674
    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
  1675
      by (simp_all add: sgn_0_0)
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1676
    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
  1677
      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
  1678
    ultimately have *: "q * l + r = sgn (s * t)
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1679
      * 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
  1680
      using l q r
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1681
      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
  1682
    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
  1683
      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
  1684
      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
  1685
    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
  1686
      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
  1687
      by auto
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1688
    ultimately show ?thesis
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1689
      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
  1690
      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
  1691
        (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
  1692
  qed
68536
e14848001c4c avoid pending shyps in global theory facts;
wenzelm
parents: 67118
diff changeset
  1693
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
  1694
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1695
end
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1696
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1697
lemma pos_mod_bound [simp]:
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1698
  "k mod l < l" if "l > 0" for k l :: int
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1699
proof -
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1700
  obtain m and s where "k = sgn s * int m"
69695
753ae9e9773d algebraized more material from theory Divides
haftmann
parents: 69593
diff changeset
  1701
    by (rule int_sgnE)
66816
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1702
  moreover from that obtain n where "l = sgn 1 * int n"
69695
753ae9e9773d algebraized more material from theory Divides
haftmann
parents: 69593
diff changeset
  1703
    by (cases l) simp_all
753ae9e9773d algebraized more material from theory Divides
haftmann
parents: 69593
diff changeset
  1704
  moreover from this that have "n > 0"
753ae9e9773d algebraized more material from theory Divides
haftmann
parents: 69593
diff changeset
  1705
    by simp
66816
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1706
  ultimately show ?thesis
69695
753ae9e9773d algebraized more material from theory Divides
haftmann
parents: 69593
diff changeset
  1707
    by (simp only: modulo_int_unfold)
753ae9e9773d algebraized more material from theory Divides
haftmann
parents: 69593
diff changeset
  1708
      (simp add: mod_greater_zero_iff_not_dvd)
753ae9e9773d algebraized more material from theory Divides
haftmann
parents: 69593
diff changeset
  1709
qed
753ae9e9773d algebraized more material from theory Divides
haftmann
parents: 69593
diff changeset
  1710
753ae9e9773d algebraized more material from theory Divides
haftmann
parents: 69593
diff changeset
  1711
lemma neg_mod_bound [simp]:
753ae9e9773d algebraized more material from theory Divides
haftmann
parents: 69593
diff changeset
  1712
  "l < k mod l" if "l < 0" for k l :: int
753ae9e9773d algebraized more material from theory Divides
haftmann
parents: 69593
diff changeset
  1713
proof -
753ae9e9773d algebraized more material from theory Divides
haftmann
parents: 69593
diff changeset
  1714
  obtain m and s where "k = sgn s * int m"
753ae9e9773d algebraized more material from theory Divides
haftmann
parents: 69593
diff changeset
  1715
    by (rule int_sgnE)
753ae9e9773d algebraized more material from theory Divides
haftmann
parents: 69593
diff changeset
  1716
  moreover from that obtain q where "l = sgn (- 1) * int (Suc q)"
753ae9e9773d algebraized more material from theory Divides
haftmann
parents: 69593
diff changeset
  1717
    by (cases l) simp_all
753ae9e9773d algebraized more material from theory Divides
haftmann
parents: 69593
diff changeset
  1718
  moreover define n where "n = Suc q"
753ae9e9773d algebraized more material from theory Divides
haftmann
parents: 69593
diff changeset
  1719
  then have "Suc q = n"
753ae9e9773d algebraized more material from theory Divides
haftmann
parents: 69593
diff changeset
  1720
    by simp
753ae9e9773d algebraized more material from theory Divides
haftmann
parents: 69593
diff changeset
  1721
  ultimately show ?thesis
753ae9e9773d algebraized more material from theory Divides
haftmann
parents: 69593
diff changeset
  1722
    by (simp only: modulo_int_unfold)
66816
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1723
      (simp add: mod_greater_zero_iff_not_dvd)
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1724
qed
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1725
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1726
lemma pos_mod_sign [simp]:
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1727
  "0 \<le> k mod l" if "l > 0" for k l :: int
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1728
proof -
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1729
  obtain m and s where "k = sgn s * int m"
69695
753ae9e9773d algebraized more material from theory Divides
haftmann
parents: 69593
diff changeset
  1730
    by (rule int_sgnE)
66816
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1731
  moreover from that obtain n where "l = sgn 1 * int n"
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1732
    by (cases l) auto
69695
753ae9e9773d algebraized more material from theory Divides
haftmann
parents: 69593
diff changeset
  1733
  moreover from this that have "n > 0"
753ae9e9773d algebraized more material from theory Divides
haftmann
parents: 69593
diff changeset
  1734
    by simp
66816
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1735
  ultimately show ?thesis
69695
753ae9e9773d algebraized more material from theory Divides
haftmann
parents: 69593
diff changeset
  1736
    by (simp only: modulo_int_unfold) simp
753ae9e9773d algebraized more material from theory Divides
haftmann
parents: 69593
diff changeset
  1737
qed
753ae9e9773d algebraized more material from theory Divides
haftmann
parents: 69593
diff changeset
  1738
753ae9e9773d algebraized more material from theory Divides
haftmann
parents: 69593
diff changeset
  1739
lemma neg_mod_sign [simp]:
753ae9e9773d algebraized more material from theory Divides
haftmann
parents: 69593
diff changeset
  1740
  "k mod l \<le> 0" if "l < 0" for k l :: int
753ae9e9773d algebraized more material from theory Divides
haftmann
parents: 69593
diff changeset
  1741
proof -
753ae9e9773d algebraized more material from theory Divides
haftmann
parents: 69593
diff changeset
  1742
  obtain m and s where "k = sgn s * int m"
753ae9e9773d algebraized more material from theory Divides
haftmann
parents: 69593
diff changeset
  1743
    by (rule int_sgnE)
753ae9e9773d algebraized more material from theory Divides
haftmann
parents: 69593
diff changeset
  1744
  moreover from that obtain q where "l = sgn (- 1) * int (Suc q)"
753ae9e9773d algebraized more material from theory Divides
haftmann
parents: 69593
diff changeset
  1745
    by (cases l) simp_all
753ae9e9773d algebraized more material from theory Divides
haftmann
parents: 69593
diff changeset
  1746
  moreover define n where "n = Suc q"
753ae9e9773d algebraized more material from theory Divides
haftmann
parents: 69593
diff changeset
  1747
  then have "Suc q = n"
753ae9e9773d algebraized more material from theory Divides
haftmann
parents: 69593
diff changeset
  1748
    by simp
753ae9e9773d algebraized more material from theory Divides
haftmann
parents: 69593
diff changeset
  1749
  ultimately show ?thesis
753ae9e9773d algebraized more material from theory Divides
haftmann
parents: 69593
diff changeset
  1750
    by (simp only: modulo_int_unfold) simp
66816
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1751
qed
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1752
72187
e4aecb0c7296 more lemmas
haftmann
parents: 71535
diff changeset
  1753
lemma div_pos_pos_trivial [simp]:
e4aecb0c7296 more lemmas
haftmann
parents: 71535
diff changeset
  1754
  "k div l = 0" if "k \<ge> 0" and "k < l" for k l :: int
e4aecb0c7296 more lemmas
haftmann
parents: 71535
diff changeset
  1755
  using that by (simp add: unique_euclidean_semiring_class.div_eq_0_iff division_segment_int_def)
e4aecb0c7296 more lemmas
haftmann
parents: 71535
diff changeset
  1756
e4aecb0c7296 more lemmas
haftmann
parents: 71535
diff changeset
  1757
lemma mod_pos_pos_trivial [simp]:
e4aecb0c7296 more lemmas
haftmann
parents: 71535
diff changeset
  1758
  "k mod l = k" if "k \<ge> 0" and "k < l" for k l :: int
e4aecb0c7296 more lemmas
haftmann
parents: 71535
diff changeset
  1759
  using that by (simp add: mod_eq_self_iff_div_eq_0)
e4aecb0c7296 more lemmas
haftmann
parents: 71535
diff changeset
  1760
e4aecb0c7296 more lemmas
haftmann
parents: 71535
diff changeset
  1761
lemma div_neg_neg_trivial [simp]:
e4aecb0c7296 more lemmas
haftmann
parents: 71535
diff changeset
  1762
  "k div l = 0" if "k \<le> 0" and "l < k" for k l :: int
e4aecb0c7296 more lemmas
haftmann
parents: 71535
diff changeset
  1763
  using that by (cases "k = 0") (simp, simp add: unique_euclidean_semiring_class.div_eq_0_iff division_segment_int_def)
e4aecb0c7296 more lemmas
haftmann
parents: 71535
diff changeset
  1764
e4aecb0c7296 more lemmas
haftmann
parents: 71535
diff changeset
  1765
lemma mod_neg_neg_trivial [simp]:
e4aecb0c7296 more lemmas
haftmann
parents: 71535
diff changeset
  1766
  "k mod l = k" if "k \<le> 0" and "l < k" for k l :: int
e4aecb0c7296 more lemmas
haftmann
parents: 71535
diff changeset
  1767
  using that by (simp add: mod_eq_self_iff_div_eq_0)
e4aecb0c7296 more lemmas
haftmann
parents: 71535
diff changeset
  1768
e4aecb0c7296 more lemmas
haftmann
parents: 71535
diff changeset
  1769
lemma div_pos_neg_trivial:
e4aecb0c7296 more lemmas
haftmann
parents: 71535
diff changeset
  1770
  "k div l = - 1" if "0 < k" and "k + l \<le> 0" for k l :: int
e4aecb0c7296 more lemmas
haftmann
parents: 71535
diff changeset
  1771
proof (cases \<open>l = - k\<close>)
e4aecb0c7296 more lemmas
haftmann
parents: 71535
diff changeset
  1772
  case True
e4aecb0c7296 more lemmas
haftmann
parents: 71535
diff changeset
  1773
  with that show ?thesis
e4aecb0c7296 more lemmas
haftmann
parents: 71535
diff changeset
  1774
    by (simp add: divide_int_def)
e4aecb0c7296 more lemmas
haftmann
parents: 71535
diff changeset
  1775
next
e4aecb0c7296 more lemmas
haftmann
parents: 71535
diff changeset
  1776
  case False
e4aecb0c7296 more lemmas
haftmann
parents: 71535
diff changeset
  1777
  show ?thesis
e4aecb0c7296 more lemmas
haftmann
parents: 71535
diff changeset
  1778
    apply (rule div_eqI [of _ "k + l"])
e4aecb0c7296 more lemmas
haftmann
parents: 71535
diff changeset
  1779
    using False that apply (simp_all add: division_segment_int_def)
e4aecb0c7296 more lemmas
haftmann
parents: 71535
diff changeset
  1780
    done
e4aecb0c7296 more lemmas
haftmann
parents: 71535
diff changeset
  1781
qed
e4aecb0c7296 more lemmas
haftmann
parents: 71535
diff changeset
  1782
e4aecb0c7296 more lemmas
haftmann
parents: 71535
diff changeset
  1783
lemma mod_pos_neg_trivial:
e4aecb0c7296 more lemmas
haftmann
parents: 71535
diff changeset
  1784
  "k mod l = k + l" if "0 < k" and "k + l \<le> 0" for k l :: int
e4aecb0c7296 more lemmas
haftmann
parents: 71535
diff changeset
  1785
proof (cases \<open>l = - k\<close>)
e4aecb0c7296 more lemmas
haftmann
parents: 71535
diff changeset
  1786
  case True
e4aecb0c7296 more lemmas
haftmann
parents: 71535
diff changeset
  1787
  with that show ?thesis
e4aecb0c7296 more lemmas
haftmann
parents: 71535
diff changeset
  1788
    by (simp add: divide_int_def)
e4aecb0c7296 more lemmas
haftmann
parents: 71535
diff changeset
  1789
next
e4aecb0c7296 more lemmas
haftmann
parents: 71535
diff changeset
  1790
  case False
e4aecb0c7296 more lemmas
haftmann
parents: 71535
diff changeset
  1791
  show ?thesis
e4aecb0c7296 more lemmas
haftmann
parents: 71535
diff changeset
  1792
    apply (rule mod_eqI [of _ _ \<open>- 1\<close>])
e4aecb0c7296 more lemmas
haftmann
parents: 71535
diff changeset
  1793
    using False that apply (simp_all add: division_segment_int_def)
e4aecb0c7296 more lemmas
haftmann
parents: 71535
diff changeset
  1794
    done
e4aecb0c7296 more lemmas
haftmann
parents: 71535
diff changeset
  1795
qed
e4aecb0c7296 more lemmas
haftmann
parents: 71535
diff changeset
  1796
e4aecb0c7296 more lemmas
haftmann
parents: 71535
diff changeset
  1797
text \<open>There is neither \<open>div_neg_pos_trivial\<close> nor \<open>mod_neg_pos_trivial\<close>
e4aecb0c7296 more lemmas
haftmann
parents: 71535
diff changeset
  1798
  because \<^term>\<open>0 div l = 0\<close> would supersede it.\<close>
e4aecb0c7296 more lemmas
haftmann
parents: 71535
diff changeset
  1799
e4aecb0c7296 more lemmas
haftmann
parents: 71535
diff changeset
  1800
text \<open>Distributive laws for function \<open>nat\<close>.\<close>
e4aecb0c7296 more lemmas
haftmann
parents: 71535
diff changeset
  1801
e4aecb0c7296 more lemmas
haftmann
parents: 71535
diff changeset
  1802
lemma nat_div_distrib:
e4aecb0c7296 more lemmas
haftmann
parents: 71535
diff changeset
  1803
  \<open>nat (x div y) = nat x div nat y\<close> if \<open>0 \<le> x\<close>
e4aecb0c7296 more lemmas
haftmann
parents: 71535
diff changeset
  1804
  using that by (simp add: divide_int_def sgn_if)
e4aecb0c7296 more lemmas
haftmann
parents: 71535
diff changeset
  1805
e4aecb0c7296 more lemmas
haftmann
parents: 71535
diff changeset
  1806
lemma nat_div_distrib':
e4aecb0c7296 more lemmas
haftmann
parents: 71535
diff changeset
  1807
  \<open>nat (x div y) = nat x div nat y\<close> if \<open>0 \<le> y\<close>
e4aecb0c7296 more lemmas
haftmann
parents: 71535
diff changeset
  1808
  using that by (simp add: divide_int_def sgn_if)
e4aecb0c7296 more lemmas
haftmann
parents: 71535
diff changeset
  1809
e4aecb0c7296 more lemmas
haftmann
parents: 71535
diff changeset
  1810
lemma nat_mod_distrib: \<comment> \<open>Fails if y<0: the LHS collapses to (nat z) but the RHS doesn't\<close>
e4aecb0c7296 more lemmas
haftmann
parents: 71535
diff changeset
  1811
  \<open>nat (x mod y) = nat x mod nat y\<close> if \<open>0 \<le> x\<close> \<open>0 \<le> y\<close>
e4aecb0c7296 more lemmas
haftmann
parents: 71535
diff changeset
  1812
  using that by (simp add: modulo_int_def sgn_if)
e4aecb0c7296 more lemmas
haftmann
parents: 71535
diff changeset
  1813
66816
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66814
diff changeset
  1814
71157
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1815
subsection \<open>Special case: euclidean rings containing the natural numbers\<close>
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1816
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1817
class unique_euclidean_semiring_with_nat = semidom + semiring_char_0 + unique_euclidean_semiring +
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1818
  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
  1819
    and division_segment_of_nat [simp]: "division_segment (of_nat n) = 1"
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1820
    and division_segment_euclidean_size [simp]: "division_segment a * of_nat (euclidean_size a) = a"
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1821
begin
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1822
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1823
lemma division_segment_eq_iff:
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1824
  "a = b" if "division_segment a = division_segment b"
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1825
    and "euclidean_size a = euclidean_size b"
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1826
  using that division_segment_euclidean_size [of a] by simp
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1827
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1828
lemma euclidean_size_of_nat [simp]:
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1829
  "euclidean_size (of_nat n) = n"
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1830
proof -
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1831
  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
  1832
    by (fact division_segment_euclidean_size)
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1833
  then show ?thesis by simp
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1834
qed
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1835
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1836
lemma of_nat_euclidean_size:
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1837
  "of_nat (euclidean_size a) = a div division_segment a"
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1838
proof -
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1839
  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
  1840
    by (subst nonzero_mult_div_cancel_left) simp_all
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1841
  also have "\<dots> = a div division_segment a"
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1842
    by simp
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1843
  finally show ?thesis .
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1844
qed
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1845
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1846
lemma division_segment_1 [simp]:
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1847
  "division_segment 1 = 1"
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1848
  using division_segment_of_nat [of 1] by simp
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1849
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1850
lemma division_segment_numeral [simp]:
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1851
  "division_segment (numeral k) = 1"
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1852
  using division_segment_of_nat [of "numeral k"] by simp
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1853
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1854
lemma euclidean_size_1 [simp]:
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1855
  "euclidean_size 1 = 1"
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1856
  using euclidean_size_of_nat [of 1] by simp
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1857
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1858
lemma euclidean_size_numeral [simp]:
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1859
  "euclidean_size (numeral k) = numeral k"
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1860
  using euclidean_size_of_nat [of "numeral k"] by simp
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1861
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1862
lemma of_nat_dvd_iff:
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1863
  "of_nat m dvd of_nat n \<longleftrightarrow> m dvd n" (is "?P \<longleftrightarrow> ?Q")
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1864
proof (cases "m = 0")
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1865
  case True
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1866
  then show ?thesis
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1867
    by simp
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1868
next
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1869
  case False
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1870
  show ?thesis
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1871
  proof
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1872
    assume ?Q
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1873
    then show ?P
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1874
      by auto
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1875
  next
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1876
    assume ?P
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1877
    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
  1878
      by simp
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1879
    then have "of_nat n = of_nat (n div m * m)"
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1880
      by (simp add: of_nat_div)
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1881
    then have "n = n div m * m"
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1882
      by (simp only: of_nat_eq_iff)
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1883
    then have "n = m * (n div m)"
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1884
      by (simp add: ac_simps)
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1885
    then show ?Q ..
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1886
  qed
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1887
qed
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1888
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1889
lemma of_nat_mod:
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1890
  "of_nat (m mod n) = of_nat m mod of_nat n"
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1891
proof -
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1892
  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
  1893
    by (simp add: div_mult_mod_eq)
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1894
  also have "of_nat m = of_nat (m div n * n + m mod n)"
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1895
    by simp
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1896
  finally show ?thesis
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1897
    by (simp only: of_nat_div of_nat_mult of_nat_add) simp
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1898
qed
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1899
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1900
lemma one_div_two_eq_zero [simp]:
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1901
  "1 div 2 = 0"
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1902
proof -
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1903
  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
  1904
    by (simp only:) simp
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1905
  then show ?thesis
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1906
    by simp
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1907
qed
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1908
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1909
lemma one_mod_two_eq_one [simp]:
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1910
  "1 mod 2 = 1"
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1911
proof -
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1912
  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
  1913
    by (simp only:) simp
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1914
  then show ?thesis
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1915
    by simp
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1916
qed
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1917
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1918
lemma one_mod_2_pow_eq [simp]:
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1919
  "1 mod (2 ^ n) = of_bool (n > 0)"
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1920
proof -
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1921
  have "1 mod (2 ^ n) = of_nat (1 mod (2 ^ n))"
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1922
    using of_nat_mod [of 1 "2 ^ n"] by simp
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1923
  also have "\<dots> = of_bool (n > 0)"
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1924
    by simp
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1925
  finally show ?thesis .
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1926
qed
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1927
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1928
lemma one_div_2_pow_eq [simp]:
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1929
  "1 div (2 ^ n) = of_bool (n = 0)"
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1930
  using div_mult_mod_eq [of 1 "2 ^ n"] by auto
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1931
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1932
lemma div_mult2_eq':
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1933
  "a div (of_nat m * of_nat n) = a div of_nat m div of_nat n"
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1934
proof (cases a "of_nat m * of_nat n" rule: divmod_cases)
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1935
  case (divides q)
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1936
  then show ?thesis
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1937
    using nonzero_mult_div_cancel_right [of "of_nat m" "q * of_nat n"]
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1938
    by (simp add: ac_simps)
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1939
next
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1940
  case (remainder q r)
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1941
  then have "division_segment r = 1"
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1942
    using division_segment_of_nat [of "m * n"] by simp
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1943
  with division_segment_euclidean_size [of r]
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1944
  have "of_nat (euclidean_size r) = r"
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1945
    by simp
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1946
  have "a mod (of_nat m * of_nat n) div (of_nat m * of_nat n) = 0"
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1947
    by simp
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1948
  with remainder(6) have "r div (of_nat m * of_nat n) = 0"
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1949
    by simp
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1950
  with \<open>of_nat (euclidean_size r) = r\<close>
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1951
  have "of_nat (euclidean_size r) div (of_nat m * of_nat n) = 0"
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1952
    by simp
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1953
  then have "of_nat (euclidean_size r div (m * n)) = 0"
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1954
    by (simp add: of_nat_div)
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1955
  then have "of_nat (euclidean_size r div m div n) = 0"
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1956
    by (simp add: div_mult2_eq)
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1957
  with \<open>of_nat (euclidean_size r) = r\<close> have "r div of_nat m div of_nat n = 0"
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1958
    by (simp add: of_nat_div)
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1959
  with remainder(1)
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1960
  have "q = (r div of_nat m + q * of_nat n * of_nat m div of_nat m) div of_nat n"
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1961
    by simp
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1962
  with remainder(5) remainder(7) show ?thesis
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1963
    using div_plus_div_distrib_dvd_right [of "of_nat m" "q * (of_nat m * of_nat n)" r]
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1964
    by (simp add: ac_simps)
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1965
next
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1966
  case by0
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1967
  then show ?thesis
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1968
    by auto
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1969
qed
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1970
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1971
lemma mod_mult2_eq':
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1972
  "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
  1973
proof -
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1974
  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
  1975
    by (simp add: combine_common_factor div_mult_mod_eq)
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1976
  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
  1977
    by (simp add: ac_simps)
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1978
  ultimately show ?thesis
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1979
    by (simp add: div_mult2_eq' mult_commute)
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1980
qed
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1981
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1982
lemma div_mult2_numeral_eq:
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1983
  "a div numeral k div numeral l = a div numeral (k * l)" (is "?A = ?B")
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1984
proof -
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1985
  have "?A = a div of_nat (numeral k) div of_nat (numeral l)"
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1986
    by simp
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1987
  also have "\<dots> = a div (of_nat (numeral k) * of_nat (numeral l))"
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1988
    by (fact div_mult2_eq' [symmetric])
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1989
  also have "\<dots> = ?B"
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1990
    by simp
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1991
  finally show ?thesis .
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1992
qed
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1993
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1994
lemma numeral_Bit0_div_2:
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1995
  "numeral (num.Bit0 n) div 2 = numeral n"
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1996
proof -
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1997
  have "numeral (num.Bit0 n) = numeral n + numeral n"
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1998
    by (simp only: numeral.simps)
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  1999
  also have "\<dots> = numeral n * 2"
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2000
    by (simp add: mult_2_right)
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2001
  finally have "numeral (num.Bit0 n) div 2 = numeral n * 2 div 2"
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2002
    by simp
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2003
  also have "\<dots> = numeral n"
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2004
    by (rule nonzero_mult_div_cancel_right) simp
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2005
  finally show ?thesis .
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2006
qed
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2007
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2008
lemma numeral_Bit1_div_2:
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2009
  "numeral (num.Bit1 n) div 2 = numeral n"
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2010
proof -
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2011
  have "numeral (num.Bit1 n) = numeral n + numeral n + 1"
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2012
    by (simp only: numeral.simps)
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2013
  also have "\<dots> = numeral n * 2 + 1"
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2014
    by (simp add: mult_2_right)
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2015
  finally have "numeral (num.Bit1 n) div 2 = (numeral n * 2 + 1) div 2"
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2016
    by simp
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2017
  also have "\<dots> = numeral n * 2 div 2 + 1 div 2"
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2018
    using dvd_triv_right by (rule div_plus_div_distrib_dvd_left)
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2019
  also have "\<dots> = numeral n * 2 div 2"
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2020
    by simp
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2021
  also have "\<dots> = numeral n"
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2022
    by (rule nonzero_mult_div_cancel_right) simp
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2023
  finally show ?thesis .
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2024
qed
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2025
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2026
lemma exp_mod_exp:
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2027
  \<open>2 ^ m mod 2 ^ n = of_bool (m < n) * 2 ^ m\<close>
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2028
proof -
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2029
  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
  2030
    by (auto simp add: not_less monoid_mult_class.power_add dest!: le_Suc_ex)
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2031
  then have \<open>of_nat ?lhs = of_nat ?rhs\<close>
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2032
    by simp
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2033
  then show ?thesis
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2034
    by (simp add: of_nat_mod)
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2035
qed
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2036
71412
96d126844adc more theorems
haftmann
parents: 71408
diff changeset
  2037
lemma mask_mod_exp:
71408
554385d4cf59 more theorems
haftmann
parents: 71157
diff changeset
  2038
  \<open>(2 ^ n - 1) mod 2 ^ m = 2 ^ min m n - 1\<close>
554385d4cf59 more theorems
haftmann
parents: 71157
diff changeset
  2039
proof -
554385d4cf59 more theorems
haftmann
parents: 71157
diff changeset
  2040
  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
  2041
  proof (cases \<open>n \<le> m\<close>)
554385d4cf59 more theorems
haftmann
parents: 71157
diff changeset
  2042
    case True
554385d4cf59 more theorems
haftmann
parents: 71157
diff changeset
  2043
    then show ?thesis
554385d4cf59 more theorems
haftmann
parents: 71157
diff changeset
  2044
      by (simp add: Suc_le_lessD min.absorb2)
554385d4cf59 more theorems
haftmann
parents: 71157
diff changeset
  2045
  next
554385d4cf59 more theorems
haftmann
parents: 71157
diff changeset
  2046
    case False
554385d4cf59 more theorems
haftmann
parents: 71157
diff changeset
  2047
    then have \<open>m < n\<close>
554385d4cf59 more theorems
haftmann
parents: 71157
diff changeset
  2048
      by simp
554385d4cf59 more theorems
haftmann
parents: 71157
diff changeset
  2049
    then obtain q where n: \<open>n = Suc q + m\<close>
554385d4cf59 more theorems
haftmann
parents: 71157
diff changeset
  2050
      by (auto dest: less_imp_Suc_add)
554385d4cf59 more theorems
haftmann
parents: 71157
diff changeset
  2051
    then have \<open>min m n = m\<close>
554385d4cf59 more theorems
haftmann
parents: 71157
diff changeset
  2052
      by simp
554385d4cf59 more theorems
haftmann
parents: 71157
diff changeset
  2053
    moreover have \<open>(2::nat) ^ m \<le> 2 * 2 ^ q * 2 ^ m\<close>
554385d4cf59 more theorems
haftmann
parents: 71157
diff changeset
  2054
      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
  2055
    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
  2056
      by (simp add: monoid_mult_class.power_add algebra_simps)
554385d4cf59 more theorems
haftmann
parents: 71157
diff changeset
  2057
    ultimately show ?thesis
554385d4cf59 more theorems
haftmann
parents: 71157
diff changeset
  2058
      by (simp only: euclidean_semiring_cancel_class.mod_mult_self3) simp
554385d4cf59 more theorems
haftmann
parents: 71157
diff changeset
  2059
  qed
554385d4cf59 more theorems
haftmann
parents: 71157
diff changeset
  2060
  then have \<open>of_nat ?lhs = of_nat ?rhs\<close>
554385d4cf59 more theorems
haftmann
parents: 71157
diff changeset
  2061
    by simp
554385d4cf59 more theorems
haftmann
parents: 71157
diff changeset
  2062
  then show ?thesis
554385d4cf59 more theorems
haftmann
parents: 71157
diff changeset
  2063
    by (simp add: of_nat_mod of_nat_diff)
554385d4cf59 more theorems
haftmann
parents: 71157
diff changeset
  2064
qed
554385d4cf59 more theorems
haftmann
parents: 71157
diff changeset
  2065
71535
b612edee9b0c more frugal simp rules for bit operations; more pervasive use of bit selector
haftmann
parents: 71413
diff changeset
  2066
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
  2067
  \<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
  2068
  by simp
b612edee9b0c more frugal simp rules for bit operations; more pervasive use of bit selector
haftmann
parents: 71413
diff changeset
  2069
71157
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2070
end
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2071
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2072
class unique_euclidean_ring_with_nat = ring + unique_euclidean_semiring_with_nat
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2073
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2074
instance nat :: unique_euclidean_semiring_with_nat
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2075
  by standard (simp_all add: dvd_eq_mod_eq_0)
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2076
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2077
instance int :: unique_euclidean_ring_with_nat
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2078
  by standard (simp_all add: dvd_eq_mod_eq_0 divide_int_def division_segment_int_def)
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2079
8bdf3c36011c tuned theory structure
haftmann
parents: 70147
diff changeset
  2080
66808
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  2081
subsection \<open>Code generation\<close>
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  2082
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  2083
code_identifier
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  2084
  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
  2085
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  2086
end