@@ -5,7 +5,7 @@ Author: Joseph Myers.
5
5
-/
6
6
import algebra.add_torsor
7
7
import data.indicator_function
8
- import linear_algebra.basis
8
+ import linear_algebra.finite_dimensional
9
9
10
10
noncomputable theory
11
11
open_locale big_operators
@@ -333,6 +333,67 @@ lemma affine_combination_indicator_subset (w : ι → k) (p : ι → P) {s₁ s
333
333
by rw [affine_combination_apply, affine_combination_apply,
334
334
weighted_vsub_of_point_indicator_subset _ _ _ _ h]
335
335
336
+ variables {V}
337
+
338
+ /-- Suppose an indexed family of points is given, along with a subset
339
+ of the index type. A vector can be expressed as
340
+ `weighted_vsub_of_point` using a `finset` lying within that subset and
341
+ with a given sum of weights if and only if it can be expressed as
342
+ `weighted_vsub_of_point` with that sum of weights for the
343
+ corresponding indexed family whose index type is the subtype
344
+ corresponding to that subset. -/
345
+ lemma eq_weighted_vsub_of_point_subset_iff_eq_weighted_vsub_of_point_subtype {v : V} {x : k}
346
+ {s : set ι} {p : ι → P} {b : P} :
347
+ (∃ (fs : finset ι) (hfs : ↑fs ⊆ s) (w : ι → k) (hw : ∑ i in fs, w i = x),
348
+ v = fs.weighted_vsub_of_point V p b w) ↔
349
+ ∃ (fs : finset s) (w : s → k) (hw : ∑ i in fs, w i = x),
350
+ v = fs.weighted_vsub_of_point V (λ (i : s), p i) b w :=
351
+ begin
352
+ simp_rw weighted_vsub_of_point_apply,
353
+ split,
354
+ { rintros ⟨fs, hfs, w, rfl, rfl⟩,
355
+ use [fs.subtype s, λ i, w i, sum_subtype_of_mem _ hfs, (sum_subtype_of_mem _ hfs).symm] },
356
+ { rintros ⟨fs, w, rfl, rfl⟩,
357
+ refine ⟨fs.map (function.embedding.subtype _), map_subtype_subset _,
358
+ λ i, if h : i ∈ s then w ⟨i, h⟩ else 0 , _, _⟩;
359
+ simp }
360
+ end
361
+
362
+ variables (k)
363
+
364
+ /-- Suppose an indexed family of points is given, along with a subset
365
+ of the index type. A vector can be expressed as `weighted_vsub` using
366
+ a `finset` lying within that subset and with sum of weights 0 if and
367
+ only if it can be expressed as `weighted_vsub` with sum of weights 0
368
+ for the corresponding indexed family whose index type is the subtype
369
+ corresponding to that subset. -/
370
+ lemma eq_weighted_vsub_subset_iff_eq_weighted_vsub_subtype {v : V} {s : set ι} {p : ι → P} :
371
+ (∃ (fs : finset ι) (hfs : ↑fs ⊆ s) (w : ι → k) (hw : ∑ i in fs, w i = 0 ),
372
+ v = fs.weighted_vsub V p w) ↔
373
+ ∃ (fs : finset s) (w : s → k) (hw : ∑ i in fs, w i = 0 ),
374
+ v = fs.weighted_vsub V (λ (i : s), p i) w :=
375
+ eq_weighted_vsub_of_point_subset_iff_eq_weighted_vsub_of_point_subtype
376
+
377
+ variables (V)
378
+
379
+ /-- Suppose an indexed family of points is given, along with a subset
380
+ of the index type. A point can be expressed as an
381
+ `affine_combination` using a `finset` lying within that subset and
382
+ with sum of weights 1 if and only if it can be expressed an
383
+ `affine_combination` with sum of weights 1 for the corresponding
384
+ indexed family whose index type is the subtype corresponding to that
385
+ subset. -/
386
+ lemma eq_affine_combination_subset_iff_eq_affine_combination_subtype {p0 : P} {s : set ι}
387
+ {p : ι → P} :
388
+ (∃ (fs : finset ι) (hfs : ↑fs ⊆ s) (w : ι → k) (hw : ∑ i in fs, w i = 1 ),
389
+ p0 = fs.affine_combination V w p) ↔
390
+ ∃ (fs : finset s) (w : s → k) (hw : ∑ i in fs, w i = 1 ),
391
+ p0 = fs.affine_combination V w (λ (i : s), p i) :=
392
+ begin
393
+ simp_rw [affine_combination_apply, eq_vadd_iff_vsub_eq],
394
+ exact eq_weighted_vsub_of_point_subset_iff_eq_weighted_vsub_of_point_subtype
395
+ end
396
+
336
397
end finset
337
398
338
399
section affine_independent
@@ -402,6 +463,79 @@ begin
402
463
exact finset.eq_zero_of_sum_eq_zero hw h2b i hi }
403
464
end
404
465
466
+ /-- A family is affinely independent if and only if any affine
467
+ combinations (with sum of weights 1) that evaluate to the same point
468
+ have equal `set.indicator`. -/
469
+ lemma affine_independent_iff_indicator_eq_of_affine_combination_eq (p : ι → P) :
470
+ affine_independent k V p ↔ ∀ (s1 s2 : finset ι) (w1 w2 : ι → k), ∑ i in s1, w1 i = 1 →
471
+ ∑ i in s2, w2 i = 1 → s1.affine_combination V w1 p = s2.affine_combination V w2 p →
472
+ set.indicator ↑s1 w1 = set.indicator ↑s2 w2 :=
473
+ begin
474
+ split,
475
+ { intros ha s1 s2 w1 w2 hw1 hw2 heq,
476
+ ext i,
477
+ by_cases hi : i ∈ (s1 ∪ s2),
478
+ { rw ←sub_eq_zero,
479
+ rw set.sum_indicator_subset _ (finset.subset_union_left s1 s2) at hw1,
480
+ rw set.sum_indicator_subset _ (finset.subset_union_right s1 s2) at hw2,
481
+ have hws : ∑ i in s1 ∪ s2, (set.indicator ↑s1 w1 - set.indicator ↑s2 w2) i = 0 ,
482
+ { simp [hw1, hw2] },
483
+ rw [finset.affine_combination_indicator_subset _ _ _ (finset.subset_union_left s1 s2),
484
+ finset.affine_combination_indicator_subset _ _ _ (finset.subset_union_right s1 s2),
485
+ ←vsub_eq_zero_iff_eq V, finset.affine_combination_vsub] at heq,
486
+ exact ha (s1 ∪ s2) (set.indicator ↑s1 w1 - set.indicator ↑s2 w2) hws heq i hi },
487
+ { rw [←finset.mem_coe, finset.coe_union] at hi,
488
+ simp [mt (set.mem_union_left ↑s2) hi, mt (set.mem_union_right ↑s1) hi] } },
489
+ { intros ha s w hw hs i0 hi0,
490
+ let w1 : ι → k := function.update (function.const ι 0 ) i0 1 ,
491
+ have hw1 : ∑ i in s, w1 i = 1 ,
492
+ { rw [finset.sum_update_of_mem hi0, finset.sum_const_zero, add_zero] },
493
+ have hw1s : s.affine_combination V w1 p = p i0 :=
494
+ s.affine_combination_of_eq_one_of_eq_zero V w1 p hi0 (function.update_same _ _ _)
495
+ (λ _ _ hne, function.update_noteq hne _ _),
496
+ let w2 := w + w1,
497
+ have hw2 : ∑ i in s, w2 i = 1 ,
498
+ { simp [w2, finset.sum_add_distrib, hw, hw1] },
499
+ have hw2s : s.affine_combination V w2 p = p i0,
500
+ { simp [w2, ←finset.weighted_vsub_vadd_affine_combination, hs, hw1s] },
501
+ replace ha := ha s s w2 w1 hw2 hw1 (hw1s.symm ▸ hw2s),
502
+ have hws : w2 i0 - w1 i0 = 0 ,
503
+ { rw ←finset.mem_coe at hi0,
504
+ rw [←set.indicator_of_mem hi0 w2, ←set.indicator_of_mem hi0 w1, ha, sub_self] },
505
+ simpa [w2] using hws }
506
+ end
507
+
508
+ variables {k V}
509
+
510
+ /-- If a family is affinely independent, so is any subfamily given by
511
+ composition of an embedding into index type with the original
512
+ family. -/
513
+ lemma affine_independent_embedding_of_affine_independent {ι2 : Type *} (f : ι2 ↪ ι) {p : ι → P}
514
+ (ha : affine_independent k V p) : affine_independent k V (p ∘ f) :=
515
+ begin
516
+ intros fs w hw hs i0 hi0,
517
+ let fs' := fs.map f,
518
+ let w' := λ i, if h : ∃ i2, f i2 = i then w h.some else 0 ,
519
+ have hw' : ∀ i2 : ι2 , w' (f i2) = w i2,
520
+ { intro i2,
521
+ have h : ∃ i : ι2 , f i = f i2 := ⟨i2, rfl⟩,
522
+ have hs : h.some = i2 := f.injective h.some_spec,
523
+ simp_rw [w', dif_pos h, hs] },
524
+ have hw's : ∑ i in fs', w' i = 0 ,
525
+ { rw [←hw, finset.sum_map],
526
+ simp [hw'] },
527
+ have hs' : fs'.weighted_vsub V p w' = 0 ,
528
+ { rw [←hs, finset.weighted_vsub_apply, finset.weighted_vsub_apply, finset.sum_map],
529
+ simp [hw'] },
530
+ rw [←ha fs' w' hw's hs' (f i0) ((finset.mem_map' _).2 hi0), hw']
531
+ end
532
+
533
+ /-- If a family is affinely independent, so is any subfamily indexed
534
+ by a subtype of the index type. -/
535
+ lemma affine_independent_subtype_of_affine_independent {p : ι → P}
536
+ (ha : affine_independent k V p) (s : set ι) : affine_independent k V (λ i : s, p i) :=
537
+ affine_independent_embedding_of_affine_independent (function.embedding.subtype _) ha
538
+
405
539
end affine_independent
406
540
407
541
namespace affine_space
@@ -763,6 +897,31 @@ mem_span_points k V p s hp
763
897
764
898
end affine_span
765
899
900
+ namespace affine_space
901
+
902
+ variables (k : Type *) (V : Type *) {P : Type *} [field k] [add_comm_group V] [module k V]
903
+ [affine_space k V P]
904
+ variables {ι : Type *}
905
+
906
+ /-- The `vector_span` of a finite set is finite-dimensional. -/
907
+ lemma finite_dimensional_vector_span_of_finite {s : set P} (h : set.finite s) :
908
+ finite_dimensional k (vector_span k V s) :=
909
+ finite_dimensional.span_of_finite k $ vsub_set_finite_of_finite V h
910
+
911
+ /-- The direction of the affine span of a finite set is
912
+ finite-dimensional. -/
913
+ lemma finite_dimensional_direction_affine_span_of_finite {s : set P} (h : set.finite s) :
914
+ finite_dimensional k (affine_span k V s).direction :=
915
+ (direction_affine_span k V s).symm ▸ finite_dimensional_vector_span_of_finite k V h
916
+
917
+ /-- The direction of the affine span of a family indexed by a
918
+ `fintype` is finite-dimensional. -/
919
+ instance finite_dimensional_direction_affine_span_of_fintype [fintype ι] (p : ι → P) :
920
+ finite_dimensional k (affine_span k V (set.range p)).direction :=
921
+ finite_dimensional_direction_affine_span_of_finite k V (set.finite_range _)
922
+
923
+ end affine_space
924
+
766
925
namespace affine_subspace
767
926
768
927
variables {k : Type *} {V : Type *} {P : Type *} [ring k] [add_comm_group V] [module k V]
@@ -873,6 +1032,12 @@ begin
873
1032
simp
874
1033
end
875
1034
1035
+ /-- A point is in the affine span of a single point if and only if
1036
+ they are equal. -/
1037
+ @[simp] lemma mem_affine_span_singleton (p1 p2 : P) :
1038
+ p1 ∈ affine_span k V ({p2} : set P) ↔ p1 = p2 :=
1039
+ by simp [←mem_coe]
1040
+
876
1041
/-- The span of a union of sets is the sup of their spans. -/
877
1042
lemma span_union (s t : set P) : affine_span k V (s ∪ t) = affine_span k V s ⊔ affine_span k V t :=
878
1043
(affine_subspace.gi k V P).gc.l_sup
@@ -1035,6 +1200,15 @@ begin
1035
1200
{ exact λ h, h.symm ▸ hp }
1036
1201
end
1037
1202
1203
+ /-- Coercing a subspace to a set then taking the affine span produces
1204
+ the original subspace. -/
1205
+ @[simp] lemma affine_span_coe (s : affine_subspace k V P) : affine_span k V (s : set P) = s :=
1206
+ begin
1207
+ refine le_antisymm _ (subset_span_points _ _ _),
1208
+ rintros p ⟨p1, hp1, v, hv, rfl⟩,
1209
+ exact vadd_mem_of_mem_direction hv hp1
1210
+ end
1211
+
1038
1212
end affine_subspace
1039
1213
1040
1214
namespace affine_space
@@ -1248,6 +1422,77 @@ end
1248
1422
1249
1423
end affine_space
1250
1424
1425
+ section affine_independent
1426
+
1427
+ open affine_space
1428
+
1429
+ variables {k : Type *} {V : Type *} {P : Type *} [ring k] [add_comm_group V] [module k V]
1430
+ variables [affine_space k V P] {ι : Type *}
1431
+
1432
+ /-- If a family is affinely independent, and the spans of points
1433
+ indexed by two subsets of the index type have a point in common, those
1434
+ subsets of the index type have an element in common, if the underlying
1435
+ ring is nontrivial. -/
1436
+ lemma exists_mem_inter_of_exists_mem_inter_affine_span_of_affine_independent [nontrivial k]
1437
+ {p : ι → P} (ha : affine_independent k V p) {s1 s2 : set ι} {p0 : P}
1438
+ (hp0s1 : p0 ∈ affine_span k V (p '' s1)) (hp0s2 : p0 ∈ affine_span k V (p '' s2)):
1439
+ ∃ (i : ι), i ∈ s1 ∩ s2 :=
1440
+ begin
1441
+ rw set.image_eq_range at hp0s1 hp0s2,
1442
+ rw [mem_affine_span_iff_eq_affine_combination,
1443
+ ←finset.eq_affine_combination_subset_iff_eq_affine_combination_subtype] at hp0s1 hp0s2,
1444
+ rcases hp0s1 with ⟨fs1, hfs1, w1, hw1, hp0s1⟩,
1445
+ rcases hp0s2 with ⟨fs2, hfs2, w2, hw2, hp0s2⟩,
1446
+ rw affine_independent_iff_indicator_eq_of_affine_combination_eq at ha,
1447
+ replace ha := ha fs1 fs2 w1 w2 hw1 hw2 (hp0s1 ▸ hp0s2),
1448
+ have hnz : ∑ i in fs1, w1 i ≠ 0 := hw1.symm ▸ one_ne_zero,
1449
+ rcases finset.exists_ne_zero_of_sum_ne_zero hnz with ⟨i, hifs1, hinz⟩,
1450
+ simp_rw [←set.indicator_of_mem (finset.mem_coe.2 hifs1) w1, ha] at hinz,
1451
+ use [i, hfs1 hifs1, hfs2 (set.mem_of_indicator_ne_zero hinz)]
1452
+ end
1453
+
1454
+ /-- If a family is affinely independent, the spans of points indexed
1455
+ by disjoint subsets of the index type are disjoint, if the underlying
1456
+ ring is nontrivial. -/
1457
+ lemma affine_span_disjoint_of_disjoint_of_affine_independent [nontrivial k] {p : ι → P}
1458
+ (ha : affine_independent k V p) {s1 s2 : set ι} (hd : s1 ∩ s2 = ∅) :
1459
+ (affine_span k V (p '' s1) : set P) ∩ affine_span k V (p '' s2) = ∅ :=
1460
+ begin
1461
+ by_contradiction hne,
1462
+ change (affine_span k V (p '' s1) : set P) ∩ affine_span k V (p '' s2) ≠ ∅ at hne,
1463
+ rw set.ne_empty_iff_nonempty at hne,
1464
+ rcases hne with ⟨p0, hp0s1, hp0s2⟩,
1465
+ cases exists_mem_inter_of_exists_mem_inter_affine_span_of_affine_independent
1466
+ ha hp0s1 hp0s2 with i hi,
1467
+ exact set.not_mem_empty i (hd ▸ hi)
1468
+ end
1469
+
1470
+ /-- If a family is affinely independent, a point in the family is in
1471
+ the span of some of the points given by a subset of the index type if
1472
+ and only if that point's index is in the subset, if the underlying
1473
+ ring is nontrivial. -/
1474
+ @[simp] lemma mem_affine_span_iff_mem_of_affine_independent [nontrivial k] {p : ι → P}
1475
+ (ha : affine_independent k V p) (i : ι) (s : set ι) :
1476
+ p i ∈ affine_span k V (p '' s) ↔ i ∈ s :=
1477
+ begin
1478
+ split,
1479
+ { intro hs,
1480
+ have h := exists_mem_inter_of_exists_mem_inter_affine_span_of_affine_independent
1481
+ ha hs (mem_affine_span k V (set.mem_image_of_mem _ (set.mem_singleton _))),
1482
+ rwa [←set.nonempty_def, set.inter_singleton_nonempty] at h },
1483
+ { exact λ h, mem_affine_span k V (set.mem_image_of_mem p h) }
1484
+ end
1485
+
1486
+ /-- If a family is affinely independent, a point in the family is not
1487
+ in the affine span of the other points, if the underlying ring is
1488
+ nontrivial. -/
1489
+ lemma not_mem_affine_span_diff_of_affine_independent [nontrivial k] {p : ι → P}
1490
+ (ha : affine_independent k V p) (i : ι) (s : set ι) :
1491
+ p i ∉ affine_span k V (p '' (s \ {i})) :=
1492
+ by simp [ha]
1493
+
1494
+ end affine_independent
1495
+
1251
1496
namespace affine_subspace
1252
1497
1253
1498
variables {k : Type *} {V : Type *} {P : Type *} [ring k] [add_comm_group V] [module k V]
0 commit comments