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

Commit ea1cff4

Browse files
committed
feat(linear_algebra/pi): ext lemma for f : (Π i, M i) →ₗ[R] N (#6233)
1 parent 0c61fc4 commit ea1cff4

File tree

3 files changed

+41
-8
lines changed

3 files changed

+41
-8
lines changed

src/algebra/big_operators/pi.lean

Lines changed: 12 additions & 7 deletions
Original file line numberDiff line numberDiff line change
@@ -61,7 +61,6 @@ begin
6161
{ intro h, exfalso, simpa using h, },
6262
end
6363

64-
@[ext]
6564
lemma add_monoid_hom.functions_ext [fintype I] (G : Type*)
6665
[add_comm_monoid G] (g h : (Π i, Z i) →+ G)
6766
(w : ∀ (i : I) (x : Z i), g (pi.single i x) = h (pi.single i x)) : g = h :=
@@ -73,22 +72,28 @@ begin
7372
apply w,
7473
end
7574

75+
/-- This is used as the ext lemma instead of `add_monoid_hom.functions_ext` for reasons explained in
76+
note [partially-applied ext lemmas]. -/
77+
@[ext]
78+
lemma add_monoid_hom.functions_ext' [fintype I] (M : Type*) [add_comm_monoid M]
79+
(g h : (Π i, Z i) →+ M)
80+
(H : ∀ i, g.comp (add_monoid_hom.single Z i) = h.comp (add_monoid_hom.single Z i)) :
81+
g = h :=
82+
have _ := λ i, add_monoid_hom.congr_fun (H i), -- elab without an expected type
83+
g.functions_ext M h this
84+
7685
end single
7786

7887
section ring_hom
7988
open pi
8089
variables {I : Type*} [decidable_eq I] {f : I → Type*}
8190
variables [Π i, semiring (f i)]
8291

83-
-- we need `apply`+`convert` because Lean fails to unify different `add_monoid` instances
84-
-- on `Π i, f i`
8592
@[ext]
8693
lemma ring_hom.functions_ext [fintype I] (G : Type*) [semiring G] (g h : (Π i, f i) →+* G)
8794
(w : ∀ (i : I) (x : f i), g (single i x) = h (single i x)) : g = h :=
88-
begin
89-
apply ring_hom.coe_add_monoid_hom_injective,
90-
convert add_monoid_hom.functions_ext _ _ _ _; assumption
91-
end
95+
ring_hom.coe_add_monoid_hom_injective $
96+
add_monoid_hom.functions_ext G (g : (Π i, f i) →+ G) h w
9297

9398
end ring_hom
9499

src/linear_algebra/matrix.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -140,7 +140,7 @@ end
140140

141141
@[simp] lemma matrix.to_lin'_one :
142142
matrix.to_lin' (1 : matrix n n R) = id :=
143-
by { ext, simp }
143+
by { ext, simp [one_apply, std_basis_apply] }
144144

145145
@[simp] lemma linear_map.to_matrix'_id :
146146
(linear_map.to_matrix' (linear_map.id : (n → R) →ₗ[R] (n → R))) = 1 :=

src/linear_algebra/std_basis.lean

Lines changed: 28 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -44,12 +44,40 @@ def std_basis (i : ι) : φ i →ₗ[R] (Πi, φ i) := pi (diag i)
4444
lemma std_basis_apply (i : ι) (b : φ i) : std_basis R φ i b = update 0 i b :=
4545
by ext j; rw [std_basis, pi_apply, diag, update_apply]; refl
4646

47+
lemma coe_std_basis (i : ι) : ⇑(std_basis R φ i) = pi.single i :=
48+
funext $ std_basis_apply R φ i
49+
4750
@[simp] lemma std_basis_same (i : ι) (b : φ i) : std_basis R φ i b i = b :=
4851
by rw [std_basis_apply, update_same]
4952

5053
lemma std_basis_ne (i j : ι) (h : j ≠ i) (b : φ i) : std_basis R φ i b j = 0 :=
5154
by rw [std_basis_apply, update_noteq h]; refl
5255

56+
section ext
57+
58+
variables {R φ} {M : Type*} [fintype ι] [add_comm_monoid M] [semimodule R M]
59+
{f g : (Π i, φ i) →ₗ[R] M}
60+
61+
lemma pi_ext (h : ∀ i x, f (pi.single i x) = g (pi.single i x)) :
62+
f = g :=
63+
to_add_monoid_hom_injective $ add_monoid_hom.functions_ext _ _ _ h
64+
65+
lemma pi_ext_iff : f = g ↔ ∀ i x, f (pi.single i x) = g (pi.single i x) :=
66+
⟨λ h i x, h ▸ rfl, pi_ext⟩
67+
68+
/-- This is used as the ext lemma instead of `linear_map.pi_ext` for reasons explained in
69+
note [partially-applied ext lemmas]. -/
70+
@[ext] lemma pi_ext' (h : ∀ i, f.comp (std_basis R φ i) = g.comp (std_basis R φ i)) : f = g :=
71+
begin
72+
refine pi_ext (λ i x, _),
73+
simpa only [comp_apply, coe_std_basis] using linear_map.congr_fun (h i) x
74+
end
75+
76+
lemma pi_ext'_iff : f = g ↔ ∀ i, f.comp (std_basis R φ i) = g.comp (std_basis R φ i) :=
77+
⟨λ h i, h ▸ rfl, pi_ext'⟩
78+
79+
end ext
80+
5381
lemma ker_std_basis (i : ι) : ker (std_basis R φ i) = ⊥ :=
5482
ker_eq_bot_of_injective $ assume f g hfg,
5583
have std_basis R φ i f i = std_basis R φ i g i := hfg ▸ rfl,

0 commit comments

Comments
 (0)