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

Commit 2a3cd41

Browse files
committed
feat(group_theory/free_product): equivalence with reduced words (#7395)
We show that each element of the free product is represented by a unique reduced word.
1 parent 7deb32c commit 2a3cd41

File tree

1 file changed

+186
-8
lines changed

1 file changed

+186
-8
lines changed

src/group_theory/free_product.lean

Lines changed: 186 additions & 8 deletions
Original file line numberDiff line numberDiff line change
@@ -5,10 +5,14 @@ Authors: David Wärn
55
-/
66
import algebra.free_monoid
77
import group_theory.congruence
8+
import data.list.chain
89
/-!
910
# The free product of groups or monoids
11+
1012
Given an `ι`-indexed family `M` of monoids, we define their free product (categorical coproduct)
11-
`free_product M`.
13+
`free_product M`. When `ι` and all `M i` have decidable equality, the free product bijects with the
14+
type `word M` of reduced words. This bijection is constructed by defining an action of
15+
`free_product M` on `word M`.
1216
1317
When `M i` are all groups, `free_product M` is also a group (and the coproduct in the category of
1418
groups).
@@ -18,21 +22,31 @@ groups).
1822
- `free_product M`: the free product, defined as a quotient of a free monoid.
1923
- `free_product.of {i} : M i →* free_product M`.
2024
- `free_product.lift : (Π {i}, M i →* N) ≃ (free_product M →* N)`: the universal property.
25+
- `free_product.word M`: the type of reduced words.
26+
- `free_product.word.equiv M : free_product M ≃ word M`.
2127
2228
## Remarks
2329
2430
There are many answers to the question "what is the free product of a family `M` of monoids?", and
25-
they are all equivalent but not obviously equivalent. We provide one almost tautological answer,
26-
namely `free_product M`, which is a quotient of the type of words in the alphabet `Σ i, M i`. It's
27-
straightforward to define and easy to prove its universal property. But this answer is not
28-
completely satisfactory, because it's difficult to tell when two elements `x y : free_product M` are
29-
distinct since `free_product M` is defined as a quotient. Soon a second answer will be given, in
30-
terms of reduced words, which lets you show that distinct elements are distinct.
31+
they are all equivalent but not obviously equivalent. We provide two answers. The first, almost
32+
tautological answer is given by `free_product M`, which is a quotient of the type of words in the
33+
alphabet `Σ i, M i`. It's straightforward to define and easy to prove its universal property. But
34+
this answer is not completely satisfactory, because it's difficult to tell when two elements
35+
`x y : free_product M` are distinct since `free_product M` is defined as a quotient.
36+
37+
The second, maximally efficient answer is given by `word M`. An element of `word M` is a word in the
38+
alphabet `Σ i, M i`, where the letter `⟨i, 1⟩` doesn't occur and no adjacent letters share an index
39+
`i`. Since we only work with reduced words, there is no need for quotienting, and it is easy to tell
40+
when two elements are distinct. However it's not obvious that this is even a monoid!
41+
42+
We prove that every element of `free_product M` can be represented by a unique reduced word, i.e.
43+
`free_product M` and `word M` are equivalent types. This means that `word M` can be given a monoid
44+
structure, and it lets us tell when two elements of `free_product M` are distinct.
3145
3246
There is also a completely tautological, maximally inefficient answer given by
3347
`algebra.category.Mon.colimits`. Whereas `free_product M` at least ensures that (any instance of)
3448
associativity holds by reflexivity, in this answer associativity holds because of quotienting. Yet
35-
another answer, which is constructively more satisfying, can be obtained by showing that
49+
another answer, which is constructively more satisfying, could be obtained by showing that
3650
`free_product.rel` is confluent.
3751
3852
## References
@@ -56,6 +70,13 @@ def free_product : Type* := (con_gen (free_product.rel M)).quotient
5670

5771
namespace free_product
5872

