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

Commit 04f8fd7

Browse files
feat(group_theory/dihedral): add dihedral groups (#5171)
Contains a subset of the content of #1076 , but implemented slightly differently. In #1076, finite and infinite dihedral groups are implemented separately, but a side effect of what I did was that `dihedral 0` corresponds to the infinite dihedral group. Co-authored-by: Bryan Gin-ge Chen <bryangingechen@gmail.com>
1 parent ee2c963 commit 04f8fd7

File tree

1 file changed

+178
-0
lines changed

1 file changed

+178
-0
lines changed

src/group_theory/dihedral.lean

Lines changed: 178 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,178 @@
1+
/-
2+
Copyright (c) 2020 Shing Tak Lam. All rights reserved.
3+
Released under Apache 2.0 license as described in the file LICENSE.
4+
Authors: Shing Tak Lam
5+
-/
6+
import data.zmod.basic
7+
import group_theory.order_of_element
8+
9+
/-!
10+
# Dihedral Groups
11+
12+
We define the dihedral groups `dihedral n`, with elements `r i` and `sr i` for `i : zmod n`.
13+
14+
For `n ≠ 0`, `dihedral n` represents the symmetry group of the regular `n`-gon. `r i` represents
15+
the rotations of the `n`-gon by `2πi/n`, and `sr i` represents the reflections of the `n`-gon.
16+
`dihedral 0` corresponds to the infinite dihedral group.
17+
-/
18+
19+
/--
20+
For `n ≠ 0`, `dihedral n` represents the symmetry group of the regular `n`-gon. `r i` represents
21+
the rotations of the `n`-gon by `2πi/n`, and `sr i` represents the reflections of the `n`-gon.
22+
`dihedral 0` corresponds to the infinite dihedral group.
23+
-/
24+
@[derive decidable_eq]
25+
inductive dihedral (n : ℕ) : Type
26+
| r : zmod n → dihedral
27+
| sr : zmod n → dihedral
28+
29+
namespace dihedral
30+
31+
variables {n : ℕ}
32+
33+
/--
34+
Multiplication of the dihedral group.
35+
-/
36+
private def mul : dihedral n → dihedral n → dihedral n
37+
| (r i) (r j) := r (i + j)
38+
| (r i) (sr j) := sr (j - i)
39+
| (sr i) (r j) := sr (i + j)
40+
| (sr i) (sr j) := r (j - i)
41+
42+
/--
43+
The identity `1` is the rotation by `0`.
44+
-/
45+
private def one : dihedral n := r 0
46+
47+
instance : inhabited (dihedral n) := ⟨one⟩
48+
49+
/--
50+
The inverse of a an element of the dihedral group.
51+
-/
52+
private def inv : dihedral n → dihedral n
53+
| (r i) := r (-i)
54+
| (sr i) := sr i
55+
56+
/--
57+
The group structure on `dihedral n`.
58+
-/
59+
instance : group (dihedral n) :=
60+
{ mul := mul,
61+
mul_assoc :=
62+
begin
63+
rintros (a | a) (b | b) (c | c);
64+
simp only [mul];
65+
ring,
66+
end,
67+
one := one,
68+
one_mul :=
69+
begin
70+
rintros (a | a),
71+
exact congr_arg r (zero_add a),
72+
exact congr_arg sr (sub_zero a),
73+
end,
74+
mul_one := begin
75+
rintros (a | a),
76+
exact congr_arg r (add_zero a),
77+
exact congr_arg sr (add_zero a),
78+
end,
79+
inv := inv,
80+
mul_left_inv := begin
81+
rintros (a | a),
82+
exact congr_arg r (neg_add_self a),
83+
exact congr_arg r (sub_self a),
84+
end }
85+
86+
@[simp] lemma r_mul_r (i j : zmod n) : r i * r j = r (i + j) := rfl
87+
@[simp] lemma r_mul_sr (i j : zmod n) : r i * sr j = sr (j - i) := rfl
88+
@[simp] lemma sr_mul_r (i j : zmod n) : sr i * r j = sr (i + j) := rfl
89+
@[simp] lemma sr_mul_sr (i j : zmod n) : sr i * sr j = r (j - i) := rfl
90+
91+
lemma one_def : (1 : dihedral n) = r 0 := rfl
92+
93+
private def fintype_helper : (zmod n ⊕ zmod n) ≃ dihedral n :=
94+
{ inv_fun := λ i, match i with
95+
| (r j) := sum.inl j
96+
| (sr j) := sum.inr j
97+
end,
98+
to_fun := λ i, match i with
99+
| (sum.inl j) := r j
100+
| (sum.inr j) := sr j
101+
end,
102+
left_inv := by rintro (x | x); refl,
103+
right_inv := by rintro (x | x); refl }
104+
105+
/--
106+
If `0 < n`, then `dihedral n` is a finite group.
107+
-/
108+
instance [fact (0 < n)] : fintype (dihedral n) := fintype.of_equiv _ fintype_helper
109+
110+
instance : nontrivial (dihedral n) := ⟨⟨r 0, sr 0, dec_trivial⟩⟩
111+
112+
/--
113+
If `0 < n`, then `dihedral n` has `2n` elements.
114+
-/
115+
lemma card [fact (0 < n)] : fintype.card (dihedral n) = 2 * n :=
116+
begin
117+
rw ←fintype.card_eq.mpr ⟨fintype_helper⟩,
118+
change fintype.card (zmod n ⊕ zmod n) = 2 * n,
119+
rw [fintype.card_sum, zmod.card, two_mul]
120+
end
121+
122+
@[simp] lemma r_one_pow (k : ℕ) : (r 1 : dihedral n) ^ k = r k :=
123+
begin
124+
induction k with k IH,
125+
{ refl },
126+
{ rw [pow_succ, IH, r_mul_r],
127+
congr' 1,
128+
norm_cast,
129+
rw nat.one_add }
130+
end
131+
132+
@[simp] lemma r_one_pow_n : (r (1 : zmod n))^n = 1 :=
133+
begin
134+
cases n,
135+
{ rw pow_zero },
136+
{ rw [r_one_pow, one_def],
137+
congr' 1,
138+
exact zmod.cast_self _, }
139+
end
140+
141+
@[simp] lemma sr_mul_self (i : zmod n) : sr i * sr i = 1 := by rw [sr_mul_sr, sub_self, one_def]
142+
143+
/--
144+
If `0 < n`, then `sr i` has order 2.
145+
-/
146+
@[simp] lemma order_of_sr [fact (0 < n)] (i : zmod n) : order_of (sr i) = 2 :=
147+
begin
148+
rw order_of_eq_prime _ _,
149+
{ exact nat.prime_two },
150+
rw [pow_two, sr_mul_self],
151+
dec_trivial,
152+
end
153+
154+
/--
155+
If `0 < n`, then `r 1` has order `n`.
156+
-/
157+
@[simp] lemma order_of_r_one [hnpos : fact (0 < n)] : order_of (r 1 : dihedral n) = n :=
158+
begin
159+
cases lt_or_eq_of_le (nat.le_of_dvd hnpos (order_of_dvd_of_pow_eq_one (@r_one_pow_n n))) with h h,
160+
{ have h1 : (r 1 : dihedral n)^(order_of (r 1)) = 1,
161+
{ exact pow_order_of_eq_one _ },
162+
rw r_one_pow at h1,
163+
injection h1 with h2,
164+
rw [←zmod.val_eq_zero, zmod.val_cast_nat, nat.mod_eq_of_lt h] at h2,
165+
exact absurd h2.symm (ne_of_lt (order_of_pos _)) },
166+
{ exact h }
167+
end
168+
169+
/--
170+
If `0 < n`, then `i : zmod n` has order `n / gcd n i`
171+
-/
172+
lemma order_of_r [fact (0 < n)] (i : zmod n) : order_of (r i) = n / nat.gcd n i.val :=
173+
begin
174+
conv_lhs { rw ←zmod.cast_val i },
175+
rw [←r_one_pow, order_of_pow, order_of_r_one]
176+
end
177+
178+
end dihedral

0 commit comments

Comments
 (0)