@@ -37,14 +37,16 @@ function with finite support, module, linear algebra
37
37
38
38
variable {ι : Type *} {R : Type *} {S : Type *} {M : ι → Type *} {N : Type *}
39
39
40
- variable [dec_ι : DecidableEq ι]
41
-
42
40
namespace DFinsupp
43
41
44
42
variable [Semiring R] [∀ i, AddCommMonoid (M i)] [∀ i, Module R (M i)]
45
43
46
44
variable [AddCommMonoid N] [Module R N]
47
45
46
+ section DecidableEq
47
+
48
+ variable [DecidableEq ι]
49
+
48
50
/-- `DFinsupp.mk` as a `LinearMap`. -/
49
51
def lmk (s : Finset ι) : (∀ i : (↑s : Set ι), M i) →ₗ[R] Π₀ i, M i where
50
52
toFun := mk s
@@ -137,8 +139,7 @@ def lsum [Semiring S] [Module S N] [SMulCommClass R S N] :
137
139
(∀ i, M i →ₗ[R] N) ≃ₗ[S] (Π₀ i, M i) →ₗ[R] N where
138
140
toFun F :=
139
141
{ toFun := sumAddHom fun i => (F i).toAddMonoidHom
140
- -- Porting note: needed to (β := M) hint below
141
- map_add' := (DFinsupp.liftAddHom (β := M) (fun (i : ι) => (F i).toAddMonoidHom)).map_add
142
+ map_add' := (DFinsupp.liftAddHom fun (i : ι) => (F i).toAddMonoidHom).map_add
142
143
map_smul' := fun c f => by
143
144
dsimp
144
145
apply DFinsupp.induction f
@@ -174,12 +175,13 @@ theorem lsum_single [Semiring S] [Module S N] [SMulCommClass R S N] (F : ∀ i,
174
175
175
176
end Lsum
176
177
178
+ end DecidableEq
179
+
177
180
/-! ### Bundled versions of `DFinsupp.mapRange`
178
181
179
182
The names should match the equivalent bundled `Finsupp.mapRange` definitions.
180
183
-/
181
184
182
-
183
185
section mapRange
184
186
185
187
variable {β β₁ β₂ : ι → Type *}
@@ -217,13 +219,10 @@ theorem mapRange.linearMap_comp (f : ∀ i, β₁ i →ₗ[R] β₂ i) (f₂ :
217
219
(fun i => (f i).map_zero) (fun i => (f₂ i).map_zero) (by simp)
218
220
#align dfinsupp.map_range.linear_map_comp DFinsupp.mapRange.linearMap_comp
219
221
220
- theorem sum_mapRange_index.linearMap [∀ (i : ι) (x : β₁ i), Decidable (x ≠ 0 )]
221
- [∀ (i : ι) (x : β₂ i), Decidable (x ≠ 0 )] {f : ∀ i, β₁ i →ₗ[R] β₂ i} {h : ∀ i, β₂ i →ₗ[R] N}
222
- {l : Π₀ i, β₁ i} :
223
- -- Porting note: Needed to add (M := ...) below
224
- (DFinsupp.lsum ℕ (M := β₂)) h (mapRange.linearMap f l)
225
- = (DFinsupp.lsum ℕ (M := β₁)) (fun i => (h i).comp (f i)) l := by
226
- simpa [DFinsupp.sumAddHom_apply] using sum_mapRange_index fun i => by simp
222
+ theorem sum_mapRange_index.linearMap [DecidableEq ι] {f : ∀ i, β₁ i →ₗ[R] β₂ i}
223
+ {h : ∀ i, β₂ i →ₗ[R] N} {l : Π₀ i, β₁ i} :
224
+ DFinsupp.lsum ℕ h (mapRange.linearMap f l) = DFinsupp.lsum ℕ (fun i => (h i).comp (f i)) l := by
225
+ classical simpa [DFinsupp.sumAddHom_apply] using sum_mapRange_index fun i => by simp
227
226
#align dfinsupp.sum_map_range_index.linear_map DFinsupp.sum_mapRange_index.linearMap
228
227
229
228
/-- `DFinsupp.mapRange.linearMap` as a `LinearEquiv`. -/
@@ -289,6 +288,10 @@ variable [Semiring R] [AddCommMonoid N] [Module R N]
289
288
290
289
open DFinsupp
291
290
291
+ section DecidableEq
292
+
293
+ variable [DecidableEq ι]
294
+
292
295
theorem dfinsupp_sum_mem {β : ι → Type *} [∀ i, Zero (β i)] [∀ (i) (x : β i), Decidable (x ≠ 0 )]
293
296
(S : Submodule R N) (f : Π₀ i, β i) (g : ∀ i, β i → N) (h : ∀ c, f c ≠ 0 → g c (f c) ∈ S) :
294
297
f.sum g ∈ S :=
@@ -356,6 +359,16 @@ theorem mem_iSup_iff_exists_dfinsupp' (p : ι → Submodule R N) [∀ (i) (x : p
356
359
LinearMap.toAddMonoidHom_coe, coeSubtype]
357
360
#align submodule.mem_supr_iff_exists_dfinsupp' Submodule.mem_iSup_iff_exists_dfinsupp'
358
361
362
+ theorem mem_biSup_iff_exists_dfinsupp (p : ι → Prop ) [DecidablePred p] (S : ι → Submodule R N)
363
+ (x : N) :
364
+ (x ∈ ⨆ (i) (_ : p i), S i) ↔
365
+ ∃ f : Π₀ i, S i,
366
+ DFinsupp.lsum ℕ (M := fun i ↦ ↥(S i)) (fun i => (S i).subtype) (f.filter p) = x :=
367
+ SetLike.ext_iff.mp (biSup_eq_range_dfinsupp_lsum p S) x
368
+ #align submodule.mem_bsupr_iff_exists_dfinsupp Submodule.mem_biSup_iff_exists_dfinsupp
369
+
370
+ end DecidableEq
371
+
359
372
lemma mem_iSup_iff_exists_finsupp (p : ι → Submodule R N) (x : N) :
360
373
x ∈ iSup p ↔ ∃ (f : ι →₀ N), (∀ i, f i ∈ p i) ∧ (f.sum fun _i xi ↦ xi) = x := by
361
374
classical
@@ -366,14 +379,6 @@ lemma mem_iSup_iff_exists_finsupp (p : ι → Submodule R N) (x : N) :
366
379
· ext; simp
367
380
· simp [Finsupp.mem_support_iff.mp hi]
368
381
369
- theorem mem_biSup_iff_exists_dfinsupp (p : ι → Prop ) [DecidablePred p] (S : ι → Submodule R N)
370
- (x : N) :
371
- (x ∈ ⨆ (i) (_ : p i), S i) ↔
372
- ∃ f : Π₀ i, S i,
373
- DFinsupp.lsum ℕ (M := fun i ↦ ↥(S i)) (fun i => (S i).subtype) (f.filter p) = x :=
374
- SetLike.ext_iff.mp (biSup_eq_range_dfinsupp_lsum p S) x
375
- #align submodule.mem_bsupr_iff_exists_dfinsupp Submodule.mem_biSup_iff_exists_dfinsupp
376
-
377
382
open BigOperators
378
383
379
384
theorem mem_iSup_finset_iff_exists_sum {s : Finset ι} (p : ι → Submodule R N) (a : N) :
@@ -421,7 +426,7 @@ open DFinsupp
421
426
422
427
section Semiring
423
428
424
- variable [Semiring R] [AddCommMonoid N] [Module R N]
429
+ variable [DecidableEq ι] [ Semiring R] [AddCommMonoid N] [Module R N]
425
430
426
431
/-- Independence of a family of submodules can be expressed as a quantifier over `DFinsupp`s.
427
432
@@ -477,7 +482,7 @@ end Semiring
477
482
478
483
section Ring
479
484
480
- variable [Ring R] [AddCommGroup N] [Module R N]
485
+ variable [DecidableEq ι] [ Ring R] [AddCommGroup N] [Module R N]
481
486
482
487
/- If `DFinsupp.sumAddHom` applied with `AddSubmonoid.subtype` is injective then the additive
483
488
subgroups are independent. -/
0 commit comments