src/HOL/Number_Theory/Cong.thy
author wenzelm
Wed, 17 Apr 2024 23:22:32 +0200
changeset 80132 ef2134570abb
parent 80084 173548e4d5d0
permissions -rw-r--r--
merged
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
41959
b460124855b8 tuned headers;
wenzelm
parents: 41541
diff changeset
     1
(*  Title:      HOL/Number_Theory/Cong.thy
66380
96ff0eb8294a misc tuning and modernization;
wenzelm
parents: 65417
diff changeset
     2
    Author:     Christophe Tabacznyj
96ff0eb8294a misc tuning and modernization;
wenzelm
parents: 65417
diff changeset
     3
    Author:     Lawrence C. Paulson
96ff0eb8294a misc tuning and modernization;
wenzelm
parents: 65417
diff changeset
     4
    Author:     Amine Chaieb
96ff0eb8294a misc tuning and modernization;
wenzelm
parents: 65417
diff changeset
     5
    Author:     Thomas M. Rasmussen
96ff0eb8294a misc tuning and modernization;
wenzelm
parents: 65417
diff changeset
     6
    Author:     Jeremy Avigad
31719
29f5b20e8ee8 Added NewNumberTheory by Jeremy Avigad
nipkow
parents:
diff changeset
     7
29f5b20e8ee8 Added NewNumberTheory by Jeremy Avigad
nipkow
parents:
diff changeset
     8
Defines congruence (notation: [x = y] (mod z)) for natural numbers and
29f5b20e8ee8 Added NewNumberTheory by Jeremy Avigad
nipkow
parents:
diff changeset
     9
integers.
29f5b20e8ee8 Added NewNumberTheory by Jeremy Avigad
nipkow
parents:
diff changeset
    10
29f5b20e8ee8 Added NewNumberTheory by Jeremy Avigad
nipkow
parents:
diff changeset
    11
This file combines and revises a number of prior developments.
29f5b20e8ee8 Added NewNumberTheory by Jeremy Avigad
nipkow
parents:
diff changeset
    12
29f5b20e8ee8 Added NewNumberTheory by Jeremy Avigad
nipkow
parents:
diff changeset
    13
The original theories "GCD" and "Primes" were by Christophe Tabacznyj
58623
2db1df2c8467 more bibtex entries;
wenzelm
parents: 57512
diff changeset
    14
and Lawrence C. Paulson, based on @{cite davenport92}. They introduced
31719
29f5b20e8ee8 Added NewNumberTheory by Jeremy Avigad
nipkow
parents:
diff changeset
    15
gcd, lcm, and prime for the natural numbers.
29f5b20e8ee8 Added NewNumberTheory by Jeremy Avigad
nipkow
parents:
diff changeset
    16
29f5b20e8ee8 Added NewNumberTheory by Jeremy Avigad
nipkow
parents:
diff changeset
    17
The original theory "IntPrimes" was by Thomas M. Rasmussen, and
29f5b20e8ee8 Added NewNumberTheory by Jeremy Avigad
nipkow
parents:
diff changeset
    18
extended gcd, lcm, primes to the integers. Amine Chaieb provided
29f5b20e8ee8 Added NewNumberTheory by Jeremy Avigad
nipkow
parents:
diff changeset
    19
another extension of the notions to the integers, and added a number
44872
a98ef45122f3 misc tuning;
wenzelm
parents: 41959
diff changeset
    20
of results to "Primes" and "GCD".
31719
29f5b20e8ee8 Added NewNumberTheory by Jeremy Avigad
nipkow
parents:
diff changeset
    21
29f5b20e8ee8 Added NewNumberTheory by Jeremy Avigad
nipkow
parents:
diff changeset
    22
The original theory, "IntPrimes", by Thomas M. Rasmussen, defined and
29f5b20e8ee8 Added NewNumberTheory by Jeremy Avigad
nipkow
parents:
diff changeset
    23
developed the congruence relations on the integers. The notion was
33718
06e9aff51d17 Fixed a typo in a comment.
webertj
parents: 32479
diff changeset
    24
extended to the natural numbers by Chaieb. Jeremy Avigad combined
31719
29f5b20e8ee8 Added NewNumberTheory by Jeremy Avigad
nipkow
parents:
diff changeset
    25
these, revised and tidied them, made the development uniform for the
29f5b20e8ee8 Added NewNumberTheory by Jeremy Avigad
nipkow
parents:
diff changeset
    26
natural numbers and the integers, and added a number of new theorems.
29f5b20e8ee8 Added NewNumberTheory by Jeremy Avigad
nipkow
parents:
diff changeset
    27
*)
29f5b20e8ee8 Added NewNumberTheory by Jeremy Avigad
nipkow
parents:
diff changeset
    28
60526
fad653acf58f isabelle update_cartouches;
wenzelm
parents: 59816
diff changeset
    29
section \<open>Congruence\<close>
31719
29f5b20e8ee8 Added NewNumberTheory by Jeremy Avigad
nipkow
parents:
diff changeset
    30
29f5b20e8ee8 Added NewNumberTheory by Jeremy Avigad
nipkow
parents:
diff changeset
    31
theory Cong
66453
cc19f7ca2ed6 session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
wenzelm
parents: 66380
diff changeset
    32
  imports "HOL-Computational_Algebra.Primes"
31719
29f5b20e8ee8 Added NewNumberTheory by Jeremy Avigad
nipkow
parents:
diff changeset
    33
begin
29f5b20e8ee8 Added NewNumberTheory by Jeremy Avigad
nipkow
parents:
diff changeset
    34
66888
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66837
diff changeset
    35
subsection \<open>Generic congruences\<close>
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66837
diff changeset
    36
 
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66837
diff changeset
    37
context unique_euclidean_semiring
31719
29f5b20e8ee8 Added NewNumberTheory by Jeremy Avigad
nipkow
parents:
diff changeset
    38
begin
29f5b20e8ee8 Added NewNumberTheory by Jeremy Avigad
nipkow
parents:
diff changeset
    39
71546
4dd5dadfc87d some uses of "' " as witness for this feature;
wenzelm
parents: 69597
diff changeset
    40
