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

Commit 6d85ff2

Browse files
committed
refactor(linear_algebra/finsupp): replace mem_span_iff_total (#7735)
This PR renames the existing `mem_span_iff_total` to `mem_span_image_iff_total` and the existing `span_eq_map_total` to `span_image_eq_map_total`, and replaces them with slightly simpler lemmas about sets in the module, rather than indexed families. As usual in the linear algebra library, there is tension between using sets and using indexed families, but as `span` is defined in terms of sets I think the new lemmas merit taking the simpler names. Co-authored-by: Scott Morrison <scott.morrison@gmail.com>
1 parent 6d90d35 commit 6d85ff2

File tree

13 files changed

+42
-32
lines changed

13 files changed

+42
-32
lines changed

src/algebraic_geometry/structure_sheaf.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -716,7 +716,7 @@ begin
716716
exact subset_vanishing_ideal_zero_locus {f} (set.mem_singleton f) },
717717

718718
replace hn := ideal.mul_mem_left _ f hn,
719-
erw [←pow_succ, finsupp.mem_span_iff_total] at hn,
719+
erw [←pow_succ, finsupp.mem_span_image_iff_total] at hn,
720720
rcases hn with ⟨b, b_supp, hb⟩,
721721
rw finsupp.total_apply_of_mem_supported R b_supp at hb,
722722
dsimp at hb,

