Skip to content
This repository was archived by the owner on Jul 24, 2024. It is now read-only.

Commit 391746b

Browse files
committed
feat(data/zmod/basic): chinese remainder theorem (#8356)
Co-authored-by: Chris Hughes <33847686+ChrisHughes24@users.noreply.github.com>
1 parent 65006d2 commit 391746b

File tree

7 files changed

+145
-1
lines changed

7 files changed

+145
-1
lines changed

src/algebra/char_p/basic.lean

Lines changed: 18 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -428,3 +428,21 @@ begin
428428
end
429429

430430
end
431+
432+
section prod
433+
434+
variables (S : Type v) [semiring R] [semiring S] (p q : ℕ) [char_p R p]
435+
436+
/-- The characteristic of the product of rings is the least common multiple of the
437+
characteristics of the two rings. -/
438+
instance [char_p S q] : char_p (R × S) (nat.lcm p q) :=
439+
{ cast_eq_zero_iff :=
440+
by simp [prod.ext_iff, char_p.cast_eq_zero_iff R p,
441+
char_p.cast_eq_zero_iff S q, nat.lcm_dvd_iff] }
442+
443+
/-- The characteristic of the product of two rings of the same characteristic
444+
is the same as the characteristic of the rings -/
445+
instance prod.char_p [char_p S p] : char_p (R × S) p :=
446+
by convert nat.lcm.char_p R S p p; simp
447+
448+
end prod

src/algebra/ring/prod.lean

Lines changed: 20 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -141,4 +141,24 @@ ring_hom.ext $ λ _, rfl
141141
(ring_hom.snd S R).comp ↑(prod_comm : R × S ≃+* S × R) = ring_hom.fst R S :=
142142
ring_hom.ext $ λ _, rfl
143143

144+
variables (R S) [subsingleton S]
145+
146+
/-- A ring `R` is isomorphic to `R × S` when `S` is the zero ring -/
147+
@[simps] def prod_zero_ring : R ≃+* R × S :=
148+
{ to_fun := λ x, (x, 0),
149+
inv_fun := prod.fst,
150+
map_add' := by simp,
151+
map_mul' := by simp,
152+
left_inv := λ x, rfl,
153+
right_inv := λ x, by cases x; simp }
154+
155+
/-- A ring `R` is isomorphic to `S × R` when `S` is the zero ring -/
156+
@[simps] def zero_ring_prod : R ≃+* S × R :=
157+
{ to_fun := λ x, (0, x),
158+
inv_fun := prod.snd,
159+
map_add' := by simp,
160+
map_mul' := by simp,
161+
left_inv := λ x, rfl,
162+
right_inv := λ x, by cases x; simp }
163+
144164
end ring_equiv

src/data/fintype/basic.lean

Lines changed: 14 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -980,6 +980,20 @@ begin
980980
rwa injective_iff_surjective_of_equiv (equiv_of_card_eq h) }
981981
end
982982

983+
lemma right_inverse_of_left_inverse_of_card_le {f : α → β} {g : β → α}
984+
(hfg : left_inverse f g) (hcard : card α ≤ card β) :
985+
right_inverse f g :=
986+
have hsurj : surjective f, from surjective_iff_has_right_inverse.2 ⟨g, hfg⟩,
987+
right_inverse_of_injective_of_left_inverse
988+
((bijective_iff_surjective_and_card _).2
989+
⟨hsurj, le_antisymm hcard (card_le_of_surjective f hsurj)⟩ ).1
990+
hfg
991+
992+
lemma left_inverse_of_right_inverse_of_card_le {f : α → β} {g : β → α}
993+
(hfg : right_inverse f g) (hcard : card β ≤ card α) :
994+
left_inverse f g :=
995+
right_inverse_of_left_inverse_of_card_le hfg hcard
996+
983997
end fintype
984998

985999
lemma fintype.coe_image_univ [fintype α] [decidable_eq β] {f : α → β} :

src/data/int/cast.lean

Lines changed: 13 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -202,6 +202,19 @@ end cast
202202

203203
end int
204204

