GroupTheory/src/Group/SubgroupTests.agda
2024-10-22 15:03:45 -07:00

107 lines
4.5 KiB
Agda
Raw Blame History

This file contains ambiguous Unicode characters

This file contains Unicode characters that might be confused with other characters. If you think that this is intentional, you can safely ignore this warning. Use the Escape button to reveal them.

open import Group.Core
open import Level using (Level)
open import Data.Product using (_,_; Σ-syntax; proj₁; proj₂)
open import Function using (_$_; _∘_; id)
import Group.Properties
module Group.SubgroupTests
{ℓ₁ ℓ₂ ℓ₃ : Level}
(𝐆 : Group {ℓ₁} {ℓ₂})
(H : Set ℓ₃)
where
open Group 𝐆
two-step : (ι : H G)
( (a b : H) Σ[ c H ] ι c ι a · ι b)
( (a : H) Σ[ a⁻¹ H ] ι a⁻¹ (ι a) ⁻¹)
Σ[ e H ] (ι e e)
Group
two-step ι ·-closed ⁻¹-closed (e , ιe≈e) =
record
{ G = H
; _·_ = λ g proj₁ ·-closed g
; _⁻¹ = proj₁ ⁻¹-closed
; _≈_ = λ g h ι g ι h
; e = e
; isEquivalence = record
{ refl = refl
; sym = sym
; trans = trans
}
; ≈-resp-·ˡ = λ {a} {b} {c} ιa≈ιb
let (a·c , ιa·c≈ιιc) = ·-closed a c in
let (b·c , ιb·c≈ιιc) = ·-closed b c in begin
ι a·c ≈⟨ ιa·c≈ιιc
ι a · ι c ≈⟨ ≈-resp-·ˡ ιa≈ιb
ι b · ι c ≈⟨ sym ιb·c≈ιιc
ι b·c
; ≈-resp-·ʳ = λ {a} {b} {c} ιa≈ιb
let (c·a , ιc·a≈ιιa) = ·-closed c a in
let (c·b , ιc·b≈ιιb) = ·-closed c b in begin
ι c·a ≈⟨ ιc·a≈ιιa
ι c · ι a ≈⟨ ≈-resp-·ʳ ιa≈ιb
ι c · ι b ≈⟨ sym ιc·b≈ιιb
ι c·b
; ≈-resp-⁻¹ = λ {a} {b} ιa≈ιb
let (a⁻¹ , hai) = ⁻¹-closed a in
let (b⁻¹ , hbi) = ⁻¹-closed b in begin
ι a⁻¹ ≈⟨ hai
(ι a) ⁻¹ ≈⟨ ≈-resp-⁻¹ ιa≈ιb
(ι b) ⁻¹ ≈⟨ sym hbi
ι b⁻¹
; identity-ˡ = λ {a}
let (e·a , ιe·a≈ιe·ιa) = ·-closed e a in begin
ι e·a ≈⟨ ιe·a≈ιe·ιa
ι e · ι a ≈⟨ ≈-resp-·ˡ ιe≈e
e · ι a ≈⟨ identity-ˡ
ι a
; identity-ʳ = λ {a}
let (a·e , ιa·eιιe) = ·-closed a e in begin
ι a·e ≈⟨ ιa·eιιe
ι a · ι e ≈⟨ ≈-resp-·ʳ ιe≈e
ι a · e ≈⟨ identity-ʳ
ι a
; assoc = λ {a} {b} {c}
let (a·b , hab) = ·-closed a b in
let (ab·c , hc) = ·-closed a·b c in
let (b·c , hbc) = ·-closed b c in
let (a·bc , ha) = ·-closed a b·c in begin
ι a·bc ≈⟨ ha
ι a · ι b·c ≈⟨ ≈-resp-·ʳ hbc
ι a · (ι b · ι c) ≈⟨ assoc
(ι a · ι b) · ι c ≈⟨ sym (≈-resp-·ˡ hab)
(ι a·b) · ι c ≈⟨ sym hc
ι ab·c
; inverse-ˡ = λ {a}
let (a⁻¹ , ha⁻¹) = ⁻¹-closed a in
let (a⁻¹·a , hboth) = ·-closed a⁻¹ a in begin
ι a⁻¹·a ≈⟨ hboth
ι a⁻¹ · ι a ≈⟨ ≈-resp-·ˡ ha⁻¹
(ι a) ⁻¹ · ι a ≈⟨ inverse-ˡ
e ≈⟨ sym ιe≈e
ι e
; inverse-ʳ = λ {a}
let (a⁻¹ , ha⁻¹) = ⁻¹-closed a in
let (a·a⁻¹ , hboth) = ·-closed a a⁻¹ in begin
ι a·a⁻¹ ≈⟨ hboth
ι a · ι a⁻¹ ≈⟨ ≈-resp-·ʳ ha⁻¹
ι a · (ι a) ⁻¹ ≈⟨ inverse-ʳ
e ≈⟨ sym ιe≈e
ι e
}
where open Reasoning
two-step-subgroup : (ι : H G)
(h₁ : ( (a b : H) Σ[ c H ] ι c ι a · ι b))
(h₂ : ( (a : H) Σ[ a⁻¹ H ] ι a⁻¹ (ι a) ⁻¹))
(h₃ : Σ[ e H ] (ι e e))
Subgroup (two-step ι h₁ h₂ h₃) 𝐆
two-step-subgroup ι h₁ h₂ h₃ = record { F = F ; inj = id }
where
𝐇 = two-step ι h₁ h₂ h₃
F = record
{ φ = ι
; homo = λ {a} {b} h₁ a b .proj₂
; φ-resp-≈ = id }