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

Commit 03d906c

Browse files
feat(field_theory/finite/basic): zmod.pow_totient is true for zero (#15771)
1 parent 286d675 commit 03d906c

File tree

1 file changed

+8
-5
lines changed

1 file changed

+8
-5
lines changed

src/field_theory/finite/basic.lean

Lines changed: 8 additions & 5 deletions
Original file line numberDiff line numberDiff line change
@@ -353,18 +353,21 @@ open zmod
353353

354354
/-- The **Fermat-Euler totient theorem**. `nat.modeq.pow_totient` is an alternative statement
355355
of the same theorem. -/
356-
@[simp] lemma zmod.pow_totient {n : ℕ} [fact (0 < n)] (x : (zmod n)ˣ) : x ^ φ n = 1 :=
357-
by rw [← card_units_eq_totient, pow_card_eq_one]
356+
@[simp] lemma zmod.pow_totient {n : ℕ} (x : (zmod n)ˣ) : x ^ φ n = 1 :=
357+
begin
358+
cases n,
359+
{ rw [nat.totient_zero, pow_zero] },
360+
{ rw [← card_units_eq_totient, pow_card_eq_one] }
361+
end
358362

359363
/-- The **Fermat-Euler totient theorem**. `zmod.pow_totient` is an alternative statement
360364
of the same theorem. -/
361365
lemma nat.modeq.pow_totient {x n : ℕ} (h : nat.coprime x n) : x ^ φ n ≡ 1 [MOD n] :=
362366
begin
363-
cases n, {simp},
364367
rw ← zmod.eq_iff_modeq_nat,
365-
let x' : units (zmod (n+1)) := zmod.unit_of_coprime _ h,
368+
let x' : units (zmod n) := zmod.unit_of_coprime _ h,
366369
have := zmod.pow_totient x',
367-
apply_fun (coe : units (zmod (n+1)) → zmod (n+1)) at this,
370+
apply_fun (coe : units (zmod n) → zmod n) at this,
368371
simpa only [-zmod.pow_totient, nat.succ_eq_add_one, nat.cast_pow, units.coe_one,
369372
nat.cast_one, coe_unit_of_coprime, units.coe_pow],
370373
end

0 commit comments

Comments
 (0)