205+
namespace prod
206+
207+
variables {α : Type*} {β : Type*} [has_zero α] [has_one α] [has_add α] [has_neg α]
208+
[has_zero β] [has_one β] [has_add β] [has_neg β]
209+
210+
@[simp] lemma fst_int_cast (n : ℤ) : (n : α × β).fst = n :=
211+
by induction n; simp *
212+
213+
@[simp] lemma snd_int_cast (n : ℤ) : (n : α × β).snd = n :=
214+
by induction n; simp *
215+
216+
end prod
217+
205218
open int
206219

207220
namespace add_monoid_hom

src/data/nat/cast.lean

Lines changed: 13 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -239,6 +239,19 @@ end linear_ordered_field
239239

240240
end nat
241241

242+
namespace prod
243+
244+
variables {α : Type*} {β : Type*} [has_zero α] [has_one α] [has_add α]
245+
[has_zero β] [has_one β] [has_add β]
246+
247+
@[simp] lemma fst_nat_cast (n : ℕ) : (n : α × β).fst = n :=
248+
by induction n; simp *
249+
250+
@[simp] lemma snd_nat_cast (n : ℕ) : (n : α × β).snd = n :=
251+
by induction n; simp *
252+
253+
end prod
254+
242255
namespace add_monoid_hom
243256

244257
variables {A B : Type*} [add_monoid A]

src/data/nat/gcd.lean

Lines changed: 10 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -191,6 +191,10 @@ or.elim (eq_zero_or_pos k)
191191
by rw [gcd_mul_lcm, ←gcd_mul_right, mul_comm n k];
192192
exact dvd_gcd (mul_dvd_mul_left _ H2) (mul_dvd_mul_right H1 _))
193193