73+
/-- The type of reduced words. A reduced word cannot contain a letter `1`, and no two adjacent
74+
letters can come from the same summand. -/
75+
@[ext] structure word :=
76+
(to_list : list (Σ i, M i))
77+
(ne_one : ∀ l ∈ to_list, sigma.snd l ≠ 1)
78+
(chain_ne : to_list.chain' (λ l l', sigma.fst l ≠ sigma.fst l'))
79+
5980
variable {M}
6081

6182
/-- The inclusion of a summand into the free product. -/
@@ -147,4 +168,161 @@ instance : group (free_product G) :=
147168

148169
end group
149170

171+
namespace word
172+
173+
/-- The empty reduced word. -/
174+
def empty : word M := { to_list := [], ne_one := λ _, false.elim, chain_ne := list.chain'_nil }
175+
176+
instance : inhabited (word M) := ⟨empty⟩
177+
178+
/-- A reduced word determines an element of the free product, given by multiplication. -/
179+
def prod (w : word M) : free_product M :=
180+
list.prod (w.to_list.map $ λ l, of l.snd)
181+
182+
@[simp] lemma prod_nil : prod (empty : word M) = 1 := rfl
183+
184+
/-- `fst_idx w` is `some i` if the first letter of `w` is `⟨i, m⟩` with `m : M i`. If `w` is empty
185+
then it's `none`. -/
186+
def fst_idx (w : word M) : option ι := w.to_list.head'.map sigma.fst
187+
188+
lemma fst_idx_ne_iff {w : word M} {i} :
189+
fst_idx w ≠ some i ↔ ∀ l ∈ w.to_list.head', i ≠ sigma.fst l :=
190+
not_iff_not.mp $ by simp [fst_idx]
191+
192+
variable (M)
193+
194+
/-- Given an index `i : ι`, `pair M i` is the type of pairs `(head, tail)` where `head : M i` and
195+
`tail : word M`, subject to the constraint that first letter of `tail` can't be `⟨i, m⟩`.
196+
By prepending `head` to `tail`, one obtains a new word. We'll show that any word can be uniquely
197+
obtained in this way. -/
198+
@[ext] structure pair (i : ι) :=
199+
(head : M i)
200+
(tail : word M)
201+
(fst_idx_ne : fst_idx tail ≠ some i)
202+
203+
instance (i : ι) : inhabited (pair M i) := ⟨⟨1, empty, by tauto⟩⟩
204+
205+
variable {M}
206+
207+
variables [∀ i, decidable_eq (M i)]
208+
209+
/-- Given a pair `(head, tail)`, we can form a word by prepending `head` to `tail`, except if `head`
210+
is `1 : M i` then we have to just return `word` since we need the result to be reduced. -/
211+
def rcons {i} (p : pair M i) : word M :=
212+
if h : p.head = 1 then p.tail
213+
else { to_list := ⟨i, p.head⟩ :: p.tail.to_list,
214+
ne_one := by { rintros l (rfl | hl), exact h, exact p.tail.ne_one l hl },
215+
chain_ne := p.tail.chain_ne.cons' (fst_idx_ne_iff.mp p.fst_idx_ne) }
216+
217+
/-- Given a word of the form `⟨l :: ls, h1, h2⟩`, we can form a word of the form `⟨ls, _, _⟩`,
218+
dropping the first letter. -/
219+
private def mk_aux {l} (ls : list (Σ i, M i)) (h1 : ∀ l' ∈ l :: ls, sigma.snd l' ≠ 1)
220+
(h2 : (l :: ls).chain' _) : word M :=
221+
⟨ls, λ l' hl, h1 _ (list.mem_cons_of_mem _ hl), h2.tail⟩
222+
223+
lemma cons_eq_rcons {i} {m : M i} {ls h1 h2} :
224+
word.mk (⟨i, m⟩ :: ls) h1 h2 = rcons ⟨m, mk_aux ls h1 h2, fst_idx_ne_iff.mpr h2.rel_head'⟩ :=
225+
by { rw [rcons, dif_neg], refl, exact h1 ⟨i, m⟩ (ls.mem_cons_self _) }
226+
227+
@[simp] lemma prod_rcons {i} (p : pair M i) :
228+
prod (rcons p) = of p.head * prod p.tail :=
229+
if hm : p.head = 1 then by rw [rcons, dif_pos hm, hm, monoid_hom.map_one, one_mul]
230+
else by rw [rcons, dif_neg hm, prod, list.map_cons, list.prod_cons, prod]
231+
232+
lemma rcons_inj {i} : function.injective (rcons : pair M i → word M) :=
233+
begin
234+
rintros ⟨m, w, h⟩ ⟨m', w', h'⟩ he,
235+
by_cases hm : m = 1;
236+
by_cases hm' : m' = 1,
237+
{ simp only [rcons, dif_pos hm, dif_pos hm'] at he, cc, },
238+
{ exfalso, simp only [rcons, dif_pos hm, dif_neg hm'] at he, rw he at h, exact h rfl },
239+
{ exfalso, simp only [rcons, dif_pos hm', dif_neg hm] at he, rw ←he at h', exact h' rfl, },
240+
{ have : m = m' ∧ w.to_list = w'.to_list,
241+
{ simpa only [rcons, dif_neg hm, dif_neg hm', true_and, eq_self_iff_true, subtype.mk_eq_mk,
242+
heq_iff_eq, ←subtype.ext_iff_val] using he },
243+
rcases this with ⟨rfl, h⟩,
244+
congr, exact word.ext _ _ h, }
245+
end
246+
247+
variable [decidable_eq ι]
248+
249+
/-- Given `i : ι`, any reduced word can be decomposed into a pair `p` such that `w = rcons p`. -/
250+
-- This definition is computable but not very nice to look at. Thankfully we don't have to inspect
251+
-- it, since `rcons` is known to be injective.
252+
private def equiv_pair_aux (i) : Π w : word M, { p : pair M i // rcons p = w }
253+
| w@⟨[], _, _⟩ := ⟨⟨1, w, by rintro ⟨⟩⟩, dif_pos rfl⟩
254+
| w@⟨⟨j, m⟩ :: ls, h1, h2⟩ := if ij : i = j then
255+
{ val := { head := ij.symm.rec m,
256+
tail := mk_aux ls h1 h2,
257+
fst_idx_ne := by cases ij; exact fst_idx_ne_iff.mpr h2.rel_head' },
258+
property := by cases ij; exact cons_eq_rcons.symm }
259+
else ⟨⟨1, w, (option.some_injective _).ne (ne.symm ij)⟩, dif_pos rfl⟩
260+
261+
/-- The equivalence between words and pairs. Given a word, it decomposes it as a pair by removing
262+
the first letter if it comes from `M i`. Given a pair, it prepends the head to the tail. -/
263+
def equiv_pair (i) : word M ≃ pair M i :=
264+
{ to_fun := λ w, (equiv_pair_aux i w).val,
265+
inv_fun := rcons,
266+
left_inv := λ w, (equiv_pair_aux i w).property,
267+
right_inv := λ p, rcons_inj (equiv_pair_aux i _).property }
268+
269+
lemma equiv_pair_symm (i) (p : pair M i) : (equiv_pair i).symm p = rcons p := rfl
270+
271+
lemma equiv_pair_eq_of_fst_idx_ne {i} {w : word M} (h : fst_idx w ≠ some i) :
272+
equiv_pair i w = ⟨1, w, h⟩ :=
273+
(equiv_pair i).apply_eq_iff_eq_symm_apply.mpr $ eq.symm (dif_pos rfl)
274+
275+
instance summand_action (i) : mul_action (M i) (word M) :=
276+
{ smul := λ m w, rcons { head := m * (equiv_pair i w).head, ..equiv_pair i w },
277+
one_smul := λ w, by { simp_rw [one_mul], apply (equiv_pair i).symm_apply_eq.mpr, ext; refl },
278+
mul_smul := λ m m' w, by simp only [mul_assoc, ←equiv_pair_symm, equiv.apply_symm_apply], }
279+
280+
instance : mul_action (free_product M) (word M) :=
281+
mul_action.of_End_hom (lift (λ i, mul_action.to_End_hom))
282+
283+
lemma of_smul_def (i) (w : word M) (m : M i) :
284+
of m • w = rcons { head := m * (equiv_pair i w).head, ..equiv_pair i w } := rfl
285+
286+
lemma cons_eq_smul {i} {m : M i} {ls h1 h2} :
287+
word.mk (⟨i, m⟩ :: ls) h1 h2 = of m • mk_aux ls h1 h2 :=
288+
by rw [cons_eq_rcons, of_smul_def, equiv_pair_eq_of_fst_idx_ne _]; simp only [mul_one]
289+
290+
lemma smul_induction {C : word M → Prop}
291+
(h_empty : C empty)
292+
(h_smul : ∀ i (m : M i) w, C w → C (of m • w))
293+
(w : word M) : C w :=
294+
begin
295+
cases w with ls h1 h2,
296+
induction ls with l ls ih,
297+
{ exact h_empty },
298+
cases l with i m,
299+
rw cons_eq_smul,
300+
exact h_smul _ _ _ (ih _ _),
301+
end
302+
303+
@[simp] lemma prod_smul (m) : ∀ w : word M, prod (m • w) = m * prod w :=
304+
begin
305+
apply m.induction_on,
306+
{ intro, rw [one_smul, one_mul] },
307+
{ intros, rw [of_smul_def, prod_rcons, of.map_mul, mul_assoc, ←prod_rcons,
308+
←equiv_pair_symm, equiv.symm_apply_apply] },
309+
{ intros x y hx hy w, rw [mul_smul, hx, hy, mul_assoc] },
310+
end
311+
312+
/-- Each element of the free product corresponds to a unique reduced word. -/
313+
def equiv : free_product M ≃ word M :=
314+
{ to_fun := λ m, m • empty,
315+
inv_fun := λ w, prod w,
316+
left_inv := λ m, by dsimp only; rw [prod_smul, prod_nil, mul_one],
317+
right_inv := begin
318+
apply smul_induction,
319+
{ dsimp only, rw [prod_nil, one_smul], },
320+
{ dsimp only, intros i m w ih, rw [prod_smul, mul_smul, ih], },
321+
end }
322+
323+
instance : decidable_eq (word M) := function.injective.decidable_eq word.ext
324+
instance : decidable_eq (free_product M) := word.equiv.decidable_eq
325+
326+
end word
327+
150328
end free_product

0 commit comments

Comments
 (0)