definition cong :: "'a \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> bool"  (\<open>(1[_ = _] '(' mod _'))\<close>)
66888
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66837
diff changeset
    41
  where "cong b c a \<longleftrightarrow> b mod a = c mod a"
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66837
diff changeset
    42
  
71546
4dd5dadfc87d some uses of "' " as witness for this feature;
wenzelm
parents: 69597
diff changeset
    43
abbreviation notcong :: "'a \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> bool"  (\<open>(1[_ \<noteq> _] '(' mod _'))\<close>)
66888
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66837
diff changeset
    44
  where "notcong b c a \<equiv> \<not> cong b c a"
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66837
diff changeset
    45
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66837
diff changeset
    46
lemma cong_refl [simp]:
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66837
diff changeset
    47
  "[b = b] (mod a)"
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66837
diff changeset
    48
  by (simp add: cong_def)
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66837
diff changeset
    49
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66837
diff changeset
    50
lemma cong_sym: 
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66837
diff changeset
    51
  "[b = c] (mod a) \<Longrightarrow> [c = b] (mod a)"
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66837
diff changeset
    52
  by (simp add: cong_def)
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66837
diff changeset
    53
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66837
diff changeset
    54
lemma cong_sym_eq:
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66837
diff changeset
    55
  "[b = c] (mod a) \<longleftrightarrow> [c = b] (mod a)"
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66837
diff changeset
    56
  by (auto simp add: cong_def)
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66837
diff changeset
    57
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66837
diff changeset
    58
lemma cong_trans [trans]:
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66837
diff changeset
    59
  "[b = c] (mod a) \<Longrightarrow> [c = d] (mod a) \<Longrightarrow> [b = d] (mod a)"
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66837
diff changeset
    60
  by (simp add: cong_def)
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66837
diff changeset
    61
67085
f5d7f37b4143 tuned and generalized
haftmann
parents: 67051
diff changeset
    62
lemma cong_mult_self_right:
f5d7f37b4143 tuned and generalized
haftmann
parents: 67051
diff changeset
    63
  "[b * a = 0] (mod a)"
f5d7f37b4143 tuned and generalized
haftmann
parents: 67051
diff changeset
    64
  by (simp add: cong_def)
f5d7f37b4143 tuned and generalized
haftmann
parents: 67051
diff changeset
    65
f5d7f37b4143 tuned and generalized
haftmann
parents: 67051
diff changeset
    66
lemma cong_mult_self_left:
f5d7f37b4143 tuned and generalized
haftmann
parents: 67051
diff changeset
    67
  "[a * b = 0] (mod a)"
f5d7f37b4143 tuned and generalized
haftmann
parents: 67051
diff changeset
    68
  by (simp add: cong_def)
f5d7f37b4143 tuned and generalized
haftmann
parents: 67051
diff changeset
    69
f5d7f37b4143 tuned and generalized
haftmann
parents: 67051
diff changeset
    70
lemma cong_mod_left [simp]:
f5d7f37b4143 tuned and generalized
haftmann
parents: 67051
diff changeset
    71
  "[b mod a = c] (mod a) \<longleftrightarrow> [b = c] (mod a)"
f5d7f37b4143 tuned and generalized
haftmann
parents: 67051
diff changeset
    72
  by (simp add: cong_def)  
f5d7f37b4143 tuned and generalized
haftmann
parents: 67051
diff changeset
    73
f5d7f37b4143 tuned and generalized
haftmann
parents: 67051
diff changeset
    74
lemma cong_mod_right [simp]:
f5d7f37b4143 tuned and generalized
haftmann
parents: 67051
diff changeset
    75
  "[b = c mod a] (mod a) \<longleftrightarrow> [b = c] (mod a)"
f5d7f37b4143 tuned and generalized
haftmann
parents: 67051
diff changeset
    76
  by (simp add: cong_def)  
f5d7f37b4143 tuned and generalized
haftmann
parents: 67051
diff changeset
    77
f5d7f37b4143 tuned and generalized
haftmann
parents: 67051
diff changeset
    78
lemma cong_0 [simp, presburger]:
f5d7f37b4143 tuned and generalized
haftmann
parents: 67051
diff changeset
    79
  "[b = c] (mod 0) \<longleftrightarrow> b = c"
f5d7f37b4143 tuned and generalized
haftmann
parents: 67051
diff changeset
    80
  by (simp add: cong_def)
f5d7f37b4143 tuned and generalized
haftmann
parents: 67051
diff changeset
    81
f5d7f37b4143 tuned and generalized
haftmann
parents: 67051
diff changeset
    82
lemma cong_1 [simp, presburger]:
f5d7f37b4143 tuned and generalized
haftmann
parents: 67051
diff changeset
    83
  "[b = c] (mod 1)"
f5d7f37b4143 tuned and generalized
haftmann
parents: 67051
diff changeset
    84
  by (simp add: cong_def)
f5d7f37b4143 tuned and generalized
haftmann
parents: 67051
diff changeset
    85
f5d7f37b4143 tuned and generalized
haftmann
parents: 67051
diff changeset
    86
lemma cong_dvd_iff:
f5d7f37b4143 tuned and generalized
haftmann
parents: 67051
diff changeset
    87
  "a dvd b \<longleftrightarrow> a dvd c" if "[b = c] (mod a)"
f5d7f37b4143 tuned and generalized
haftmann
parents: 67051
diff changeset
    88
  using that by (auto simp: cong_def dvd_eq_mod_eq_0)
f5d7f37b4143 tuned and generalized
haftmann
parents: 67051
diff changeset
    89
f5d7f37b4143 tuned and generalized
haftmann
parents: 67051
diff changeset
    90
lemma cong_0_iff: "[b = 0] (mod a) \<longleftrightarrow> a dvd b"
f5d7f37b4143 tuned and generalized
haftmann
parents: 67051
diff changeset
    91
  by (simp add: cong_def dvd_eq_mod_eq_0)
f5d7f37b4143 tuned and generalized
haftmann
parents: 67051
diff changeset
    92
66888
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66837
diff changeset
    93
lemma cong_add:
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66837
diff changeset
    94
  "[b = c] (mod a) \<Longrightarrow> [d = e] (mod a) \<Longrightarrow> [b + d = c + e] (mod a)"
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66837
diff changeset
    95
  by (auto simp add: cong_def intro: mod_add_cong)
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66837
diff changeset
    96
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66837
diff changeset
    97
lemma cong_mult:
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66837
diff changeset
    98
  "[b = c] (mod a) \<Longrightarrow> [d = e] (mod a) \<Longrightarrow> [b * d = c * e] (mod a)"
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66837
diff changeset
    99
  by (auto simp add: cong_def intro: mod_mult_cong)
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66837
diff changeset
   100
67085
f5d7f37b4143 tuned and generalized
haftmann
parents: 67051
diff changeset
   101
lemma cong_scalar_right:
f5d7f37b4143 tuned and generalized
haftmann
parents: 67051
diff changeset
   102
  "[b = c] (mod a) \<Longrightarrow> [b * d = c * d] (mod a)"
f5d7f37b4143 tuned and generalized
haftmann
parents: 67051
diff changeset
   103
  by (simp add: cong_mult)
f5d7f37b4143 tuned and generalized
haftmann
parents: 67051
diff changeset
   104
f5d7f37b4143 tuned and generalized
haftmann
parents: 67051
diff changeset
   105
lemma cong_scalar_left:
f5d7f37b4143 tuned and generalized
haftmann
parents: 67051
diff changeset
   106
  "[b = c] (mod a) \<Longrightarrow> [d * b = d * c] (mod a)"
f5d7f37b4143 tuned and generalized
haftmann
parents: 67051
diff changeset
   107
  by (simp add: cong_mult)
f5d7f37b4143 tuned and generalized
haftmann
parents: 67051
diff changeset
   108
66888
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66837
diff changeset
   109
lemma cong_pow:
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66837
diff changeset
   110
  "[b = c] (mod a) \<Longrightarrow> [b ^ n = c ^ n] (mod a)"
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66837
diff changeset
   111
  by (simp add: cong_def power_mod [symmetric, of b n a] power_mod [symmetric, of c n a])
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66837
diff changeset
   112
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66837
diff changeset
   113
lemma cong_sum:
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66837
diff changeset
   114
  "[sum f A = sum g A] (mod a)" if "\<And>x. x \<in> A \<Longrightarrow> [f x = g x] (mod a)"
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66837
diff changeset
   115
  using that by (induct A rule: infinite_finite_induct) (auto intro: cong_add)
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66837
diff changeset
   116
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66837
diff changeset
   117
lemma cong_prod:
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66837
diff changeset
   118
  "[prod f A = prod g A] (mod a)" if "(\<And>x. x \<in> A \<Longrightarrow> [f x = g x] (mod a))"
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66837
diff changeset
   119
  using that by (induct A rule: infinite_finite_induct) (auto intro: cong_mult)
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66837
diff changeset
   120
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66837
diff changeset
   121
lemma mod_mult_cong_right:
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66837
diff changeset
   122
  "[c mod (a * b) = d] (mod a) \<longleftrightarrow> [c = d] (mod a)"
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66837
diff changeset
   123
  by (simp add: cong_def mod_mod_cancel mod_add_left_eq)
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66837
diff changeset
   124
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66837
diff changeset
   125
lemma mod_mult_cong_left:
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66837
diff changeset
   126
  "[c mod (b * a) = d] (mod a) \<longleftrightarrow> [c = d] (mod a)"
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66837
diff changeset
   127
  using mod_mult_cong_right [of c a b d] by (simp add: ac_simps)
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66837
diff changeset
   128
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66837
diff changeset
   129
end
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66837
diff changeset
   130
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66837
diff changeset
   131
context unique_euclidean_ring
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66837
diff changeset
   132
begin
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66837
diff changeset
   133
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66837
diff changeset
   134
lemma cong_diff:
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66837
diff changeset
   135
  "[b = c] (mod a) \<Longrightarrow> [d = e] (mod a) \<Longrightarrow> [b - d = c - e] (mod a)"
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66837
diff changeset
   136
  by (auto simp add: cong_def intro: mod_diff_cong)
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66837
diff changeset
   137
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66837
diff changeset
   138
lemma cong_diff_iff_cong_0:
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66837
diff changeset
   139
  "[b - c = 0] (mod a) \<longleftrightarrow> [b = c] (mod a)" (is "?P \<longleftrightarrow> ?Q")
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66837
diff changeset
   140
proof
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66837
diff changeset
   141
  assume ?P
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66837
diff changeset
   142
  then have "[b - c + c = 0 + c] (mod a)"
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66837
diff changeset
   143
    by (rule cong_add) simp
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66837
diff changeset
   144
  then show ?Q
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66837
diff changeset
   145
    by simp
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66837
diff changeset
   146
next
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66837
diff changeset
   147
  assume ?Q
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66837
diff changeset
   148
  with cong_diff [of b c a c c] show ?P
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66837
diff changeset
   149
    by simp
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66837
diff changeset
   150
qed
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66837
diff changeset
   151
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66837
diff changeset
   152
lemma cong_minus_minus_iff:
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66837
diff changeset
   153
  "[- b = - c] (mod a) \<longleftrightarrow> [b = c] (mod a)"
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66837
diff changeset
   154
  using cong_diff_iff_cong_0 [of b c a] cong_diff_iff_cong_0 [of "- b" "- c" a]
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66837
diff changeset
   155
  by (simp add: cong_0_iff dvd_diff_commute)
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66837
diff changeset
   156
67115
2977773a481c generalized more lemmas
haftmann
parents: 67087
diff changeset
   157
lemma cong_modulus_minus_iff [iff]:
67087
733017b19de9 generalized more lemmas
haftmann
parents: 67086
diff changeset
   158
  "[b = c] (mod - a) \<longleftrightarrow> [b = c] (mod a)"
66888
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66837
diff changeset
   159
  using cong_diff_iff_cong_0 [of b c a] cong_diff_iff_cong_0 [of b c " -a"]
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66837
diff changeset
   160
  by (simp add: cong_0_iff)
31719
29f5b20e8ee8 Added NewNumberTheory by Jeremy Avigad
nipkow
parents:
diff changeset
   161
67087
733017b19de9 generalized more lemmas
haftmann
parents: 67086
diff changeset
   162
lemma cong_iff_dvd_diff:
733017b19de9 generalized more lemmas
haftmann
parents: 67086
diff changeset
   163
  "[a = b] (mod m) \<longleftrightarrow> m dvd (a - b)"
733017b19de9 generalized more lemmas
haftmann
parents: 67086
diff changeset
   164
  by (simp add: cong_0_iff [symmetric] cong_diff_iff_cong_0)
733017b19de9 generalized more lemmas
haftmann
parents: 67086
diff changeset
   165
733017b19de9 generalized more lemmas
haftmann
parents: 67086
diff changeset
   166
lemma cong_iff_lin:
733017b19de9 generalized more lemmas
haftmann
parents: 67086
diff changeset
   167
  "[a = b] (mod m) \<longleftrightarrow> (\<exists>k. b = a + m * k)" (is "?P \<longleftrightarrow> ?Q")
733017b19de9 generalized more lemmas
haftmann
parents: 67086
diff changeset
   168
proof -
733017b19de9 generalized more lemmas
haftmann
parents: 67086
diff changeset
   169
  have "?P \<longleftrightarrow> m dvd b - a"
733017b19de9 generalized more lemmas
haftmann
parents: 67086
diff changeset
   170
    by (simp add: cong_iff_dvd_diff dvd_diff_commute)
733017b19de9 generalized more lemmas
haftmann
parents: 67086
diff changeset
   171
  also have "\<dots> \<longleftrightarrow> ?Q"
733017b19de9 generalized more lemmas
haftmann
parents: 67086
diff changeset
   172
    by (auto simp add: algebra_simps elim!: dvdE)
733017b19de9 generalized more lemmas
haftmann
parents: 67086
diff changeset
   173
  finally show ?thesis
733017b19de9 generalized more lemmas
haftmann
parents: 67086
diff changeset
   174
    by simp
733017b19de9 generalized more lemmas
haftmann
parents: 67086
diff changeset
   175
qed
733017b19de9 generalized more lemmas
haftmann
parents: 67086
diff changeset
   176
67115
2977773a481c generalized more lemmas
haftmann
parents: 67087
diff changeset
   177
lemma cong_add_lcancel:
2977773a481c generalized more lemmas
haftmann
parents: 67087
diff changeset
   178
  "[a + x = a + y] (mod n) \<longleftrightarrow> [x = y] (mod n)"
2977773a481c generalized more lemmas
haftmann
parents: 67087
diff changeset
   179
  by (simp add: cong_iff_lin algebra_simps)
2977773a481c generalized more lemmas
haftmann
parents: 67087
diff changeset
   180
2977773a481c generalized more lemmas
haftmann
parents: 67087
diff changeset
   181
lemma cong_add_rcancel:
2977773a481c generalized more lemmas
haftmann
parents: 67087
diff changeset
   182
  "[x + a = y + a] (mod n) \<longleftrightarrow> [x = y] (mod n)"
2977773a481c generalized more lemmas
haftmann
parents: 67087
diff changeset
   183
  by (simp add: cong_iff_lin algebra_simps)
2977773a481c generalized more lemmas
haftmann
parents: 67087
diff changeset
   184
2977773a481c generalized more lemmas
haftmann
parents: 67087
diff changeset
   185
lemma cong_add_lcancel_0:
2977773a481c generalized more lemmas
haftmann
parents: 67087
diff changeset
   186
  "[a + x = a] (mod n) \<longleftrightarrow> [x = 0] (mod n)"
2977773a481c generalized more lemmas
haftmann
parents: 67087
diff changeset
   187
  using cong_add_lcancel [of a x 0 n] by simp
2977773a481c generalized more lemmas
haftmann
parents: 67087
diff changeset
   188
2977773a481c generalized more lemmas
haftmann
parents: 67087
diff changeset
   189
lemma cong_add_rcancel_0:
2977773a481c generalized more lemmas
haftmann
parents: 67087
diff changeset
   190
  "[x + a = a] (mod n) \<longleftrightarrow> [x = 0] (mod n)"
2977773a481c generalized more lemmas
haftmann
parents: 67087
diff changeset
   191
  using cong_add_rcancel [of x a 0 n] by simp
2977773a481c generalized more lemmas
haftmann
parents: 67087
diff changeset
   192
2977773a481c generalized more lemmas
haftmann
parents: 67087
diff changeset
   193
lemma cong_dvd_modulus:
2977773a481c generalized more lemmas
haftmann
parents: 67087
diff changeset
   194
  "[x = y] (mod n)" if "[x = y] (mod m)" and "n dvd m"
2977773a481c generalized more lemmas
haftmann
parents: 67087
diff changeset
   195
  using that by (auto intro: dvd_trans simp add: cong_iff_dvd_diff)
2977773a481c generalized more lemmas
haftmann
parents: 67087
diff changeset
   196
2977773a481c generalized more lemmas
haftmann
parents: 67087
diff changeset
   197
lemma cong_modulus_mult:
2977773a481c generalized more lemmas
haftmann
parents: 67087
diff changeset
   198
  "[x = y] (mod m)" if "[x = y] (mod m * n)"
2977773a481c generalized more lemmas
haftmann
parents: 67087
diff changeset
   199
  using that by (simp add: cong_iff_dvd_diff) (rule dvd_mult_left)
2977773a481c generalized more lemmas
haftmann
parents: 67087
diff changeset
   200
31719
29f5b20e8ee8 Added NewNumberTheory by Jeremy Avigad
nipkow
parents:
diff changeset
   201
end
29f5b20e8ee8 Added NewNumberTheory by Jeremy Avigad
nipkow
parents:
diff changeset
   202
67115
2977773a481c generalized more lemmas
haftmann
parents: 67087
diff changeset
   203
lemma cong_abs [simp]:
2977773a481c generalized more lemmas
haftmann
parents: 67087
diff changeset
   204
  "[x = y] (mod \<bar>m\<bar>) \<longleftrightarrow> [x = y] (mod m)"
2977773a481c generalized more lemmas
haftmann
parents: 67087
diff changeset
   205
  for x y :: "'a :: {unique_euclidean_ring, linordered_idom}"
2977773a481c generalized more lemmas
haftmann
parents: 67087
diff changeset
   206
  by (simp add: cong_iff_dvd_diff)
2977773a481c generalized more lemmas
haftmann
parents: 67087
diff changeset
   207
2977773a481c generalized more lemmas
haftmann
parents: 67087
diff changeset
   208
lemma cong_square:
2977773a481c generalized more lemmas
haftmann
parents: 67087
diff changeset
   209
  "prime p \<Longrightarrow> 0 < a \<Longrightarrow> [a * a = 1] (mod p) \<Longrightarrow> [a = 1] (mod p) \<or> [a = - 1] (mod p)"
2977773a481c generalized more lemmas
haftmann
parents: 67087
diff changeset
   210
  for a p :: "'a :: {normalization_semidom, linordered_idom, unique_euclidean_ring}"
2977773a481c generalized more lemmas
haftmann
parents: 67087
diff changeset
   211
  by (auto simp add: cong_iff_dvd_diff square_diff_one_factored dest: prime_dvd_multD)
2977773a481c generalized more lemmas
haftmann
parents: 67087
diff changeset
   212
2977773a481c generalized more lemmas
haftmann
parents: 67087
diff changeset
   213
lemma cong_mult_rcancel:
2977773a481c generalized more lemmas
haftmann
parents: 67087
diff changeset
   214
  "[a * k = b * k] (mod m) \<longleftrightarrow> [a = b] (mod m)"
2977773a481c generalized more lemmas
haftmann
parents: 67087
diff changeset
   215
  if "coprime k m" for a k m :: "'a::{unique_euclidean_ring, ring_gcd}"
2977773a481c generalized more lemmas
haftmann
parents: 67087
diff changeset
   216
  using that by (auto simp add: cong_iff_dvd_diff left_diff_distrib [symmetric] ac_simps coprime_dvd_mult_right_iff)
2977773a481c generalized more lemmas
haftmann
parents: 67087
diff changeset
   217
2977773a481c generalized more lemmas
haftmann
parents: 67087
diff changeset
   218
lemma cong_mult_lcancel:
2977773a481c generalized more lemmas
haftmann
parents: 67087
diff changeset
   219
  "[k * a = k * b] (mod m) = [a = b] (mod m)"
2977773a481c generalized more lemmas
haftmann
parents: 67087
diff changeset
   220
  if "coprime k m" for a k m :: "'a::{unique_euclidean_ring, ring_gcd}"
2977773a481c generalized more lemmas
haftmann
parents: 67087
diff changeset
   221
  using that cong_mult_rcancel [of k m a b] by (simp add: ac_simps)
2977773a481c generalized more lemmas
haftmann
parents: 67087
diff changeset
   222
2977773a481c generalized more lemmas
haftmann
parents: 67087
diff changeset
   223
lemma coprime_cong_mult:
2977773a481c generalized more lemmas
haftmann
parents: 67087
diff changeset
   224
  "[a = b] (mod m) \<Longrightarrow> [a = b] (mod n) \<Longrightarrow> coprime m n \<Longrightarrow> [a = b] (mod m * n)"
2977773a481c generalized more lemmas
haftmann
parents: 67087
diff changeset
   225
  for a b :: "'a :: {unique_euclidean_ring, semiring_gcd}"
2977773a481c generalized more lemmas
haftmann
parents: 67087
diff changeset
   226
  by (simp add: cong_iff_dvd_diff divides_mult)
2977773a481c generalized more lemmas
haftmann
parents: 67087
diff changeset
   227
2977773a481c generalized more lemmas
haftmann
parents: 67087
diff changeset
   228
lemma cong_gcd_eq:
2977773a481c generalized more lemmas
haftmann
parents: 67087
diff changeset
   229
  "gcd a m = gcd b m" if "[a = b] (mod m)"
2977773a481c generalized more lemmas
haftmann
parents: 67087
diff changeset
   230
  for a b :: "'a :: {unique_euclidean_semiring, euclidean_semiring_gcd}"
2977773a481c generalized more lemmas
haftmann
parents: 67087
diff changeset
   231
proof (cases "m = 0")
2977773a481c generalized more lemmas
haftmann
parents: 67087
diff changeset
   232
  case True
2977773a481c generalized more lemmas
haftmann
parents: 67087
diff changeset
   233
  with that show ?thesis
2977773a481c generalized more lemmas
haftmann
parents: 67087
diff changeset
   234
    by simp
2977773a481c generalized more lemmas
haftmann
parents: 67087
diff changeset
   235
next
2977773a481c generalized more lemmas
haftmann
parents: 67087
diff changeset
   236
  case False
2977773a481c generalized more lemmas
haftmann
parents: 67087
diff changeset
   237
  moreover have "gcd (a mod m) m = gcd (b mod m) m"
2977773a481c generalized more lemmas
haftmann
parents: 67087
diff changeset
   238
    using that by (simp add: cong_def)
2977773a481c generalized more lemmas
haftmann
parents: 67087
diff changeset
   239
  ultimately show ?thesis
2977773a481c generalized more lemmas
haftmann
parents: 67087
diff changeset
   240
    by simp
2977773a481c generalized more lemmas
haftmann
parents: 67087
diff changeset
   241
qed 
2977773a481c generalized more lemmas
haftmann
parents: 67087
diff changeset
   242
2977773a481c generalized more lemmas
haftmann
parents: 67087
diff changeset
   243
lemma cong_imp_coprime:
2977773a481c generalized more lemmas
haftmann
parents: 67087
diff changeset
   244
  "[a = b] (mod m) \<Longrightarrow> coprime a m \<Longrightarrow> coprime b m"
2977773a481c generalized more lemmas
haftmann
parents: 67087
diff changeset
   245
  for a b :: "'a :: {unique_euclidean_semiring, euclidean_semiring_gcd}"
2977773a481c generalized more lemmas
haftmann
parents: 67087
diff changeset
   246
  by (auto simp add: coprime_iff_gcd_eq_1 dest: cong_gcd_eq)
2977773a481c generalized more lemmas
haftmann
parents: 67087
diff changeset
   247
2977773a481c generalized more lemmas
haftmann
parents: 67087
diff changeset
   248
lemma cong_cong_prod_coprime:
2977773a481c generalized more lemmas
haftmann
parents: 67087
diff changeset
   249
  "[x = y] (mod (\<Prod>i\<in>A. m i))" if
2977773a481c generalized more lemmas
haftmann
parents: 67087
diff changeset
   250
    "(\<forall>i\<in>A. [x = y] (mod m i))"
2977773a481c generalized more lemmas
haftmann
parents: 67087
diff changeset
   251
    "(\<forall>i\<in>A. (\<forall>j\<in>A. i \<noteq> j \<longrightarrow> coprime (m i) (m j)))"
2977773a481c generalized more lemmas
haftmann
parents: 67087
diff changeset
   252
  for x y :: "'a :: {unique_euclidean_ring, semiring_gcd}"
2977773a481c generalized more lemmas
haftmann
parents: 67087
diff changeset
   253
  using that by (induct A rule: infinite_finite_induct)
2977773a481c generalized more lemmas
haftmann
parents: 67087
diff changeset
   254
    (auto intro!: coprime_cong_mult prod_coprime_right)
2977773a481c generalized more lemmas
haftmann
parents: 67087
diff changeset
   255
66380
96ff0eb8294a misc tuning and modernization;
wenzelm
parents: 65417
diff changeset
   256
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 68707
diff changeset
   257
subsection \<open>Congruences on \<^typ>\<open>nat\<close> and \<^typ>\<open>int\<close>\<close>
31719
29f5b20e8ee8 Added NewNumberTheory by Jeremy Avigad
nipkow
parents:
diff changeset
   258
66888
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66837
diff changeset
   259
lemma cong_int_iff:
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66837
diff changeset
   260
  "[int m = int q] (mod int n) \<longleftrightarrow> [m = q] (mod n)"
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66837
diff changeset
   261
  by (simp add: cong_def of_nat_mod [symmetric])
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66837
diff changeset
   262
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66837
diff changeset
   263
lemma cong_Suc_0 [simp, presburger]:
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66837
diff changeset
   264
  "[m = n] (mod Suc 0)"
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66837
diff changeset
   265
  using cong_1 [of m n] by simp
31719
29f5b20e8ee8 Added NewNumberTheory by Jeremy Avigad
nipkow
parents:
diff changeset
   266
31952
40501bb2d57c renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents: 31792
diff changeset
   267
lemma cong_diff_nat:
66888
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66837
diff changeset
   268
  "[a - c = b - d] (mod m)" if "[a = b] (mod m)" "[c = d] (mod m)"
66954
0230af0f3c59 removed ancient nat-int transfer
haftmann
parents: 66888
diff changeset
   269
    and "a \<ge> c" "b \<ge> d" for a b c d m :: nat
0230af0f3c59 removed ancient nat-int transfer
haftmann
parents: 66888
diff changeset
   270
proof -
0230af0f3c59 removed ancient nat-int transfer
haftmann
parents: 66888
diff changeset
   271
  have "[c + (a - c) = d + (b - d)] (mod m)"
0230af0f3c59 removed ancient nat-int transfer
haftmann
parents: 66888
diff changeset
   272
    using that by simp
0230af0f3c59 removed ancient nat-int transfer
haftmann
parents: 66888
diff changeset
   273
  with \<open>[c = d] (mod m)\<close> have "[c + (a - c) = c + (b - d)] (mod m)"
0230af0f3c59 removed ancient nat-int transfer
haftmann
parents: 66888
diff changeset
   274
    using mod_add_cong by (auto simp add: cong_def) fastforce
0230af0f3c59 removed ancient nat-int transfer
haftmann
parents: 66888
diff changeset
   275
  then show ?thesis
0230af0f3c59 removed ancient nat-int transfer
haftmann
parents: 66888
diff changeset
   276
    by (simp add: cong_def nat_mod_eq_iff)
0230af0f3c59 removed ancient nat-int transfer
haftmann
parents: 66888
diff changeset
   277
qed
31719
29f5b20e8ee8 Added NewNumberTheory by Jeremy Avigad
nipkow
parents:
diff changeset
   278
66888
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66837
diff changeset
   279
lemma cong_diff_iff_cong_0_nat:
66954
0230af0f3c59 removed ancient nat-int transfer
haftmann
parents: 66888
diff changeset
   280
  "[a - b = 0] (mod m) \<longleftrightarrow> [a = b] (mod m)" if "a \<ge> b" for a b :: nat
76224
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 71546
diff changeset
   281
  using that by (simp add: cong_0_iff) (simp add: cong_def mod_eq_dvd_iff_nat)
31719
29f5b20e8ee8 Added NewNumberTheory by Jeremy Avigad
nipkow
parents:
diff changeset
   282
66954
0230af0f3c59 removed ancient nat-int transfer
haftmann
parents: 66888
diff changeset
   283
lemma cong_diff_iff_cong_0_nat':
0230af0f3c59 removed ancient nat-int transfer
haftmann
parents: 66888
diff changeset
   284
  "[nat \<bar>int a - int b\<bar> = 0] (mod m) \<longleftrightarrow> [a = b] (mod m)"
0230af0f3c59 removed ancient nat-int transfer
haftmann
parents: 66888
diff changeset
   285
proof (cases "b \<le> a")
0230af0f3c59 removed ancient nat-int transfer
haftmann
parents: 66888
diff changeset
   286
  case True
0230af0f3c59 removed ancient nat-int transfer
haftmann
parents: 66888
diff changeset
   287
  then show ?thesis
0230af0f3c59 removed ancient nat-int transfer
haftmann
parents: 66888
diff changeset
   288
    by (simp add: nat_diff_distrib' cong_diff_iff_cong_0_nat [of b a m])
0230af0f3c59 removed ancient nat-int transfer
haftmann
parents: 66888
diff changeset
   289
next
0230af0f3c59 removed ancient nat-int transfer
haftmann
parents: 66888
diff changeset
   290
  case False
0230af0f3c59 removed ancient nat-int transfer
haftmann
parents: 66888
diff changeset
   291
  then have "a \<le> b"
0230af0f3c59 removed ancient nat-int transfer
haftmann
parents: 66888
diff changeset
   292
    by simp
0230af0f3c59 removed ancient nat-int transfer
haftmann
parents: 66888
diff changeset
   293
  then show ?thesis
0230af0f3c59 removed ancient nat-int transfer
haftmann
parents: 66888
diff changeset
   294
    by (simp add: nat_diff_distrib' cong_diff_iff_cong_0_nat [of a b m])
0230af0f3c59 removed ancient nat-int transfer
haftmann
parents: 66888
diff changeset
   295
      (auto simp add: cong_def)
0230af0f3c59 removed ancient nat-int transfer
haftmann
parents: 66888
diff changeset
   296
qed
0230af0f3c59 removed ancient nat-int transfer
haftmann
parents: 66888
diff changeset
   297
0230af0f3c59 removed ancient nat-int transfer
haftmann
parents: 66888
diff changeset
   298
lemma cong_altdef_nat:
0230af0f3c59 removed ancient nat-int transfer
haftmann
parents: 66888
diff changeset
   299
  "a \<ge> b \<Longrightarrow> [a = b] (mod m) \<longleftrightarrow> m dvd (a - b)"
66380
96ff0eb8294a misc tuning and modernization;
wenzelm
parents: 65417
diff changeset
   300
  for a b :: nat
66888
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66837
diff changeset
   301
  by (simp add: cong_0_iff [symmetric] cong_diff_iff_cong_0_nat)
31719
29f5b20e8ee8 Added NewNumberTheory by Jeremy Avigad
nipkow
parents:
diff changeset
   302
66954
0230af0f3c59 removed ancient nat-int transfer
haftmann
parents: 66888
diff changeset
   303
lemma cong_altdef_nat':
0230af0f3c59 removed ancient nat-int transfer
haftmann
parents: 66888
diff changeset
   304
  "[a = b] (mod m) \<longleftrightarrow> m dvd nat \<bar>int a - int b\<bar>"
67085
f5d7f37b4143 tuned and generalized
haftmann
parents: 67051
diff changeset
   305
  using cong_diff_iff_cong_0_nat' [of a b m]
f5d7f37b4143 tuned and generalized
haftmann
parents: 67051
diff changeset
   306
  by (simp only: cong_0_iff [symmetric])
66954
0230af0f3c59 removed ancient nat-int transfer
haftmann
parents: 66888
diff changeset
   307
0230af0f3c59 removed ancient nat-int transfer
haftmann
parents: 66888
diff changeset
   308
lemma cong_mult_rcancel_nat:
0230af0f3c59 removed ancient nat-int transfer
haftmann
parents: 66888
diff changeset
   309
  "[a * k = b * k] (mod m) \<longleftrightarrow> [a = b] (mod m)"
0230af0f3c59 removed ancient nat-int transfer
haftmann
parents: 66888
diff changeset
   310
  if "coprime k m" for a k m :: nat
0230af0f3c59 removed ancient nat-int transfer
haftmann
parents: 66888
diff changeset
   311
proof -
0230af0f3c59 removed ancient nat-int transfer
haftmann
parents: 66888
diff changeset
   312
  have "[a * k = b * k] (mod m) \<longleftrightarrow> m dvd nat \<bar>int (a * k) - int (b * k)\<bar>"
0230af0f3c59 removed ancient nat-int transfer
haftmann
parents: 66888
diff changeset
   313
    by (simp add: cong_altdef_nat')
0230af0f3c59 removed ancient nat-int transfer
haftmann
parents: 66888
diff changeset
   314
  also have "\<dots> \<longleftrightarrow> m dvd nat \<bar>(int a - int b) * int k\<bar>"
0230af0f3c59 removed ancient nat-int transfer
haftmann
parents: 66888
diff changeset
   315
    by (simp add: algebra_simps)
0230af0f3c59 removed ancient nat-int transfer
haftmann
parents: 66888
diff changeset
   316
  also have "\<dots> \<longleftrightarrow> m dvd nat \<bar>int a - int b\<bar> * k"
0230af0f3c59 removed ancient nat-int transfer
haftmann
parents: 66888
diff changeset
   317
    by (simp add: abs_mult nat_times_as_int)
0230af0f3c59 removed ancient nat-int transfer
haftmann
parents: 66888
diff changeset
   318
  also have "\<dots> \<longleftrightarrow> m dvd nat \<bar>int a - int b\<bar>"
67051
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66954
diff changeset
   319
    by (rule coprime_dvd_mult_left_iff) (use \<open>coprime k m\<close> in \<open>simp add: ac_simps\<close>)
66954
0230af0f3c59 removed ancient nat-int transfer
haftmann
parents: 66888
diff changeset
   320
  also have "\<dots> \<longleftrightarrow> [a = b] (mod m)"
0230af0f3c59 removed ancient nat-int transfer
haftmann
parents: 66888
diff changeset
   321
    by (simp add: cong_altdef_nat')
0230af0f3c59 removed ancient nat-int transfer
haftmann
parents: 66888
diff changeset
   322
  finally show ?thesis .
0230af0f3c59 removed ancient nat-int transfer
haftmann
parents: 66888
diff changeset
   323
qed
31719
29f5b20e8ee8 Added NewNumberTheory by Jeremy Avigad
nipkow
parents:
diff changeset
   324
66954
0230af0f3c59 removed ancient nat-int transfer
haftmann
parents: 66888
diff changeset
   325
lemma cong_mult_lcancel_nat:
0230af0f3c59 removed ancient nat-int transfer
haftmann
parents: 66888
diff changeset
   326
  "[k * a = k * b] (mod m) = [a = b] (mod m)"
0230af0f3c59 removed ancient nat-int transfer
haftmann
parents: 66888
diff changeset
   327
  if "coprime k m" for a k m :: nat
0230af0f3c59 removed ancient nat-int transfer
haftmann
parents: 66888
diff changeset
   328
  using that by (simp add: cong_mult_rcancel_nat ac_simps)
31719
29f5b20e8ee8 Added NewNumberTheory by Jeremy Avigad
nipkow
parents:
diff changeset
   329
31952
40501bb2d57c renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents: 31792
diff changeset
   330
lemma coprime_cong_mult_nat:
66380
96ff0eb8294a misc tuning and modernization;
wenzelm
parents: 65417
diff changeset
   331
  "[a = b] (mod m) \<Longrightarrow> [a = b] (mod n) \<Longrightarrow> coprime m n \<Longrightarrow> [a = b] (mod m * n)"
96ff0eb8294a misc tuning and modernization;
wenzelm
parents: 65417
diff changeset
   332
  for a b :: nat
66954
0230af0f3c59 removed ancient nat-int transfer
haftmann
parents: 66888
diff changeset
   333
  by (simp add: cong_altdef_nat' divides_mult)
31719
29f5b20e8ee8 Added NewNumberTheory by Jeremy Avigad
nipkow
parents:
diff changeset
   334
66380
96ff0eb8294a misc tuning and modernization;
wenzelm
parents: 65417
diff changeset
   335
lemma cong_less_imp_eq_nat: "0 \<le> a \<Longrightarrow> a < m \<Longrightarrow> 0 \<le> b \<Longrightarrow> b < m \<Longrightarrow> [a = b] (mod m) \<Longrightarrow> a = b"
96ff0eb8294a misc tuning and modernization;
wenzelm
parents: 65417
diff changeset
   336
  for a b :: nat
66888
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66837
diff changeset
   337
  by (auto simp add: cong_def)
31719
29f5b20e8ee8 Added NewNumberTheory by Jeremy Avigad
nipkow
parents:
diff changeset
   338
66380
96ff0eb8294a misc tuning and modernization;
wenzelm
parents: 65417
diff changeset
   339
lemma cong_less_imp_eq_int: "0 \<le> a \<Longrightarrow> a < m \<Longrightarrow> 0 \<le> b \<Longrightarrow> b < m \<Longrightarrow> [a = b] (mod m) \<Longrightarrow> a = b"
96ff0eb8294a misc tuning and modernization;
wenzelm
parents: 65417
diff changeset
   340
  for a b :: int
66888
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66837
diff changeset
   341
  by (auto simp add: cong_def)
31719
29f5b20e8ee8 Added NewNumberTheory by Jeremy Avigad
nipkow
parents:
diff changeset
   342
66380
96ff0eb8294a misc tuning and modernization;
wenzelm
parents: 65417
diff changeset
   343
lemma cong_less_unique_nat: "0 < m \<Longrightarrow> (\<exists>!b. 0 \<le> b \<and> b < m \<and> [a = b] (mod m))"
96ff0eb8294a misc tuning and modernization;
wenzelm
parents: 65417
diff changeset
   344
  for a m :: nat
76231
8a48e18f081e reduce prominence of facts
haftmann
parents: 76224
diff changeset
   345
  by (auto simp: cong_def) (metis mod_mod_trivial mod_less_divisor)
31719
29f5b20e8ee8 Added NewNumberTheory by Jeremy Avigad
nipkow
parents:
diff changeset
   346
66380
96ff0eb8294a misc tuning and modernization;
wenzelm
parents: 65417
diff changeset
   347
lemma cong_less_unique_int: "0 < m \<Longrightarrow> (\<exists>!b. 0 \<le> b \<and> b < m \<and> [a = b] (mod m))"
96ff0eb8294a misc tuning and modernization;
wenzelm
parents: 65417
diff changeset
   348
  for a m :: int
76231
8a48e18f081e reduce prominence of facts
haftmann
parents: 76224
diff changeset
   349
  by (auto simp add: cong_def) (metis mod_mod_trivial pos_mod_bound pos_mod_sign)
31719
29f5b20e8ee8 Added NewNumberTheory by Jeremy Avigad
nipkow
parents:
diff changeset
   350
76224
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 71546
diff changeset
   351
lemma cong_iff_lin_nat: "[a = b] (mod m) \<longleftrightarrow> (\<exists>k1 k2. b + k1 * m = a + k2 * m)"
66380
96ff0eb8294a misc tuning and modernization;
wenzelm
parents: 65417
diff changeset
   352
  for a b :: nat
76224
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 71546
diff changeset
   353
  apply (auto simp add: cong_def nat_mod_eq_iff)
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 71546
diff changeset
   354
   apply (metis mult.commute)
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 71546
diff changeset
   355
  apply (metis mult.commute)
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 71546
diff changeset
   356
  done
31719
29f5b20e8ee8 Added NewNumberTheory by Jeremy Avigad
nipkow
parents:
diff changeset
   357
66380
96ff0eb8294a misc tuning and modernization;
wenzelm
parents: 65417
diff changeset
   358
lemma cong_cong_mod_nat: "[a = b] (mod m) \<longleftrightarrow> [a mod m = b mod m] (mod m)"
96ff0eb8294a misc tuning and modernization;
wenzelm
parents: 65417
diff changeset
   359
  for a b :: nat
66888
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66837
diff changeset
   360
  by simp
31719
29f5b20e8ee8 Added NewNumberTheory by Jeremy Avigad
nipkow
parents:
diff changeset
   361
66380
96ff0eb8294a misc tuning and modernization;
wenzelm
parents: 65417
diff changeset
   362
lemma cong_cong_mod_int: "[a = b] (mod m) \<longleftrightarrow> [a mod m = b mod m] (mod m)"
96ff0eb8294a misc tuning and modernization;
wenzelm
parents: 65417
diff changeset
   363
  for a b :: int
66888
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66837
diff changeset
   364
  by simp
31719
29f5b20e8ee8 Added NewNumberTheory by Jeremy Avigad
nipkow
parents:
diff changeset
   365
66380
96ff0eb8294a misc tuning and modernization;
wenzelm
parents: 65417
diff changeset
   366
lemma cong_add_lcancel_nat: "[a + x = a + y] (mod n) \<longleftrightarrow> [x = y] (mod n)"
96ff0eb8294a misc tuning and modernization;
wenzelm
parents: 65417
diff changeset
   367
  for a x y :: nat
31952
40501bb2d57c renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents: 31792
diff changeset
   368
  by (simp add: cong_iff_lin_nat)
31719
29f5b20e8ee8 Added NewNumberTheory by Jeremy Avigad
nipkow
parents:
diff changeset
   369
66380
96ff0eb8294a misc tuning and modernization;
wenzelm
parents: 65417
diff changeset
   370
lemma cong_add_rcancel_nat: "[x + a = y + a] (mod n) \<longleftrightarrow> [x = y] (mod n)"
96ff0eb8294a misc tuning and modernization;
wenzelm
parents: 65417
diff changeset
   371
  for a x y :: nat
31952
40501bb2d57c renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents: 31792
diff changeset
   372
  by (simp add: cong_iff_lin_nat)
31719
29f5b20e8ee8 Added NewNumberTheory by Jeremy Avigad
nipkow
parents:
diff changeset
   373
66380
96ff0eb8294a misc tuning and modernization;
wenzelm
parents: 65417
diff changeset
   374
lemma cong_add_lcancel_0_nat: "[a + x = a] (mod n) \<longleftrightarrow> [x = 0] (mod n)"
96ff0eb8294a misc tuning and modernization;
wenzelm
parents: 65417
diff changeset
   375
  for a x :: nat
67085
f5d7f37b4143 tuned and generalized
haftmann
parents: 67051
diff changeset
   376
  using cong_add_lcancel_nat [of a x 0 n] by simp
31719
29f5b20e8ee8 Added NewNumberTheory by Jeremy Avigad
nipkow
parents:
diff changeset
   377
66380
96ff0eb8294a misc tuning and modernization;
wenzelm
parents: 65417
diff changeset
   378
lemma cong_add_rcancel_0_nat: "[x + a = a] (mod n) \<longleftrightarrow> [x = 0] (mod n)"
96ff0eb8294a misc tuning and modernization;
wenzelm
parents: 65417
diff changeset
   379
  for a x :: nat
67085
f5d7f37b4143 tuned and generalized
haftmann
parents: 67051
diff changeset
   380
  using cong_add_rcancel_nat [of x a 0 n] by simp
66380
96ff0eb8294a misc tuning and modernization;
wenzelm
parents: 65417
diff changeset
   381
96ff0eb8294a misc tuning and modernization;
wenzelm
parents: 65417
diff changeset
   382
lemma cong_dvd_modulus_nat: "[x = y] (mod m) \<Longrightarrow> n dvd m \<Longrightarrow> [x = y] (mod n)"
96ff0eb8294a misc tuning and modernization;
wenzelm
parents: 65417
diff changeset
   383
  for x y :: nat
76224
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 71546
diff changeset
   384
  by (auto simp add: cong_altdef_nat')
31719
29f5b20e8ee8 Added NewNumberTheory by Jeremy Avigad
nipkow
parents:
diff changeset
   385
66380
96ff0eb8294a misc tuning and modernization;
wenzelm
parents: 65417
diff changeset
   386
lemma cong_to_1_nat:
96ff0eb8294a misc tuning and modernization;
wenzelm
parents: 65417
diff changeset
   387
  fixes a :: nat
96ff0eb8294a misc tuning and modernization;
wenzelm
parents: 65417
diff changeset
   388
  assumes "[a = 1] (mod n)"
96ff0eb8294a misc tuning and modernization;
wenzelm
parents: 65417
diff changeset
   389
  shows "n dvd (a - 1)"
96ff0eb8294a misc tuning and modernization;
wenzelm
parents: 65417
diff changeset
   390
proof (cases "a = 0")
96ff0eb8294a misc tuning and modernization;
wenzelm
parents: 65417
diff changeset
   391
  case True
96ff0eb8294a misc tuning and modernization;
wenzelm
parents: 65417
diff changeset
   392
  then show ?thesis by force
96ff0eb8294a misc tuning and modernization;
wenzelm
parents: 65417
diff changeset
   393
next
96ff0eb8294a misc tuning and modernization;
wenzelm
parents: 65417
diff changeset
   394
  case False
96ff0eb8294a misc tuning and modernization;
wenzelm
parents: 65417
diff changeset
   395
  with assms show ?thesis by (metis cong_altdef_nat leI less_one)
96ff0eb8294a misc tuning and modernization;
wenzelm
parents: 65417
diff changeset
   396
qed
96ff0eb8294a misc tuning and modernization;
wenzelm
parents: 65417
diff changeset
   397
96ff0eb8294a misc tuning and modernization;
wenzelm
parents: 65417
diff changeset
   398
lemma cong_0_1_nat': "[0 = Suc 0] (mod n) \<longleftrightarrow> n = Suc 0"
66888
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66837
diff changeset
   399
  by (auto simp: cong_def)
66380
96ff0eb8294a misc tuning and modernization;
wenzelm
parents: 65417
diff changeset
   400
96ff0eb8294a misc tuning and modernization;
wenzelm
parents: 65417
diff changeset
   401
lemma cong_0_1_nat: "[0 = 1] (mod n) \<longleftrightarrow> n = 1"
96ff0eb8294a misc tuning and modernization;
wenzelm
parents: 65417
diff changeset
   402
  for n :: nat
66888
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66837
diff changeset
   403
  by (auto simp: cong_def)
66380
96ff0eb8294a misc tuning and modernization;
wenzelm
parents: 65417
diff changeset
   404
96ff0eb8294a misc tuning and modernization;
wenzelm
parents: 65417
diff changeset
   405
lemma cong_0_1_int: "[0 = 1] (mod n) \<longleftrightarrow> n = 1 \<or> n = - 1"
96ff0eb8294a misc tuning and modernization;
wenzelm
parents: 65417
diff changeset
   406
  for n :: int
66888
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66837
diff changeset
   407
  by (auto simp: cong_def zmult_eq_1_iff)
66380
96ff0eb8294a misc tuning and modernization;
wenzelm
parents: 65417
diff changeset
   408
96ff0eb8294a misc tuning and modernization;
wenzelm
parents: 65417
diff changeset
   409
lemma cong_to_1'_nat: "[a = 1] (mod n) \<longleftrightarrow> a = 0 \<and> n = 1 \<or> (\<exists>m. a = 1 + m * n)"
96ff0eb8294a misc tuning and modernization;
wenzelm
parents: 65417
diff changeset
   410
  for a :: nat
96ff0eb8294a misc tuning and modernization;
wenzelm
parents: 65417
diff changeset
   411
  by (metis add.right_neutral cong_0_1_nat cong_iff_lin_nat cong_to_1_nat
96ff0eb8294a misc tuning and modernization;
wenzelm
parents: 65417
diff changeset
   412
      dvd_div_mult_self leI le_add_diff_inverse less_one mult_eq_if)
96ff0eb8294a misc tuning and modernization;
wenzelm
parents: 65417
diff changeset
   413
96ff0eb8294a misc tuning and modernization;
wenzelm
parents: 65417
diff changeset
   414
lemma cong_le_nat: "y \<le> x \<Longrightarrow> [x = y] (mod n) \<longleftrightarrow> (\<exists>q. x = q * n + y)"
96ff0eb8294a misc tuning and modernization;
wenzelm
parents: 65417
diff changeset
   415
  for x y :: nat
76224
64e8d4afcf10 moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents: 71546
diff changeset
   416
  by (auto simp add: cong_altdef_nat le_imp_diff_is_add)
68707
5b269897df9d de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents: 67115
diff changeset
   417
66380
96ff0eb8294a misc tuning and modernization;
wenzelm
parents: 65417
diff changeset
   418
lemma cong_solve_nat:
96ff0eb8294a misc tuning and modernization;
wenzelm
parents: 65417
diff changeset
   419
  fixes a :: nat
96ff0eb8294a misc tuning and modernization;
wenzelm
parents: 65417
diff changeset
   420
  shows "\<exists>x. [a * x = gcd a n] (mod n)"
68707
5b269897df9d de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents: 67115
diff changeset
   421
proof (cases "a = 0 \<or> n = 0")
66380
96ff0eb8294a misc tuning and modernization;
wenzelm
parents: 65417
diff changeset
   422
  case True
68707
5b269897df9d de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents: 67115
diff changeset
   423
  then show ?thesis
5b269897df9d de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents: 67115
diff changeset
   424
    by (force simp add: cong_0_iff cong_sym)
66380
96ff0eb8294a misc tuning and modernization;
wenzelm
parents: 65417
diff changeset
   425
next
96ff0eb8294a misc tuning and modernization;
wenzelm
parents: 65417
diff changeset
   426
  case False
96ff0eb8294a misc tuning and modernization;
wenzelm
parents: 65417
diff changeset
   427
  then show ?thesis
68707
5b269897df9d de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents: 67115
diff changeset
   428
    using bezout_nat [of a n]
66888
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66837
diff changeset
   429
    by auto (metis cong_add_rcancel_0_nat cong_mult_self_left)
66380
96ff0eb8294a misc tuning and modernization;
wenzelm
parents: 65417
diff changeset
   430
qed
96ff0eb8294a misc tuning and modernization;
wenzelm
parents: 65417
diff changeset
   431
68707
5b269897df9d de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents: 67115
diff changeset
   432
lemma cong_solve_int:
5b269897df9d de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents: 67115
diff changeset
   433
  fixes a :: int
5b269897df9d de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents: 67115
diff changeset
   434
  shows "\<exists>x. [a * x = gcd a n] (mod n)"
5b269897df9d de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents: 67115
diff changeset
   435
    by (metis bezout_int cong_iff_lin mult.commute)
31719
29f5b20e8ee8 Added NewNumberTheory by Jeremy Avigad
nipkow
parents:
diff changeset
   436
44872
a98ef45122f3 misc tuning;
wenzelm
parents: 41959
diff changeset
   437
lemma cong_solve_dvd_nat:
66380
96ff0eb8294a misc tuning and modernization;
wenzelm
parents: 65417
diff changeset
   438
  fixes a :: nat
68707
5b269897df9d de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents: 67115
diff changeset
   439
  assumes "gcd a n dvd d"
66380
96ff0eb8294a misc tuning and modernization;
wenzelm
parents: 65417
diff changeset
   440
  shows "\<exists>x. [a * x = d] (mod n)"
31719
29f5b20e8ee8 Added NewNumberTheory by Jeremy Avigad
nipkow
parents:
diff changeset
   441
proof -
68707
5b269897df9d de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents: 67115
diff changeset
   442
  from cong_solve_nat [of a] obtain x where "[a * x = gcd a n](mod n)"
31719
29f5b20e8ee8 Added NewNumberTheory by Jeremy Avigad
nipkow
parents:
diff changeset
   443
    by auto
44872
a98ef45122f3 misc tuning;
wenzelm
parents: 41959
diff changeset
   444
  then have "[(d div gcd a n) * (a * x) = (d div gcd a n) * gcd a n] (mod n)"
66888
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66837
diff changeset
   445
    using cong_scalar_left by blast
68707
5b269897df9d de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents: 67115
diff changeset
   446
  also from assms have "(d div gcd a n) * gcd a n = d"
31719
29f5b20e8ee8 Added NewNumberTheory by Jeremy Avigad
nipkow
parents:
diff changeset
   447
    by (rule dvd_div_mult_self)
29f5b20e8ee8 Added NewNumberTheory by Jeremy Avigad
nipkow
parents:
diff changeset
   448
  also have "(d div gcd a n) * (a * x) = a * (d div gcd a n * x)"
29f5b20e8ee8 Added NewNumberTheory by Jeremy Avigad
nipkow
parents:
diff changeset
   449
    by auto
29f5b20e8ee8 Added NewNumberTheory by Jeremy Avigad
nipkow
parents:
diff changeset
   450
  finally show ?thesis
29f5b20e8ee8 Added NewNumberTheory by Jeremy Avigad
nipkow
parents:
diff changeset
   451
    by auto
29f5b20e8ee8 Added NewNumberTheory by Jeremy Avigad
nipkow
parents:
diff changeset
   452
qed
29f5b20e8ee8 Added NewNumberTheory by Jeremy Avigad
nipkow
parents:
diff changeset
   453
44872
a98ef45122f3 misc tuning;
wenzelm
parents: 41959
diff changeset
   454
lemma cong_solve_dvd_int:
68707
5b269897df9d de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents: 67115
diff changeset
   455
  fixes a::int
5b269897df9d de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents: 67115
diff changeset
   456
  assumes b: "gcd a n dvd d"
66380
96ff0eb8294a misc tuning and modernization;
wenzelm
parents: 65417
diff changeset
   457
  shows "\<exists>x. [a * x = d] (mod n)"
31719
29f5b20e8ee8 Added NewNumberTheory by Jeremy Avigad
nipkow
parents:
diff changeset
   458
proof -
68707
5b269897df9d de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents: 67115
diff changeset
   459
  from cong_solve_int [of a] obtain x where "[a * x = gcd a n](mod n)"
31719
29f5b20e8ee8 Added NewNumberTheory by Jeremy Avigad
nipkow
parents:
diff changeset
   460
    by auto
44872
a98ef45122f3 misc tuning;
wenzelm
parents: 41959
diff changeset
   461
  then have "[(d div gcd a n) * (a * x) = (d div gcd a n) * gcd a n] (mod n)"
66888
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66837
diff changeset
   462
    using cong_scalar_left by blast
31719
29f5b20e8ee8 Added NewNumberTheory by Jeremy Avigad
nipkow
parents:
diff changeset
   463
  also from b have "(d div gcd a n) * gcd a n = d"
29f5b20e8ee8 Added NewNumberTheory by Jeremy Avigad
nipkow
parents:
diff changeset
   464
    by (rule dvd_div_mult_self)
29f5b20e8ee8 Added NewNumberTheory by Jeremy Avigad
nipkow
parents:
diff changeset
   465
  also have "(d div gcd a n) * (a * x) = a * (d div gcd a n * x)"
29f5b20e8ee8 Added NewNumberTheory by Jeremy Avigad
nipkow
parents:
diff changeset
   466
    by auto
29f5b20e8ee8 Added NewNumberTheory by Jeremy Avigad
nipkow
parents:
diff changeset
   467
  finally show ?thesis
29f5b20e8ee8 Added NewNumberTheory by Jeremy Avigad
nipkow
parents:
diff changeset
   468
    by auto
29f5b20e8ee8 Added NewNumberTheory by Jeremy Avigad
nipkow
parents:
diff changeset
   469
qed
29f5b20e8ee8 Added NewNumberTheory by Jeremy Avigad
nipkow
parents:
diff changeset
   470
66380
96ff0eb8294a misc tuning and modernization;
wenzelm
parents: 65417
diff changeset
   471
lemma cong_solve_coprime_nat:
67051
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66954
diff changeset
   472
  "\<exists>x. [a * x = Suc 0] (mod n)" if "coprime a n"
68707
5b269897df9d de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents: 67115
diff changeset
   473
  using that cong_solve_nat [of a n] by auto
31719
29f5b20e8ee8 Added NewNumberTheory by Jeremy Avigad
nipkow
parents:
diff changeset
   474
67051
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66954
diff changeset
   475
lemma cong_solve_coprime_int:
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66954
diff changeset
   476
  "\<exists>x. [a * x = 1] (mod n)" if "coprime a n" for a n x :: int
68707
5b269897df9d de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents: 67115
diff changeset
   477
  using that cong_solve_int [of a n] by (auto simp add: zabs_def split: if_splits)
55161
8eb891539804 minor adjustments
paulson <lp15@cam.ac.uk>
parents: 55130
diff changeset
   478
62349
7c23469b5118 cleansed junk-producing interpretations for gcd/lcm on nat altogether
haftmann
parents: 62348
diff changeset
   479
lemma coprime_iff_invertible_nat:
67085
f5d7f37b4143 tuned and generalized
haftmann
parents: 67051
diff changeset
   480
  "coprime a m \<longleftrightarrow> (\<exists>x. [a * x = Suc 0] (mod m))" (is "?P \<longleftrightarrow> ?Q")
f5d7f37b4143 tuned and generalized
haftmann
parents: 67051
diff changeset
   481
proof
f5d7f37b4143 tuned and generalized
haftmann
parents: 67051
diff changeset
   482
  assume ?P then show ?Q
f5d7f37b4143 tuned and generalized
haftmann
parents: 67051
diff changeset
   483
    by (auto dest!: cong_solve_coprime_nat)
f5d7f37b4143 tuned and generalized
haftmann
parents: 67051
diff changeset
   484
next
f5d7f37b4143 tuned and generalized
haftmann
parents: 67051
diff changeset
   485
  assume ?Q
f5d7f37b4143 tuned and generalized
haftmann
parents: 67051
diff changeset
   486
  then obtain b where "[a * b = Suc 0] (mod m)"
f5d7f37b4143 tuned and generalized
haftmann
parents: 67051
diff changeset
   487
    by blast
f5d7f37b4143 tuned and generalized
haftmann
parents: 67051
diff changeset
   488
  with coprime_mod_left_iff [of m "a * b"] show ?P
f5d7f37b4143 tuned and generalized
haftmann
parents: 67051
diff changeset
   489
    by (cases "m = 0 \<or> m = 1")
f5d7f37b4143 tuned and generalized
haftmann
parents: 67051
diff changeset
   490
      (unfold cong_def, auto simp add: cong_def)
f5d7f37b4143 tuned and generalized
haftmann
parents: 67051
diff changeset
   491
qed
66380
96ff0eb8294a misc tuning and modernization;
wenzelm
parents: 65417
diff changeset
   492
67051
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66954
diff changeset
   493
lemma coprime_iff_invertible_int:
67085
f5d7f37b4143 tuned and generalized
haftmann
parents: 67051
diff changeset
   494
  "coprime a m \<longleftrightarrow> (\<exists>x. [a * x = 1] (mod m))" (is "?P \<longleftrightarrow> ?Q") for m :: int
f5d7f37b4143 tuned and generalized
haftmann
parents: 67051
diff changeset
   495
proof
f5d7f37b4143 tuned and generalized
haftmann
parents: 67051
diff changeset
   496
  assume ?P then show ?Q
f5d7f37b4143 tuned and generalized
haftmann
parents: 67051
diff changeset
   497
    by (auto dest: cong_solve_coprime_int)
f5d7f37b4143 tuned and generalized
haftmann
parents: 67051
diff changeset
   498
next
f5d7f37b4143 tuned and generalized
haftmann
parents: 67051
diff changeset
   499
  assume ?Q
f5d7f37b4143 tuned and generalized
haftmann
parents: 67051
diff changeset
   500
  then obtain b where "[a * b = 1] (mod m)"
f5d7f37b4143 tuned and generalized
haftmann
parents: 67051
diff changeset
   501
    by blast
f5d7f37b4143 tuned and generalized
haftmann
parents: 67051
diff changeset
   502
  with coprime_mod_left_iff [of m "a * b"] show ?P
f5d7f37b4143 tuned and generalized
haftmann
parents: 67051
diff changeset
   503
    by (cases "m = 0 \<or> m = 1")
f5d7f37b4143 tuned and generalized
haftmann
parents: 67051
diff changeset
   504
      (unfold cong_def, auto simp add: zmult_eq_1_iff)
f5d7f37b4143 tuned and generalized
haftmann
parents: 67051
diff changeset
   505
qed
31719
29f5b20e8ee8 Added NewNumberTheory by Jeremy Avigad
nipkow
parents:
diff changeset
   506
66380
96ff0eb8294a misc tuning and modernization;
wenzelm
parents: 65417
diff changeset
   507
lemma coprime_iff_invertible'_nat:
68707
5b269897df9d de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents: 67115
diff changeset
   508
  assumes "m > 0"
5b269897df9d de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents: 67115
diff changeset
   509
  shows "coprime a m \<longleftrightarrow> (\<exists>x. 0 \<le> x \<and> x < m \<and> [a * x = Suc 0] (mod m))"
5b269897df9d de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents: 67115
diff changeset
   510
proof -
5b269897df9d de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents: 67115
diff changeset
   511
  have "\<And>b. \<lbrakk>0 < m; [a * b = Suc 0] (mod m)\<rbrakk> \<Longrightarrow> \<exists>b'<m. [a * b' = Suc 0] (mod m)"
5b269897df9d de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents: 67115
diff changeset
   512
    by (metis cong_def mod_less_divisor [OF assms] mod_mult_right_eq)
5b269897df9d de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents: 67115
diff changeset
   513
  then show ?thesis
5b269897df9d de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents: 67115
diff changeset
   514
    using assms coprime_iff_invertible_nat by auto
5b269897df9d de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents: 67115
diff changeset
   515
qed
31719
29f5b20e8ee8 Added NewNumberTheory by Jeremy Avigad
nipkow
parents:
diff changeset
   516
66380
96ff0eb8294a misc tuning and modernization;
wenzelm
parents: 65417
diff changeset
   517
lemma coprime_iff_invertible'_int:
68707
5b269897df9d de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents: 67115
diff changeset
   518
  fixes m :: int
5b269897df9d de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents: 67115
diff changeset
   519
  assumes "m > 0"
5b269897df9d de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents: 67115
diff changeset
   520
  shows "coprime a m \<longleftrightarrow> (\<exists>x. 0 \<le> x \<and> x < m \<and> [a * x = 1] (mod m))"
76231
8a48e18f081e reduce prominence of facts
haftmann
parents: 76224
diff changeset
   521
  using assms by (simp add: coprime_iff_invertible_int)
8a48e18f081e reduce prominence of facts
haftmann
parents: 76224
diff changeset
   522
    (metis assms cong_mod_left mod_mult_right_eq pos_mod_bound pos_mod_sign)
31719
29f5b20e8ee8 Added NewNumberTheory by Jeremy Avigad
nipkow
parents:
diff changeset
   523
66380
96ff0eb8294a misc tuning and modernization;
wenzelm
parents: 65417
diff changeset
   524
lemma cong_cong_lcm_nat: "[x = y] (mod a) \<Longrightarrow> [x = y] (mod b) \<Longrightarrow> [x = y] (mod lcm a b)"
96ff0eb8294a misc tuning and modernization;
wenzelm
parents: 65417
diff changeset
   525
  for x y :: nat
68707
5b269897df9d de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents: 67115
diff changeset
   526
  by (meson cong_altdef_nat' lcm_least)
31719
29f5b20e8ee8 Added NewNumberTheory by Jeremy Avigad
nipkow
parents:
diff changeset
   527
66380
96ff0eb8294a misc tuning and modernization;
wenzelm
parents: 65417
diff changeset
   528
lemma cong_cong_lcm_int: "[x = y] (mod a) \<Longrightarrow> [x = y] (mod b) \<Longrightarrow> [x = y] (mod lcm a b)"
96ff0eb8294a misc tuning and modernization;
wenzelm
parents: 65417
diff changeset
   529
  for x y :: int
67115
2977773a481c generalized more lemmas
haftmann
parents: 67087
diff changeset
   530
  by (auto simp add: cong_iff_dvd_diff lcm_least)
31719
29f5b20e8ee8 Added NewNumberTheory by Jeremy Avigad
nipkow
parents:
diff changeset
   531
66888
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66837
diff changeset
   532
lemma cong_cong_prod_coprime_nat:
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66837
diff changeset
   533
  "[x = y] (mod (\<Prod>i\<in>A. m i))" if
67115
2977773a481c generalized more lemmas
haftmann
parents: 67087
diff changeset
   534
    "(\<forall>i\<in>A. [x = y] (mod m i))"
66888
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66837
diff changeset
   535
    "(\<forall>i\<in>A. (\<forall>j\<in>A. i \<noteq> j \<longrightarrow> coprime (m i) (m j)))"
67115
2977773a481c generalized more lemmas
haftmann
parents: 67087
diff changeset
   536
  for x y :: nat
2977773a481c generalized more lemmas
haftmann
parents: 67087
diff changeset
   537
  using that by (induct A rule: infinite_finite_induct)
2977773a481c generalized more lemmas
haftmann
parents: 67087
diff changeset
   538
    (auto intro!: coprime_cong_mult_nat prod_coprime_right)
31719
29f5b20e8ee8 Added NewNumberTheory by Jeremy Avigad
nipkow
parents:
diff changeset
   539
31952
40501bb2d57c renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents: 31792
diff changeset
   540
lemma binary_chinese_remainder_nat:
66380
96ff0eb8294a misc tuning and modernization;
wenzelm
parents: 65417
diff changeset
   541
  fixes m1 m2 :: nat
96ff0eb8294a misc tuning and modernization;
wenzelm
parents: 65417
diff changeset
   542
  assumes a: "coprime m1 m2"
96ff0eb8294a misc tuning and modernization;
wenzelm
parents: 65417
diff changeset
   543
  shows "\<exists>x. [x = u1] (mod m1) \<and> [x = u2] (mod m2)"
31719
29f5b20e8ee8 Added NewNumberTheory by Jeremy Avigad
nipkow
parents:
diff changeset
   544
proof -
67086
haftmann
parents: 67085
diff changeset
   545
  have "\<exists>b1 b2. [b1 = 1] (mod m1) \<and> [b1 = 0] (mod m2) \<and> [b2 = 0] (mod m1) \<and> [b2 = 1] (mod m2)"
haftmann
parents: 67085
diff changeset
   546
  proof -
haftmann
parents: 67085
diff changeset
   547
    from cong_solve_coprime_nat [OF a] obtain x1 where 1: "[m1 * x1 = 1] (mod m2)"
haftmann
parents: 67085
diff changeset
   548
      by auto
haftmann
parents: 67085
diff changeset
   549
    from a have b: "coprime m2 m1"
haftmann
parents: 67085
diff changeset
   550
      by (simp add: ac_simps)
haftmann
parents: 67085
diff changeset
   551
    from cong_solve_coprime_nat [OF b] obtain x2 where 2: "[m2 * x2 = 1] (mod m1)"
haftmann
parents: 67085
diff changeset
   552
      by auto
haftmann
parents: 67085
diff changeset
   553
    have "[m1 * x1 = 0] (mod m1)"
haftmann
parents: 67085
diff changeset
   554
      by (simp add: cong_mult_self_left)
haftmann
parents: 67085
diff changeset
   555
    moreover have "[m2 * x2 = 0] (mod m2)"
haftmann
parents: 67085
diff changeset
   556
      by (simp add: cong_mult_self_left)
haftmann
parents: 67085
diff changeset
   557
    ultimately show ?thesis
haftmann
parents: 67085
diff changeset
   558
      using 1 2 by blast
haftmann
parents: 67085
diff changeset
   559
  qed
haftmann
parents: 67085
diff changeset
   560
  then obtain b1 b2
66380
96ff0eb8294a misc tuning and modernization;
wenzelm
parents: 65417
diff changeset
   561
    where "[b1 = 1] (mod m1)" and "[b1 = 0] (mod m2)"
96ff0eb8294a misc tuning and modernization;
wenzelm
parents: 65417
diff changeset
   562
      and "[b2 = 0] (mod m1)" and "[b2 = 1] (mod m2)"
31719
29f5b20e8ee8 Added NewNumberTheory by Jeremy Avigad
nipkow
parents:
diff changeset
   563
    by blast
29f5b20e8ee8 Added NewNumberTheory by Jeremy Avigad
nipkow
parents:
diff changeset
   564
  let ?x = "u1 * b1 + u2 * b2"
29f5b20e8ee8 Added NewNumberTheory by Jeremy Avigad
nipkow
parents:
diff changeset
   565
  have "[?x = u1 * 1 + u2 * 0] (mod m1)"
66888
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66837
diff changeset
   566
    using \<open>[b1 = 1] (mod m1)\<close> \<open>[b2 = 0] (mod m1)\<close> cong_add cong_scalar_left by blast
44872
a98ef45122f3 misc tuning;
wenzelm
parents: 41959
diff changeset
   567
  then have "[?x = u1] (mod m1)" by simp
31719
29f5b20e8ee8 Added NewNumberTheory by Jeremy Avigad
nipkow
parents:
diff changeset
   568
  have "[?x = u1 * 0 + u2 * 1] (mod m2)"
66888
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66837
diff changeset
   569
    using \<open>[b1 = 0] (mod m2)\<close> \<open>[b2 = 1] (mod m2)\<close> cong_add cong_scalar_left by blast
66380
96ff0eb8294a misc tuning and modernization;
wenzelm
parents: 65417
diff changeset
   570
  then have "[?x = u2] (mod m2)"
96ff0eb8294a misc tuning and modernization;
wenzelm
parents: 65417
diff changeset
   571
    by simp
96ff0eb8294a misc tuning and modernization;
wenzelm
parents: 65417
diff changeset
   572
  with \<open>[?x = u1] (mod m1)\<close> show ?thesis
96ff0eb8294a misc tuning and modernization;
wenzelm
parents: 65417
diff changeset
   573
    by blast
31719
29f5b20e8ee8 Added NewNumberTheory by Jeremy Avigad
nipkow
parents:
diff changeset
   574
qed
29f5b20e8ee8 Added NewNumberTheory by Jeremy Avigad
nipkow
parents:
diff changeset
   575
31952
40501bb2d57c renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents: 31792
diff changeset
   576
lemma binary_chinese_remainder_int:
66380
96ff0eb8294a misc tuning and modernization;
wenzelm
parents: 65417
diff changeset
   577
  fixes m1 m2 :: int
96ff0eb8294a misc tuning and modernization;
wenzelm
parents: 65417
diff changeset
   578
  assumes a: "coprime m1 m2"
96ff0eb8294a misc tuning and modernization;
wenzelm
parents: 65417
diff changeset
   579
  shows "\<exists>x. [x = u1] (mod m1) \<and> [x = u2] (mod m2)"
31719
29f5b20e8ee8 Added NewNumberTheory by Jeremy Avigad
nipkow
parents:
diff changeset
   580
proof -
67086
haftmann
parents: 67085
diff changeset
   581
  have "\<exists>b1 b2. [b1 = 1] (mod m1) \<and> [b1 = 0] (mod m2) \<and> [b2 = 0] (mod m1) \<and> [b2 = 1] (mod m2)"
haftmann
parents: 67085
diff changeset
   582
  proof -
haftmann
parents: 67085
diff changeset
   583
    from cong_solve_coprime_int [OF a] obtain x1 where 1: "[m1 * x1 = 1] (mod m2)"
haftmann
parents: 67085
diff changeset
   584
      by auto
haftmann
parents: 67085
diff changeset
   585
    from a have b: "coprime m2 m1"
haftmann
parents: 67085
diff changeset
   586
      by (simp add: ac_simps)
haftmann
parents: 67085
diff changeset
   587
    from cong_solve_coprime_int [OF b] obtain x2 where 2: "[m2 * x2 = 1] (mod m1)"
haftmann
parents: 67085
diff changeset
   588
      by auto
haftmann
parents: 67085
diff changeset
   589
    have "[m1 * x1 = 0] (mod m1)"
haftmann
parents: 67085
diff changeset
   590
     by (simp add: cong_mult_self_left)
haftmann
parents: 67085
diff changeset
   591
    moreover have "[m2 * x2 = 0] (mod m2)"
haftmann
parents: 67085
diff changeset
   592
      by (simp add: cong_mult_self_left)
haftmann
parents: 67085
diff changeset
   593
    ultimately show ?thesis
haftmann
parents: 67085
diff changeset
   594
      using 1 2 by blast
haftmann
parents: 67085
diff changeset
   595
  qed
haftmann
parents: 67085
diff changeset
   596
  then obtain b1 b2
66380
96ff0eb8294a misc tuning and modernization;
wenzelm
parents: 65417
diff changeset
   597
    where "[b1 = 1] (mod m1)" and "[b1 = 0] (mod m2)"
96ff0eb8294a misc tuning and modernization;
wenzelm
parents: 65417
diff changeset
   598
      and "[b2 = 0] (mod m1)" and "[b2 = 1] (mod m2)"
31719
29f5b20e8ee8 Added NewNumberTheory by Jeremy Avigad
nipkow
parents:
diff changeset
   599
    by blast
29f5b20e8ee8 Added NewNumberTheory by Jeremy Avigad
nipkow
parents:
diff changeset
   600
  let ?x = "u1 * b1 + u2 * b2"
29f5b20e8ee8 Added NewNumberTheory by Jeremy Avigad
nipkow
parents:
diff changeset
   601
  have "[?x = u1 * 1 + u2 * 0] (mod m1)"
66888
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66837
diff changeset
   602
    using \<open>[b1 = 1] (mod m1)\<close> \<open>[b2 = 0] (mod m1)\<close> cong_add cong_scalar_left by blast
44872
a98ef45122f3 misc tuning;
wenzelm
parents: 41959
diff changeset
   603
  then have "[?x = u1] (mod m1)" by simp
31719
29f5b20e8ee8 Added NewNumberTheory by Jeremy Avigad
nipkow
parents:
diff changeset
   604
  have "[?x = u1 * 0 + u2 * 1] (mod m2)"
66888
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66837
diff changeset
   605
    using \<open>[b1 = 0] (mod m2)\<close> \<open>[b2 = 1] (mod m2)\<close> cong_add cong_scalar_left by blast
44872
a98ef45122f3 misc tuning;
wenzelm
parents: 41959
diff changeset
   606
  then have "[?x = u2] (mod m2)" by simp
66380
96ff0eb8294a misc tuning and modernization;
wenzelm
parents: 65417
diff changeset
   607
  with \<open>[?x = u1] (mod m1)\<close> show ?thesis
96ff0eb8294a misc tuning and modernization;
wenzelm
parents: 65417
diff changeset
   608
    by blast
31719
29f5b20e8ee8 Added NewNumberTheory by Jeremy Avigad
nipkow
parents:
diff changeset
   609
qed
29f5b20e8ee8 Added NewNumberTheory by Jeremy Avigad
nipkow
parents:
diff changeset
   610
66380
96ff0eb8294a misc tuning and modernization;
wenzelm
parents: 65417
diff changeset
   611
lemma cong_modulus_mult_nat: "[x = y] (mod m * n) \<Longrightarrow> [x = y] (mod m)"
96ff0eb8294a misc tuning and modernization;
wenzelm
parents: 65417
diff changeset
   612
  for x y :: nat
68707
5b269897df9d de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents: 67115
diff changeset
   613
  by (metis cong_def mod_mult_cong_right)
31719
29f5b20e8ee8 Added NewNumberTheory by Jeremy Avigad
nipkow
parents:
diff changeset
   614
66380
96ff0eb8294a misc tuning and modernization;
wenzelm
parents: 65417
diff changeset
   615
lemma cong_less_modulus_unique_nat: "[x = y] (mod m) \<Longrightarrow> x < m \<Longrightarrow> y < m \<Longrightarrow> x = y"
96ff0eb8294a misc tuning and modernization;
wenzelm
parents: 65417
diff changeset
   616
  for x y :: nat
66888
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66837
diff changeset
   617
  by (simp add: cong_def)
31719
29f5b20e8ee8 Added NewNumberTheory by Jeremy Avigad
nipkow
parents:
diff changeset
   618
31952
40501bb2d57c renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents: 31792
diff changeset
   619
lemma binary_chinese_remainder_unique_nat:
66380
96ff0eb8294a misc tuning and modernization;
wenzelm
parents: 65417
diff changeset
   620
  fixes m1 m2 :: nat
96ff0eb8294a misc tuning and modernization;
wenzelm
parents: 65417
diff changeset
   621
  assumes a: "coprime m1 m2"
44872
a98ef45122f3 misc tuning;
wenzelm
parents: 41959
diff changeset
   622
    and nz: "m1 \<noteq> 0" "m2 \<noteq> 0"
63901
4ce989e962e0 more symbols;
wenzelm
parents: 63167
diff changeset
   623
  shows "\<exists>!x. x < m1 * m2 \<and> [x = u1] (mod m1) \<and> [x = u2] (mod m2)"
31719
29f5b20e8ee8 Added NewNumberTheory by Jeremy Avigad
nipkow
parents:
diff changeset
   624
proof -
68707
5b269897df9d de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents: 67115
diff changeset
   625
  obtain y where y1: "[y = u1] (mod m1)" and y2: "[y = u2] (mod m2)"
5b269897df9d de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents: 67115
diff changeset
   626
    using binary_chinese_remainder_nat [OF a] by blast
31719
29f5b20e8ee8 Added NewNumberTheory by Jeremy Avigad
nipkow
parents:
diff changeset
   627
  let ?x = "y mod (m1 * m2)"
29f5b20e8ee8 Added NewNumberTheory by Jeremy Avigad
nipkow
parents:
diff changeset
   628
  from nz have less: "?x < m1 * m2"
44872
a98ef45122f3 misc tuning;
wenzelm
parents: 41959
diff changeset
   629
    by auto
66380
96ff0eb8294a misc tuning and modernization;
wenzelm
parents: 65417
diff changeset
   630
  have 1: "[?x = u1] (mod m1)"
68707
5b269897df9d de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents: 67115
diff changeset
   631
    using y1 mod_mult_cong_right by blast
66380
96ff0eb8294a misc tuning and modernization;
wenzelm
parents: 65417
diff changeset
   632
  have 2: "[?x = u2] (mod m2)"
68707
5b269897df9d de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents: 67115
diff changeset
   633
    using y2 mod_mult_cong_left by blast
5b269897df9d de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents: 67115
diff changeset
   634
  have "z = ?x" if "z < m1 * m2" "[z = u1] (mod m1)"  "[z = u2] (mod m2)" for z
5b269897df9d de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents: 67115
diff changeset
   635
  proof -
31719
29f5b20e8ee8 Added NewNumberTheory by Jeremy Avigad
nipkow
parents:
diff changeset
   636
    have "[?x = z] (mod m1)"
68707
5b269897df9d de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents: 67115
diff changeset
   637
      by (metis "1" cong_def that(2))
31719
29f5b20e8ee8 Added NewNumberTheory by Jeremy Avigad
nipkow
parents:
diff changeset
   638
    moreover have "[?x = z] (mod m2)"
68707
5b269897df9d de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents: 67115
diff changeset
   639
      by (metis "2" cong_def that(3))
31719
29f5b20e8ee8 Added NewNumberTheory by Jeremy Avigad
nipkow
parents:
diff changeset
   640
    ultimately have "[?x = z] (mod m1 * m2)"
66888
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66837
diff changeset
   641
      using a by (auto intro: coprime_cong_mult_nat simp add: mod_mult_cong_left mod_mult_cong_right)
60526
fad653acf58f isabelle update_cartouches;
wenzelm
parents: 59816
diff changeset
   642
    with \<open>z < m1 * m2\<close> \<open>?x < m1 * m2\<close> show "z = ?x"
67085
f5d7f37b4143 tuned and generalized
haftmann
parents: 67051
diff changeset
   643
      by (auto simp add: cong_def)
44872
a98ef45122f3 misc tuning;
wenzelm
parents: 41959
diff changeset
   644
  qed
68707
5b269897df9d de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents: 67115
diff changeset
   645
  with less 1 2 show ?thesis
5b269897df9d de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents: 67115
diff changeset
   646
    by blast
31719
29f5b20e8ee8 Added NewNumberTheory by Jeremy Avigad
nipkow
parents:
diff changeset
   647
 qed
29f5b20e8ee8 Added NewNumberTheory by Jeremy Avigad
nipkow
parents:
diff changeset
   648
31952
40501bb2d57c renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents: 31792
diff changeset
   649
lemma chinese_remainder_nat:
44872
a98ef45122f3 misc tuning;
wenzelm
parents: 41959
diff changeset
   650
  fixes A :: "'a set"
a98ef45122f3 misc tuning;
wenzelm
parents: 41959
diff changeset
   651
    and m :: "'a \<Rightarrow> nat"
a98ef45122f3 misc tuning;
wenzelm
parents: 41959
diff changeset
   652
    and u :: "'a \<Rightarrow> nat"
a98ef45122f3 misc tuning;
wenzelm
parents: 41959
diff changeset
   653
  assumes fin: "finite A"
66380
96ff0eb8294a misc tuning and modernization;
wenzelm
parents: 65417
diff changeset
   654
    and cop: "\<forall>i \<in> A. \<forall>j \<in> A. i \<noteq> j \<longrightarrow> coprime (m i) (m j)"
96ff0eb8294a misc tuning and modernization;
wenzelm
parents: 65417
diff changeset
   655
  shows "\<exists>x. \<forall>i \<in> A. [x = u i] (mod m i)"
31719
29f5b20e8ee8 Added NewNumberTheory by Jeremy Avigad
nipkow
parents:
diff changeset
   656
proof -
67086
haftmann
parents: 67085
diff changeset
   657
  have "\<exists>b. (\<forall>i \<in> A. [b i = 1] (mod m i) \<and> [b i = 0] (mod (\<Prod>j \<in> A - {i}. m j)))"
haftmann
parents: 67085
diff changeset
   658
  proof (rule finite_set_choice, rule fin, rule ballI)
haftmann
parents: 67085
diff changeset
   659
    fix i
haftmann
parents: 67085
diff changeset
   660
    assume "i \<in> A"
haftmann
parents: 67085
diff changeset
   661
    with cop have "coprime (\<Prod>j \<in> A - {i}. m j) (m i)"
haftmann
parents: 67085
diff changeset
   662
      by (intro prod_coprime_left) auto
haftmann
parents: 67085
diff changeset
   663
    then have "\<exists>x. [(\<Prod>j \<in> A - {i}. m j) * x = Suc 0] (mod m i)"
haftmann
parents: 67085
diff changeset
   664
      by (elim cong_solve_coprime_nat)
haftmann
parents: 67085
diff changeset
   665
    then obtain x where "[(\<Prod>j \<in> A - {i}. m j) * x = 1] (mod m i)"
haftmann
parents: 67085
diff changeset
   666
      by auto
haftmann
parents: 67085
diff changeset
   667
    moreover have "[(\<Prod>j \<in> A - {i}. m j) * x = 0] (mod (\<Prod>j \<in> A - {i}. m j))"
haftmann
parents: 67085
diff changeset
   668
      by (simp add: cong_0_iff)
haftmann
parents: 67085
diff changeset
   669
    ultimately show "\<exists>a. [a = 1] (mod m i) \<and> [a = 0] (mod prod m (A - {i}))"
haftmann
parents: 67085
diff changeset
   670
      by blast
haftmann
parents: 67085
diff changeset
   671
  qed
68707
5b269897df9d de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents: 67115
diff changeset
   672
  then obtain b where b: "\<And>i. i \<in> A \<Longrightarrow> [b i = 1] (mod m i) \<and> [b i = 0] (mod (\<Prod>j \<in> A - {i}. m j))"
31719
29f5b20e8ee8 Added NewNumberTheory by Jeremy Avigad
nipkow
parents:
diff changeset
   673
    by blast
61954
1d43f86f48be more symbols;
wenzelm
parents: 60688
diff changeset
   674
  let ?x = "\<Sum>i\<in>A. (u i) * (b i)"
66380
96ff0eb8294a misc tuning and modernization;
wenzelm
parents: 65417
diff changeset
   675
  show ?thesis
31719
29f5b20e8ee8 Added NewNumberTheory by Jeremy Avigad
nipkow
parents:
diff changeset
   676
  proof (rule exI, clarify)
29f5b20e8ee8 Added NewNumberTheory by Jeremy Avigad
nipkow
parents:
diff changeset
   677
    fix i
66380
96ff0eb8294a misc tuning and modernization;
wenzelm
parents: 65417
diff changeset
   678
    assume a: "i \<in> A"
44872
a98ef45122f3 misc tuning;
wenzelm
parents: 41959
diff changeset
   679
    show "[?x = u i] (mod m i)"
31719
29f5b20e8ee8 Added NewNumberTheory by Jeremy Avigad
nipkow
parents:
diff changeset
   680
    proof -
66380
96ff0eb8294a misc tuning and modernization;
wenzelm
parents: 65417
diff changeset
   681
      from fin a have "?x = (\<Sum>j \<in> {i}. u j * b j) + (\<Sum>j \<in> A - {i}. u j * b j)"
96ff0eb8294a misc tuning and modernization;
wenzelm
parents: 65417
diff changeset
   682
        by (subst sum.union_disjoint [symmetric]) (auto intro: sum.cong)
61954
1d43f86f48be more symbols;
wenzelm
parents: 60688
diff changeset
   683
      then have "[?x = u i * b i + (\<Sum>j \<in> A - {i}. u j * b j)] (mod m i)"
31719
29f5b20e8ee8 Added NewNumberTheory by Jeremy Avigad
nipkow
parents:
diff changeset
   684
        by auto
61954
1d43f86f48be more symbols;
wenzelm
parents: 60688
diff changeset
   685
      also have "[u i * b i + (\<Sum>j \<in> A - {i}. u j * b j) =
1d43f86f48be more symbols;
wenzelm
parents: 60688
diff changeset
   686
                  u i * 1 + (\<Sum>j \<in> A - {i}. u j * 0)] (mod m i)"
68707
5b269897df9d de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents: 67115
diff changeset
   687
      proof (intro cong_add cong_scalar_left cong_sum)
5b269897df9d de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents: 67115
diff changeset
   688
        show "[b i = 1] (mod m i)"
5b269897df9d de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents: 67115
diff changeset
   689
          using a b by blast
5b269897df9d de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents: 67115
diff changeset
   690
        show "[b x = 0] (mod m i)" if "x \<in> A - {i}" for x
5b269897df9d de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents: 67115
diff changeset
   691
        proof -
5b269897df9d de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents: 67115
diff changeset
   692
          have "x \<in> A" "x \<noteq> i"
5b269897df9d de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents: 67115
diff changeset
   693
            using that by auto
5b269897df9d de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents: 67115
diff changeset
   694
          then show ?thesis
5b269897df9d de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents: 67115
diff changeset
   695
            using a b [OF \<open>x \<in> A\<close>] cong_dvd_modulus_nat fin by blast
5b269897df9d de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents: 67115
diff changeset
   696
        qed
5b269897df9d de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents: 67115
diff changeset
   697
      qed
31719
29f5b20e8ee8 Added NewNumberTheory by Jeremy Avigad
nipkow
parents:
diff changeset
   698
      finally show ?thesis
29f5b20e8ee8 Added NewNumberTheory by Jeremy Avigad
nipkow
parents:
diff changeset
   699
        by simp
29f5b20e8ee8 Added NewNumberTheory by Jeremy Avigad
nipkow
parents:
diff changeset
   700
    qed
29f5b20e8ee8 Added NewNumberTheory by Jeremy Avigad
nipkow
parents:
diff changeset
   701
  qed
29f5b20e8ee8 Added NewNumberTheory by Jeremy Avigad
nipkow
parents:
diff changeset
   702
qed
29f5b20e8ee8 Added NewNumberTheory by Jeremy Avigad
nipkow
parents:
diff changeset
   703
68707
5b269897df9d de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents: 67115
diff changeset
   704
lemma coprime_cong_prod_nat: "[x = y] (mod (\<Prod>i\<in>A. m i))"
5b269897df9d de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents: 67115
diff changeset
   705
  if "\<And>i j. \<lbrakk>i \<in> A; j \<in> A; i \<noteq> j\<rbrakk> \<Longrightarrow> coprime (m i) (m j)"
5b269897df9d de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents: 67115
diff changeset
   706
    and "\<And>i. i \<in> A \<Longrightarrow> [x = y] (mod m i)" for x y :: nat
5b269897df9d de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents: 67115
diff changeset
   707
  using that 
5b269897df9d de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents: 67115
diff changeset
   708
proof (induct A rule: infinite_finite_induct)
5b269897df9d de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents: 67115
diff changeset
   709
  case (insert x A)
5b269897df9d de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents: 67115
diff changeset
   710
  then show ?case
5b269897df9d de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents: 67115
diff changeset
   711
    by simp (metis coprime_cong_mult_nat prod_coprime_right)
5b269897df9d de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents: 67115
diff changeset
   712
qed auto
31719
29f5b20e8ee8 Added NewNumberTheory by Jeremy Avigad
nipkow
parents:
diff changeset
   713
31952
40501bb2d57c renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents: 31792
diff changeset
   714
lemma chinese_remainder_unique_nat:
44872
a98ef45122f3 misc tuning;
wenzelm
parents: 41959
diff changeset
   715
  fixes A :: "'a set"
a98ef45122f3 misc tuning;
wenzelm
parents: 41959
diff changeset
   716
    and m :: "'a \<Rightarrow> nat"
a98ef45122f3 misc tuning;
wenzelm
parents: 41959
diff changeset
   717
    and u :: "'a \<Rightarrow> nat"
a98ef45122f3 misc tuning;
wenzelm
parents: 41959
diff changeset
   718
  assumes fin: "finite A"
61954
1d43f86f48be more symbols;
wenzelm
parents: 60688
diff changeset
   719
    and nz: "\<forall>i\<in>A. m i \<noteq> 0"
66380
96ff0eb8294a misc tuning and modernization;
wenzelm
parents: 65417
diff changeset
   720
    and cop: "\<forall>i\<in>A. \<forall>j\<in>A. i \<noteq> j \<longrightarrow> coprime (m i) (m j)"
63901
4ce989e962e0 more symbols;
wenzelm
parents: 63167
diff changeset
   721
  shows "\<exists>!x. x < (\<Prod>i\<in>A. m i) \<and> (\<forall>i\<in>A. [x = u i] (mod m i))"
31719
29f5b20e8ee8 Added NewNumberTheory by Jeremy Avigad
nipkow
parents:
diff changeset
   722
proof -
44872
a98ef45122f3 misc tuning;
wenzelm
parents: 41959
diff changeset
   723
  from chinese_remainder_nat [OF fin cop]
66380
96ff0eb8294a misc tuning and modernization;
wenzelm
parents: 65417
diff changeset
   724
  obtain y where one: "(\<forall>i\<in>A. [y = u i] (mod m i))"
31719
29f5b20e8ee8 Added NewNumberTheory by Jeremy Avigad
nipkow
parents:
diff changeset
   725
    by blast
61954
1d43f86f48be more symbols;
wenzelm
parents: 60688
diff changeset
   726
  let ?x = "y mod (\<Prod>i\<in>A. m i)"
1d43f86f48be more symbols;
wenzelm
parents: 60688
diff changeset
   727
  from fin nz have prodnz: "(\<Prod>i\<in>A. m i) \<noteq> 0"
31719
29f5b20e8ee8 Added NewNumberTheory by Jeremy Avigad
nipkow
parents:
diff changeset
   728
    by auto
61954
1d43f86f48be more symbols;
wenzelm
parents: 60688
diff changeset
   729
  then have less: "?x < (\<Prod>i\<in>A. m i)"
31719
29f5b20e8ee8 Added NewNumberTheory by Jeremy Avigad
nipkow
parents:
diff changeset
   730
    by auto
66380
96ff0eb8294a misc tuning and modernization;
wenzelm
parents: 65417
diff changeset
   731
  have cong: "\<forall>i\<in>A. [?x = u i] (mod m i)"
67085
f5d7f37b4143 tuned and generalized
haftmann
parents: 67051
diff changeset
   732
    using fin one
f5d7f37b4143 tuned and generalized
haftmann
parents: 67051
diff changeset
   733
    by (auto simp add: cong_def dvd_prod_eqI mod_mod_cancel) 
66380
96ff0eb8294a misc tuning and modernization;
wenzelm
parents: 65417
diff changeset
   734
  have unique: "\<forall>z. z < (\<Prod>i\<in>A. m i) \<and> (\<forall>i\<in>A. [z = u i] (mod m i)) \<longrightarrow> z = ?x"
96ff0eb8294a misc tuning and modernization;
wenzelm
parents: 65417
diff changeset
   735
  proof clarify
31719
29f5b20e8ee8 Added NewNumberTheory by Jeremy Avigad
nipkow
parents:
diff changeset
   736
    fix z
61954
1d43f86f48be more symbols;
wenzelm
parents: 60688
diff changeset
   737
    assume zless: "z < (\<Prod>i\<in>A. m i)"
66380
96ff0eb8294a misc tuning and modernization;
wenzelm
parents: 65417
diff changeset
   738
    assume zcong: "(\<forall>i\<in>A. [z = u i] (mod m i))"
96ff0eb8294a misc tuning and modernization;
wenzelm
parents: 65417
diff changeset
   739
    have "\<forall>i\<in>A. [?x = z] (mod m i)"
67085
f5d7f37b4143 tuned and generalized
haftmann
parents: 67051
diff changeset
   740
      using cong zcong by (auto simp add: cong_def)
61954
1d43f86f48be more symbols;
wenzelm
parents: 60688
diff changeset
   741
    with fin cop have "[?x = z] (mod (\<Prod>i\<in>A. m i))"
67085
f5d7f37b4143 tuned and generalized
haftmann
parents: 67051
diff changeset
   742
      by (intro coprime_cong_prod_nat) auto
31719
29f5b20e8ee8 Added NewNumberTheory by Jeremy Avigad
nipkow
parents:
diff changeset
   743
    with zless less show "z = ?x"
67085
f5d7f37b4143 tuned and generalized
haftmann
parents: 67051
diff changeset
   744
      by (auto simp add: cong_def)
44872
a98ef45122f3 misc tuning;
wenzelm
parents: 41959
diff changeset
   745
  qed
66380
96ff0eb8294a misc tuning and modernization;
wenzelm
parents: 65417
diff changeset
   746
  from less cong unique show ?thesis
96ff0eb8294a misc tuning and modernization;
wenzelm
parents: 65417
diff changeset
   747
    by blast
31719
29f5b20e8ee8 Added NewNumberTheory by Jeremy Avigad
nipkow
parents:
diff changeset
   748
qed
29f5b20e8ee8 Added NewNumberTheory by Jeremy Avigad
nipkow
parents:
diff changeset
   749
80084
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 76231
diff changeset
   750
lemma (in semiring_1_cancel) of_nat_eq_iff_cong_CHAR:
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 76231
diff changeset
   751
  "of_nat x = (of_nat y :: 'a) \<longleftrightarrow> [x = y] (mod CHAR('a))"
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 76231
diff changeset
   752
proof (induction x y rule: linorder_wlog)
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 76231
diff changeset
   753
  case (le x y)
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 76231
diff changeset
   754
  define z where "z = y - x"
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 76231
diff changeset
   755
  have [simp]: "y = x + z"
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 76231
diff changeset
   756
    using le by (auto simp: z_def)
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 76231
diff changeset
   757
  have "(CHAR('a) dvd z) = [x = x + z] (mod CHAR('a))"
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 76231
diff changeset
   758
    by (metis \<open>y = x + z\<close> cong_def le mod_eq_dvd_iff_nat z_def)
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 76231
diff changeset
   759
  thus ?case
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 76231
diff changeset
   760
    by (simp add: of_nat_eq_0_iff_char_dvd)
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 76231
diff changeset
   761
qed (simp add: eq_commute cong_sym_eq)   
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 76231
diff changeset
   762
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 76231
diff changeset
   763
lemma (in ring_1) of_int_eq_iff_cong_CHAR:
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 76231
diff changeset
   764
  "of_int x = (of_int y :: 'a) \<longleftrightarrow> [x = y] (mod int CHAR('a))"
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 76231
diff changeset
   765
proof -
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 76231
diff changeset
   766
  have "of_int x = (of_int y :: 'a) \<longleftrightarrow> of_int (x - y) = (0 :: 'a)"
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 76231
diff changeset
   767
    by auto
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 76231
diff changeset
   768
  also have "\<dots> \<longleftrightarrow> (int CHAR('a) dvd x - y)"
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 76231
diff changeset
   769
    by (rule of_int_eq_0_iff_char_dvd)
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 76231
diff changeset
   770
  also have "\<dots> \<longleftrightarrow> [x = y] (mod int CHAR('a))"
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 76231
diff changeset
   771
    by (simp add: cong_iff_dvd_diff)
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 76231
diff changeset
   772
  finally show ?thesis .
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 76231
diff changeset
   773
qed
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 76231
diff changeset
   774
31719
29f5b20e8ee8 Added NewNumberTheory by Jeremy Avigad
nipkow
parents:
diff changeset
   775
end