194+
lemma lcm_dvd_iff {m n k : ℕ} : lcm m n ∣ k ↔ m ∣ k ∧ n ∣ k :=
195+
⟨λ h, ⟨dvd_trans (dvd_lcm_left _ _) h, dvd_trans (dvd_lcm_right _ _) h⟩,
196+
and_imp.2 lcm_dvd⟩
197+
194198
theorem lcm_assoc (m n k : ℕ) : lcm (lcm m n) k = lcm m (lcm n k) :=
195199
dvd_antisymm
196200
(lcm_dvd
@@ -358,6 +362,12 @@ by simp [coprime]
358362
@[simp] theorem coprime_self (n : ℕ) : coprime n n ↔ n = 1 :=
359363
by simp [coprime]
360364

365+
lemma coprime.eq_of_mul_eq_zero {m n : ℕ} (h : m.coprime n) (hmn : m * n = 0) :
366+
m = 0 ∧ n = 1 ∨ m = 1 ∧ n = 0 :=
367+
(nat.eq_zero_of_mul_eq_zero hmn).imp
368+
(λ hm, ⟨hm, n.coprime_zero_left.mp $ hm ▸ h⟩)
369+
(λ hn, ⟨m.coprime_zero_left.mp $ hn ▸ h.symm, hn⟩)
370+
361371
/-- Represent a divisor of `m * n` as a product of a divisor of `m` and a divisor of `n`. -/
362372
def prod_dvd_and_dvd_of_dvd_prod {m n k : ℕ} (H : k ∣ m * n) :
363373
{ d : {m' // m' ∣ m} × {n' // n' ∣ n} // k = d.1 * d.2 } :=

src/data/zmod/basic.lean

Lines changed: 57 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -83,7 +83,7 @@ instance fintype : Π (n : ℕ) [fact (0 < n)], fintype (zmod n)
8383
| 0 h := false.elim $ nat.not_lt_zero 0 h.1
8484
| (n+1) _ := fin.fintype (n+1)
8585

86-
lemma card (n : ℕ) [fact (0 < n)] : fintype.card (zmod n) = n :=
86+
@[simp] lemma card (n : ℕ) [fact (0 < n)] : fintype.card (zmod n) = n :=
8787
begin
8888
casesI n,
8989
{ exfalso, exact nat.not_lt_zero 0 (fact.out _) },
@@ -176,6 +176,14 @@ def cast : Π {n : ℕ}, zmod n → R
176176
@[simp] lemma cast_zero : ((0 : zmod n) : R) = 0 :=
177177
by { cases n; refl }
178178

179+
variables {S : Type*} [has_zero S] [has_one S] [has_add S] [has_neg S]
180+
181+
@[simp] lemma _root_.prod.fst_zmod_cast (a : zmod n) : (a : R × S).fst = a :=
182+
by cases n; simp
183+
184+
@[simp] lemma _root_.prod.snd_zmod_cast (a : zmod n) : (a : R × S).snd = a :=
185+
by cases n; simp
186+
179187
end
180188

181189
/-- So-named because the coercion is `nat.cast` into `zmod`. For `nat.cast` into an arbitrary ring,
@@ -577,6 +585,54 @@ def units_equiv_coprime {n : ℕ} [fact (0 < n)] :
577585
left_inv := λ ⟨_, _, _, _⟩, units.ext (nat_cast_zmod_val _),
578586
right_inv := λ ⟨_, _⟩, by simp }
579587

588+
/-- The **Chinese remainder theorem**. For a pair of coprime natural numbers, `m` and `n`,
589+
the rings `zmod (m * n)` and `zmod m × zmod n` are isomorphic.
590+
591+
See `ideal.quotient_inf_ring_equiv_pi_quotient` for the Chinese remainder theorem for ideals in any
592+
ring.
593+
-/
594+
def chinese_remainder {m n : ℕ} (h : m.coprime n) :
595+
zmod (m * n) ≃+* zmod m × zmod n :=
596+
let to_fun : zmod (m * n) → zmod m × zmod n :=
597+
zmod.cast_hom (show m.lcm n ∣ m * n, by simp [nat.lcm_dvd_iff]) (zmod m × zmod n) in
598+
let inv_fun : zmod m × zmod n → zmod (m * n) :=
599+
λ x, if m * n = 0
600+
then if m = 1
601+
then ring_hom.snd _ _ x
602+
else ring_hom.fst _ _ x
603+
else nat.modeq.chinese_remainder h x.1.val x.2.val in
604+
have inv : function.left_inverse inv_fun to_fun ∧ function.right_inverse inv_fun to_fun :=
605+
if hmn0 : m * n = 0
606+
then begin
607+
rcases h.eq_of_mul_eq_zero hmn0 with ⟨rfl, rfl⟩ | ⟨rfl, rfl⟩;
608+
simp [inv_fun, to_fun, function.left_inverse, function.right_inverse,
609+
ring_hom.eq_int_cast, prod.ext_iff]
610+
end
611+
else
612+
begin
613+
haveI : fact (0 < (m * n)) := ⟨nat.pos_of_ne_zero hmn0⟩,
614+
haveI : fact (0 < m) := ⟨nat.pos_of_ne_zero $ left_ne_zero_of_mul hmn0⟩,
615+
haveI : fact (0 < n) := ⟨nat.pos_of_ne_zero $ right_ne_zero_of_mul hmn0⟩,
616+
have left_inv : function.left_inverse inv_fun to_fun,
617+
{ intro x,
618+
dsimp only [dvd_mul_left, dvd_mul_right, zmod.cast_hom_apply, coe_coe, inv_fun, to_fun],
619+
conv_rhs { rw ← zmod.nat_cast_zmod_val x },
620+
rw [if_neg hmn0, zmod.eq_iff_modeq_nat, ← nat.modeq.modeq_and_modeq_iff_modeq_mul h,
621+
prod.fst_zmod_cast, prod.snd_zmod_cast],
622+
refine
623+
⟨(nat.modeq.chinese_remainder h (x : zmod m).val (x : zmod n).val).2.left.trans _,
624+
(nat.modeq.chinese_remainder h (x : zmod m).val (x : zmod n).val).2.right.trans _⟩,
625+
{ rw [← zmod.eq_iff_modeq_nat, zmod.nat_cast_zmod_val, zmod.nat_cast_val] },
626+
{ rw [← zmod.eq_iff_modeq_nat, zmod.nat_cast_zmod_val, zmod.nat_cast_val] } },
627+
exact ⟨left_inv, fintype.right_inverse_of_left_inverse_of_card_le left_inv (by simp)⟩,
628+
end,
629+
{ to_fun := to_fun,
630+
inv_fun := inv_fun,
631+
map_mul' := ring_hom.map_mul _,
632+
map_add' := ring_hom.map_add _,
633+
left_inv := inv.1,
634+
right_inv := inv.2 }
635+
580636
section totient
581637
open_locale nat
582638

0 commit comments

Comments
 (0)