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

Commit b77aa3a

Browse files
committed
feat(linear_algebra/affine_space/affine_subspace): prove that a set whose affine span is top cannot be empty. (#9113)
The lemma `finset.card_sdiff_add_card` is unrelated but I've been meaning to add it and now seemed like a good time since I'm touching `data/finset/basic.lean` anyway.
1 parent f59dbf2 commit b77aa3a

File tree

3 files changed

+26
-0
lines changed

3 files changed

+26
-0
lines changed

src/data/finset/basic.lean

Lines changed: 5 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -313,6 +313,8 @@ protected def nonempty (s : finset α) : Prop := ∃ x:α, x ∈ s
313313

314314
@[simp, norm_cast] lemma coe_nonempty {s : finset α} : (s:set α).nonempty ↔ s.nonempty := iff.rfl
315315

316+
@[simp] lemma nonempty_coe_sort (s : finset α) : nonempty ↥s ↔ s.nonempty := nonempty_subtype
317+
316318
alias coe_nonempty ↔ _ finset.nonempty.to_set
317319

318320
lemma nonempty.bex {s : finset α} (h : s.nonempty) : ∃ x:α, x ∈ s := h
@@ -2862,6 +2864,9 @@ theorem card_sdiff {s t : finset α} (h : s ⊆ t) : card (t \ s) = card t - car
28622864
suffices card (t \ s) = card ((t \ s) ∪ s) - card s, by rwa sdiff_union_of_subset h at this,
28632865
by rw [card_disjoint_union sdiff_disjoint, nat.add_sub_cancel]
28642866

2867+
lemma card_sdiff_add_card {s t : finset α} : (s \ t).card + t.card = (s ∪ t).card :=
2868+
by rw [← card_disjoint_union sdiff_disjoint, sdiff_union_self_eq_union]
2869+
28652870
lemma disjoint_filter {s : finset α} {p q : α → Prop} [decidable_pred p] [decidable_pred q] :
28662871
disjoint (s.filter p) (s.filter q) ↔ (∀ x ∈ s, p x → ¬ q x) :=
28672872
by split; simp [disjoint_left] {contextual := tt}

src/data/set/basic.lean

Lines changed: 2 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -291,6 +291,8 @@ in theorem assumptions instead of `∃ x, x ∈ s` or `s ≠ ∅` as it gives ac
291291
to the dot notation. -/
292292
protected def nonempty (s : set α) : Prop := ∃ x, x ∈ s
293293

294+
@[simp] lemma nonempty_coe_sort (s : set α) : nonempty ↥s ↔ s.nonempty := nonempty_subtype
295+
294296
lemma nonempty_def : s.nonempty ↔ ∃ x, x ∈ s := iff.rfl
295297

296298
lemma nonempty_of_mem {x} (h : x ∈ s) : s.nonempty := ⟨x, h⟩

src/linear_algebra/affine_space/affine_subspace.lean

Lines changed: 19 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -351,6 +351,10 @@ begin
351351
exact h
352352
end
353353

354+
@[simp] lemma ext_iff (s₁ s₂ : affine_subspace k P) :
355+
(s₁ : set P) = s₂ ↔ s₁ = s₂ :=
356+
⟨ext, by tidy⟩
357+
354358
/-- Two affine subspaces with the same direction and nonempty
355359
intersection are equal. -/
356360
lemma ext_of_direction_eq {s1 s2 : affine_subspace k P} (hd : s1.direction = s2.direction)
@@ -667,6 +671,21 @@ end
667671
@[simp] lemma bot_coe : ((⊥ : affine_subspace k P) : set P) = ∅ :=
668672
rfl
669673

674+
lemma bot_ne_top : (⊥ : affine_subspace k P) ≠ ⊤ :=
675+
begin
676+
intros contra,
677+
rw [← ext_iff, bot_coe, top_coe] at contra,
678+
exact set.empty_ne_univ contra,
679+
end
680+
681+
lemma nonempty_of_affine_span_eq_top {s : set P} (h : affine_span k s = ⊤) : s.nonempty :=
682+
begin
683+
rw ← set.ne_empty_iff_nonempty,
684+
rintros rfl,
685+
rw affine_subspace.span_empty at h,
686+
exact bot_ne_top k V P h,
687+
end
688+
670689
variables {P}
671690

672691
/-- No points are in `⊥`. -/

0 commit comments

Comments
 (0)