@@ -83,7 +83,7 @@ instance fintype : Π (n : ℕ) [fact (0 < n)], fintype (zmod n)
83
83
| 0 h := false.elim $ nat.not_lt_zero 0 h.1
84
84
| (n+1 ) _ := fin.fintype (n+1 )
85
85
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 :=
87
87
begin
88
88
casesI n,
89
89
{ exfalso, exact nat.not_lt_zero 0 (fact.out _) },
@@ -176,6 +176,14 @@ def cast : Π {n : ℕ}, zmod n → R
176
176
@[simp] lemma cast_zero : ((0 : zmod n) : R) = 0 :=
177
177
by { cases n; refl }
178
178
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
+
179
187
end
180
188
181
189
/-- 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)] :
577
585
left_inv := λ ⟨_, _, _, _⟩, units.ext (nat_cast_zmod_val _),
578
586
right_inv := λ ⟨_, _⟩, by simp }
579
587
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
+
580
636
section totient
581
637
open_locale nat
582
638
0 commit comments