Skip to content

Commit

Permalink
feat(linear_algebra): elements of a dim zero space (#3054)
Browse files Browse the repository at this point in the history
Co-authored-by: Rob Lewis <rob.y.lewis@gmail.com>
  • Loading branch information
robertylewis and robertylewis committed Jun 13, 2020
1 parent 51ad2a1 commit 1645542
Show file tree
Hide file tree
Showing 3 changed files with 68 additions and 0 deletions.
20 changes: 20 additions & 0 deletions src/linear_algebra/dimension.lean
Original file line number Diff line number Diff line change
Expand Up @@ -412,6 +412,26 @@ by rw [rank, rank, linear_map.range_comp]; exact dim_map_le _ _

end rank

lemma dim_zero_iff_forall_zero : vector_space.dim K V = 0 ↔ ∀ x : V, x = 0 :=
begin
split,
{ intros h x,
cases exists_is_basis K V with w hw,
have card_mk_range := hw.mk_range_eq_dim,
rw [h, cardinal.mk_emptyc_iff, set.range_coe_subtype] at card_mk_range,
simpa [card_mk_range] using hw.mem_span x },
{ intro h,
have : (⊤ : submodule K V) = ⊥,
{ ext x, simp [h x] },
rw [←dim_top, this, dim_bot] }
end

lemma dim_pos_iff_exists_ne_zero : 0 < vector_space.dim K V ↔ ∃ x : V, x ≠ 0 :=
begin
rw ←not_iff_not,
simpa using dim_zero_iff_forall_zero
end

end vector_space

section unconstrained_universes
Expand Down
37 changes: 37 additions & 0 deletions src/linear_algebra/finite_dimensional.lean
Original file line number Diff line number Diff line change
Expand Up @@ -366,3 +366,40 @@ theorem findim_range_add_findim_ker [finite_dimensional K V] (f : V →ₗ[K] V
by { rw [← f.quot_ker_equiv_range.findim_eq], exact submodule.findim_quotient_add_findim _ }

end linear_map

section zero_dim

open vector_space finite_dimensional

lemma finite_dimensional_of_dim_eq_zero (h : vector_space.dim K V = 0) : finite_dimensional K V :=
by rw [finite_dimensional_iff_dim_lt_omega, h]; exact cardinal.omega_pos

lemma findim_eq_zero_of_dim_eq_zero [finite_dimensional K V] (h : vector_space.dim K V = 0) :
findim K V = 0 :=
begin
convert findim_eq_dim K V,
rw h, norm_cast
end

variables (K V)

lemma finite_dimensional_bot : finite_dimensional K (⊥ : submodule K V) :=
finite_dimensional_of_dim_eq_zero $ by simp

lemma findim_bot : findim K (⊥ : submodule K V) = 0 :=
begin
haveI := finite_dimensional_bot K V,
convert findim_eq_dim K (⊥ : submodule K V),
rw dim_bot, norm_cast
end

variables {K V}

lemma bot_eq_top_of_dim_eq_zero (h : vector_space.dim K V = 0) : (⊥ : submodule K V) = ⊤ :=
begin
haveI := finite_dimensional_of_dim_eq_zero h,
apply eq_top_of_findim_eq,
rw [findim_bot, findim_eq_zero_of_dim_eq_zero h]
end

end zero_dim
11 changes: 11 additions & 0 deletions src/set_theory/cardinal.lean
Original file line number Diff line number Diff line change
Expand Up @@ -853,6 +853,17 @@ theorem mk_subtype_le_of_subset {α : Type u} {p q : α → Prop} (h : ∀ ⦃x
@[simp] theorem mk_emptyc (α : Type u) : mk (∅ : set α) = 0 :=
quotient.sound ⟨equiv.set.pempty α⟩

lemma mk_emptyc_iff {α : Type u} {s : set α} : mk s = 0 ↔ s = ∅ :=
begin
split,
{ intro h,
have h2 : cardinal.mk s = cardinal.mk pempty, by simp [h],
refine set.eq_empty_iff_forall_not_mem.mpr (λ _ hx, _),
rcases cardinal.eq.mp h2 with ⟨f, _⟩,
cases f ⟨_, hx⟩ },
{ intro, convert mk_emptyc _ }
end

theorem mk_univ {α : Type u} : mk (@univ α) = mk α :=
quotient.sound ⟨equiv.set.univ α⟩

Expand Down

0 comments on commit 1645542

Please sign in to comment.