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

Commit bf0f15a

Browse files
committed
chore(algebra/algebra/basic): add missing lemmas (#7412)
1 parent e3de4e3 commit bf0f15a

File tree

1 file changed

+22
-6
lines changed

1 file changed

+22
-6
lines changed

src/algebra/algebra/basic.lean

Lines changed: 22 additions & 6 deletions
Original file line numberDiff line numberDiff line change
@@ -529,17 +529,25 @@ protected def id : A →ₐ[R] A :=
529529
{ commutes' := λ _, rfl,
530530
..ring_hom.id A }
531531

532+
@[simp] lemma coe_id : ⇑(alg_hom.id R A) = id := rfl
533+
534+
@[simp] lemma id_to_ring_hom : (alg_hom.id R A : A →+* A) = ring_hom.id _ := rfl
535+
532536
end
533537

534-
@[simp] lemma id_apply (p : A) : alg_hom.id R A p = p := rfl
538+
lemma id_apply (p : A) : alg_hom.id R A p = p := rfl
535539

536540
/-- Composition of algebra homeomorphisms. -/
537541
def comp (φ₁ : B →ₐ[R] C) (φ₂ : A →ₐ[R] B) : A →ₐ[R] C :=
538542
{ commutes' := λ r : R, by rw [← φ₁.commutes, ← φ₂.commutes]; refl,
539543
.. φ₁.to_ring_hom.comp ↑φ₂ }
540544

541-
@[simp] lemma comp_apply (φ₁ : B →ₐ[R] C) (φ₂ : A →ₐ[R] B) (p : A) :
542-
φ₁.comp φ₂ p = φ₁ (φ₂ p) := rfl
545+
@[simp] lemma coe_comp (φ₁ : B →ₐ[R] C) (φ₂ : A →ₐ[R] B) : ⇑(φ₁.comp φ₂) = φ₁ ∘ φ₂ := rfl
546+
547+
lemma comp_apply (φ₁ : B →ₐ[R] C) (φ₂ : A →ₐ[R] B) (p : A) : φ₁.comp φ₂ p = φ₁ (φ₂ p) := rfl
548+
549+
lemma comp_to_ring_hom (φ₁ : B →ₐ[R] C) (φ₂ : A →ₐ[R] B) :
550+
⇑(φ₁.comp φ₂ : A →+* C) = (φ₁ : B →+* C).comp ↑φ₂ := rfl
543551

544552
@[simp] theorem comp_id : φ.comp (alg_hom.id R A) = φ :=
545553
ext $ λ x, rfl
@@ -750,8 +758,9 @@ instance : inhabited (A₁ ≃ₐ[R] A₁) := ⟨1⟩
750758
@[refl]
751759
def refl : A₁ ≃ₐ[R] A₁ := 1
752760

753-
@[simp] lemma coe_refl : (@refl R A₁ _ _ _ : A₁ →ₐ[R] A₁) = alg_hom.id R A₁ :=
754-
alg_hom.ext (λ x, rfl)
761+
@[simp] lemma refl_to_alg_hom : ↑(refl : A₁ ≃ₐ[R] A₁) = alg_hom.id R A₁ := rfl
762+
763+
@[simp] lemma coe_refl : ⇑(refl : A₁ ≃ₐ[R] A₁) = id := rfl
755764

756765
/-- Algebra equivalences are symmetric. -/
757766
@[symm]
@@ -794,7 +803,10 @@ def trans (e₁ : A₁ ≃ₐ[R] A₂) (e₂ : A₂ ≃ₐ[R] A₃) : A₁ ≃
794803
@[simp] lemma symm_apply_apply (e : A₁ ≃ₐ[R] A₂) : ∀ x, e.symm (e x) = x :=
795804
e.to_equiv.symm_apply_apply
796805

797-
@[simp] lemma trans_apply (e₁ : A₁ ≃ₐ[R] A₂) (e₂ : A₂ ≃ₐ[R] A₃) (x : A₁) :
806+
@[simp] lemma coe_trans (e₁ : A₁ ≃ₐ[R] A₂) (e₂ : A₂ ≃ₐ[R] A₃) :
807+
⇑(e₁.trans e₂) = e₂ ∘ e₁ := rfl
808+
809+
lemma trans_apply (e₁ : A₁ ≃ₐ[R] A₂) (e₂ : A₂ ≃ₐ[R] A₃) (x : A₁) :
798810
(e₁.trans e₂) x = e₂ (e₁ x) := rfl
799811

800812
@[simp] lemma comp_symm (e : A₁ ≃ₐ[R] A₂) :
@@ -805,6 +817,10 @@ by { ext, simp }
805817
alg_hom.comp ↑e.symm (e : A₁ →ₐ[R] A₂) = alg_hom.id R A₁ :=
806818
by { ext, simp }
807819

820+
theorem left_inverse_symm (e : A₁ ≃ₐ[R] A₂) : function.left_inverse e.symm e := e.left_inv
821+
822+
theorem right_inverse_symm (e : A₁ ≃ₐ[R] A₂) : function.right_inverse e.symm e := e.right_inv
823+
808824
/-- If `A₁` is equivalent to `A₁'` and `A₂` is equivalent to `A₂'`, then the type of maps
809825
`A₁ →ₐ[R] A₂` is equivalent to the type of maps `A₁' →ₐ[R] A₂'`. -/
810826
def arrow_congr {A₁' A₂' : Type*} [semiring A₁'] [semiring A₂'] [algebra R A₁'] [algebra R A₂']

0 commit comments

Comments
 (0)