Skip to content

Commit

Permalink
feat(algebra/algebra/basic): An algebra isomorphism induces a group i…
Browse files Browse the repository at this point in the history
…somorphism between automorphism groups (#6622)

Constructs the group isomorphism induced from an algebra isomorphism.



Co-authored-by: tb65536 <tb65536@users.noreply.github.com>
  • Loading branch information
tb65536 and tb65536 committed Mar 14, 2021
1 parent 552ebeb commit ade8889
Showing 1 changed file with 19 additions and 0 deletions.
19 changes: 19 additions & 0 deletions src/algebra/algebra/basic.lean
Expand Up @@ -911,6 +911,25 @@ instance aut : group (A₁ ≃ₐ[R] A₁) :=
inv := symm,
mul_left_inv := λ ϕ, by { ext, exact symm_apply_apply ϕ a } }

@[simp] lemma mul_apply (e₁ e₂ : A₁ ≃ₐ[R] A₁) (x : A₁) : (e₁ * e₂) x = e₁ (e₂ x) := rfl

/-- An algebra isomorphism induces a group isomorphism between automorphism groups -/
@[simps apply]
def aut_congr (ϕ : A₁ ≃ₐ[R] A₂) : (A₁ ≃ₐ[R] A₁) ≃* (A₂ ≃ₐ[R] A₂) :=
{ to_fun := λ ψ, ϕ.symm.trans (ψ.trans ϕ),
inv_fun := λ ψ, ϕ.trans (ψ.trans ϕ.symm),
left_inv := λ ψ, by { ext, simp_rw [trans_apply, symm_apply_apply] },
right_inv := λ ψ, by { ext, simp_rw [trans_apply, apply_symm_apply] },
map_mul' := λ ψ χ, by { ext, simp only [mul_apply, trans_apply, symm_apply_apply] } }

@[simp] lemma aut_congr_refl : aut_congr (alg_equiv.refl) = mul_equiv.refl (A₁ ≃ₐ[R] A₁) :=
by { ext, refl }

@[simp] lemma aut_congr_symm (ϕ : A₁ ≃ₐ[R] A₂) : (aut_congr ϕ).symm = aut_congr ϕ.symm := rfl

@[simp] lemma aut_congr_trans (ϕ : A₁ ≃ₐ[R] A₂) (ψ : A₂ ≃ₐ[R] A₃) :
(aut_congr ϕ).trans (aut_congr ψ) = aut_congr (ϕ.trans ψ) := rfl

end semiring

section comm_semiring
Expand Down

0 comments on commit ade8889

Please sign in to comment.