src/analysis/normed_space/inner_product.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -2739,7 +2739,7 @@ begin
27392739
have hxv' : (⟨x, hxu⟩ : u) ∉ (coe ⁻¹' v : set u) := by simp [huv, hxv],
27402740
obtain ⟨l, hl, rfl⟩ :
27412741
∃ l ∈ finsupp.supported 𝕜 𝕜 (coe ⁻¹' v : set u), (finsupp.total ↥u E 𝕜 coe) l = y,
2742-
{ rw ← finsupp.mem_span_iff_total,
2742+
{ rw ← finsupp.mem_span_image_iff_total,
27432743
simp [huv, inter_eq_self_of_subset_left, hy] },
27442744
exact hu.inner_finsupp_eq_zero hxv' hl }
27452745
end

src/field_theory/algebraic_closure.lean

Lines changed: 2 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -156,7 +156,8 @@ by { rw [to_splitting_field, eval_X_self, ← alg_hom.coe_to_ring_hom, hom_eval
156156

157157
theorem span_eval_ne_top : span_eval k ≠ ⊤ :=
158158
begin
159-
rw [ideal.ne_top_iff_one, span_eval, ideal.span, ← set.image_univ, finsupp.mem_span_iff_total],
159+
rw [ideal.ne_top_iff_one, span_eval, ideal.span, ← set.image_univ,
160+
finsupp.mem_span_image_iff_total],
160161
rintros ⟨v, _, hv⟩,
161162
replace hv := congr_arg (to_splitting_field k v.support) hv,
162163
rw [alg_hom.map_one, finsupp.total_apply, finsupp.sum, alg_hom.map_sum, finset.sum_eq_zero] at hv,

src/field_theory/fixed.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -78,7 +78,7 @@ begin
7878
rw coe_insert at hs ⊢,
7979
rw linear_independent_insert (mt mem_coe.1 has) at hs,
8080
rw linear_independent_insert' (mt mem_coe.1 has), refine ⟨ih hs.1, λ ha, _⟩,
81-
rw finsupp.mem_span_iff_total at ha, rcases ha with ⟨l, hl, hla⟩,
81+
rw finsupp.mem_span_image_iff_total at ha, rcases ha with ⟨l, hl, hla⟩,
8282
rw [finsupp.total_apply_of_mem_supported F hl] at hla,
8383
suffices : ∀ i ∈ s, l i ∈ fixed_points G F,
8484
{ replace hla := (sum_apply _ _ (λ i, l i • to_fun G F i)).symm.trans (congr_fun hla 1),

src/linear_algebra/affine_space/combination.lean

Lines changed: 2 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -537,7 +537,7 @@ begin
537537
by_cases hn : nonempty ι,
538538
{ cases hn with i0,
539539
rw [vector_span_range_eq_span_range_vsub_right k p i0, ←set.image_univ,
540-
finsupp.mem_span_iff_total,
540+
finsupp.mem_span_image_iff_total,
541541
finset.weighted_vsub_eq_weighted_vsub_of_point_of_sum_eq_zero s w p h (p i0),
542542
finset.weighted_vsub_of_point_apply],
543543
let w' := set.indicator ↑s w,
@@ -587,7 +587,7 @@ begin
587587
{ by_cases hn : nonempty ι,
588588
{ cases hn with i0,
589589
rw [vector_span_range_eq_span_range_vsub_right k p i0, ←set.image_univ,
590-
finsupp.mem_span_iff_total],
590+
finsupp.mem_span_image_iff_total],
591591
rintros ⟨l, hl, hv⟩,
592592
use insert i0 l.support,
593593
set w := (l : ι → k) -

src/linear_algebra/basis.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -395,7 +395,7 @@ lemma constr_range [nonempty ι] {f : ι → M'} :
395395
(b.constr S f).range = span R (range f) :=
396396
by rw [b.constr_def S f, linear_map.range_comp, linear_map.range_comp, linear_equiv.range,
397397
← finsupp.supported_univ, finsupp.lmap_domain_supported, ←set.image_univ,
398-
← finsupp.span_eq_map_total, set.image_id]
398+
← finsupp.span_image_eq_map_total, set.image_id]
399399

400400
end constr
401401

src/linear_algebra/dual.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -388,7 +388,7 @@ lemma mem_of_mem_span {H : set ι} {x : M} (hmem : x ∈ submodule.span R (e ''
388388
∀ i : ι, ε i x ≠ 0 → i ∈ H :=
389389
begin
390390
intros i hi,
391-
rcases (finsupp.mem_span_iff_total _).mp hmem with ⟨l, supp_l, rfl⟩,
391+
rcases (finsupp.mem_span_image_iff_total _).mp hmem with ⟨l, supp_l, rfl⟩,
392392
apply not_imp_comm.mp ((finsupp.mem_supported' _ _).mp supp_l i),
393393
rwa [← lc_def, h.dual_lc] at hi
394394
end

src/linear_algebra/finsupp.lean

Lines changed: 17 additions & 8 deletions
Original file line numberDiff line numberDiff line change
@@ -468,7 +468,16 @@ begin
468468
apply total_emb_domain R ⟨f, hf⟩ l
469469
end
470470

471-
theorem span_eq_map_total (s : set α):
471+
/-- A version of `finsupp.range_total` which is useful for going in the other direction -/
472+
theorem span_eq_range_total (s : set M) :
473+
span R s = (finsupp.total s M R coe).range :=
474+
by rw [range_total, subtype.range_coe_subtype, set.set_of_mem_eq]
475+
476+
theorem mem_span_iff_total (s : set M) (x : M) :
477+
x ∈ span R s ↔ ∃ l : s →₀ R, finsupp.total s M R coe l = x :=
478+
(set_like.ext_iff.1 $ span_eq_range_total _ _) x
479+
480+
theorem span_image_eq_map_total (s : set α):
472481
span R (v '' s) = submodule.map (finsupp.total α M R v) (supported R R s) :=
473482
begin
474483
apply span_eq_of_le,
@@ -487,9 +496,9 @@ begin
487496
refine sum_mem _ _, simp [this] }
488497
end
489498

490-
theorem mem_span_iff_total {s : set α} {x : M} :
499+
theorem mem_span_image_iff_total {s : set α} {x : M} :
491500
x ∈ span R (v '' s) ↔ ∃ l ∈ supported R R s, finsupp.total α M R v l = x :=
492-
by rw span_eq_map_total; simp
501+
by { rw span_image_eq_map_total, simp, }
493502

494503
variables (α) (M) (v)
495504

@@ -500,15 +509,15 @@ The subset is indicated by a set `s : set α` of indices.
500509
-/
501510
protected def total_on (s : set α) : supported R R s →ₗ[R] span R (v '' s) :=
502511
linear_map.cod_restrict _ ((finsupp.total _ _ _ v).comp (submodule.subtype (supported R R s))) $
503-
λ ⟨l, hl⟩, (mem_span_iff_total _).2 ⟨l, hl, rfl⟩
512+
λ ⟨l, hl⟩, (mem_span_image_iff_total _).2 ⟨l, hl, rfl⟩
504513

505514
variables {α} {M} {v}
506515

507516
theorem total_on_range (s : set α) : (finsupp.total_on α M R v s).range = ⊤ :=
508517
begin
509518
rw [finsupp.total_on, linear_map.range_eq_map, linear_map.map_cod_restrict,
510519
← linear_map.range_le_iff_comap, range_subtype, map_top, linear_map.range_comp, range_subtype],
511-
exact (span_eq_map_total _ _).le
520+
exact (span_image_eq_map_total _ _).le
512521
end
513522

514523
theorem total_comp (f : α' → α) :
@@ -763,7 +772,7 @@ lemma submodule.exists_finset_of_mem_supr
763772
begin
764773
obtain ⟨f, hf, rfl⟩ : ∃ f ∈ finsupp.supported R R (⋃ i, ↑(p i)), finsupp.total M M R id f = m,
765774
{ have aux : (id : M → M) '' (⋃ (i : ι), ↑(p i)) = (⋃ (i : ι), ↑(p i)) := set.image_id _,
766-
rwa [supr_eq_span, ← aux, finsupp.mem_span_iff_total R] at hm },
775+
rwa [supr_eq_span, ← aux, finsupp.mem_span_image_iff_total R] at hm },
767776
let t : finset M := f.support,
768777
have ht : ∀ x : {x // x ∈ t}, ∃ i, ↑x ∈ p i,
769778
{ intros x,
@@ -786,7 +795,7 @@ end
786795

787796
lemma mem_span_finset {s : finset M} {x : M} :
788797
x ∈ span R (↑s : set M) ↔ ∃ f : M → R, ∑ i in s, f i • i = x :=
789-
⟨λ hx, let ⟨v, hvs, hvx⟩ := (finsupp.mem_span_iff_total _).1
798+
⟨λ hx, let ⟨v, hvs, hvx⟩ := (finsupp.mem_span_image_iff_total _).1
790799
(show x ∈ span R (id '' (↑s : set M)), by rwa set.image_id) in
791800
⟨v, hvx ▸ (finsupp.total_apply_of_mem_supported _ hvs).symm⟩,
792801
λ ⟨f, hf⟩, hf ▸ sum_mem _ (λ i hi, smul_mem _ _ $ subset_span hi)⟩
@@ -799,7 +808,7 @@ lemma mem_span_set {m : M} {s : set M} :
799808
begin
800809
conv_lhs { rw ←set.image_id s },
801810
simp_rw ←exists_prop,
802-
exact finsupp.mem_span_iff_total R,
811+
exact finsupp.mem_span_image_iff_total R,
803812
end
804813

805814
/-- If `subsingleton R`, then `M ≃ₗ[R] ι →₀ R` for any type `ι`. -/

src/linear_algebra/linear_independent.lean

Lines changed: 4 additions & 4 deletions
Original file line numberDiff line numberDiff line change
@@ -186,7 +186,7 @@ family of vectors. See also `linear_independent.map'` for a special case assumin
186186
lemma linear_independent.map (hv : linear_independent R v) {f : M →ₗ[R] M'}
187187
(hf_inj : disjoint (span R (range v)) f.ker) : linear_independent R (f ∘ v) :=
188188
begin
189-
rw [disjoint, ← set.image_univ, finsupp.span_eq_map_total, map_inf_eq_map_inf_comap,
189+
rw [disjoint, ← set.image_univ, finsupp.span_image_eq_map_total, map_inf_eq_map_inf_comap,
190190
map_le_iff_le_comap, comap_bot, finsupp.supported_univ, top_inf_eq] at hf_inj,
191191
unfold linear_independent at hv ⊢,
192192
rw [hv, le_bot_iff] at hf_inj,
@@ -453,7 +453,7 @@ lemma linear_independent.disjoint_span_image (hv : linear_independent R v) {s t
453453
(hs : disjoint s t) :
454454
disjoint (submodule.span R $ v '' s) (submodule.span R $ v '' t) :=
455455
begin
456-
simp only [disjoint_def, finsupp.mem_span_iff_total],
456+
simp only [disjoint_def, finsupp.mem_span_image_iff_total],
457457
rintros _ ⟨l₁, hl₁, rfl⟩ ⟨l₂, hl₂, H⟩,
458458
rw [hv.injective_total.eq_iff] at H, subst l₂,
459459
have : l₁ = 0 := finsupp.disjoint_supported_supported hs (submodule.mem_inf.2 ⟨hl₁, hl₂⟩),
@@ -621,7 +621,7 @@ end
621621
lemma linear_independent_iff_not_smul_mem_span :
622622
linear_independent R v ↔ (∀ (i : ι) (a : R), a • (v i) ∈ span R (v '' (univ \ {i})) → a = 0) :=
623623
⟨ λ hv i a ha, begin
624-
rw [finsupp.span_eq_map_total, mem_map] at ha,
624+
rw [finsupp.span_image_eq_map_total, mem_map] at ha,
625625
rcases ha with ⟨l, hl, e⟩,
626626
rw sub_eq_zero.1 (linear_independent_iff.1 hv (l - finsupp.single i a) (by simp [e])) at hl,
627627
by_contra hn,
@@ -630,7 +630,7 @@ end, λ H, linear_independent_iff.2 $ λ l hl, begin
630630
ext i, simp only [finsupp.zero_apply],
631631
by_contra hn,
632632
refine hn (H i _ _),
633-
refine (finsupp.mem_span_iff_total _).2 ⟨finsupp.single i (l i) - l, _, _⟩,
633+
refine (finsupp.mem_span_image_iff_total _).2 ⟨finsupp.single i (l i) - l, _, _⟩,
634634
{ rw finsupp.mem_supported',
635635
intros j hj,
636636
have hij : j = i :=

src/ring_theory/adjoin/basic.lean

Lines changed: 2 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -220,14 +220,14 @@ begin
220220
change r ∈ (adjoin (adjoin R s) t).to_submodule at hr,
221221
haveI := classical.dec_eq A,
222222
haveI := classical.dec_eq R,
223-
rw [← hq', ← set.image_id q, finsupp.mem_span_iff_total (adjoin R s)] at hr,
223+
rw [← hq', ← set.image_id q, finsupp.mem_span_image_iff_total (adjoin R s)] at hr,
224224
rcases hr with ⟨l, hlq, rfl⟩,
225225
have := @finsupp.total_apply A A (adjoin R s),
226226
rw [this, finsupp.sum],
227227
refine sum_mem _ _,
228228
intros z hz, change (l z).1 * _ ∈ _,
229229
have : (l z).1 ∈ (adjoin R s).to_submodule := (l z).2,
230-
rw [← hp', ← set.image_id p, finsupp.mem_span_iff_total R] at this,
230+
rw [← hp', ← set.image_id p, finsupp.mem_span_image_iff_total R] at this,
231231
rcases this with ⟨l2, hlp, hl⟩,
232232
have := @finsupp.total_apply A A R,
233233
rw this at hl,

0 commit comments

Comments
 (0)