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

Commit 67da272

Browse files
committed
feat(analysis/inner_product_space): Gram-Schmidt Basis (#14514)
When the Gram-Schmidt procedure is given a basis, it produces a basis. This pull request also refactors the lemma `gram_schmidt_span`. The new versions of the lemmas cover the span of `Iio`, the span of `Iic` and the span of the complete set of input vectors. I was also able to remove some type classes in the process.
1 parent 988f160 commit 67da272

File tree

2 files changed

+112
-46
lines changed

2 files changed

+112
-46
lines changed

src/analysis/inner_product_space/gram_schmidt_ortho.lean

Lines changed: 109 additions & 46 deletions
Original file line numberDiff line numberDiff line change
@@ -1,11 +1,12 @@
11
/-
22
Copyright (c) 2022 Jiale Miao. All rights reserved.
33
Released under Apache 2.0 license as described in the file LICENSE.
4-
Authors: Jiale Miao, Kevin Buzzard
4+
Authors: Jiale Miao, Kevin Buzzard, Alexander Bentkamp
55
-/
66

77
import analysis.inner_product_space.projection
88
import order.well_founded_set
9+
import analysis.inner_product_space.pi_L2
910

1011
/-!
1112
# Gram-Schmidt Orthogonalization and Orthonormalization
@@ -25,6 +26,8 @@ and outputs a set of orthogonal vectors which have the same span.
2526
- `gram_schmidt_ne_zero` :
2627
If the input vectors of `gram_schmidt` are linearly independent,
2728
then the output vectors are non-zero.
29+
- `gram_schmidt_basis` :
30+
The basis produced by the Gram-Schmidt process when given a basis as input.
2831
- `gram_schmidt_normed` :
2932
the normalized `gram_schmidt` (i.e each vector in `gram_schmidt_normed` has unit length.)
3033
- `gram_schmidt_orthornormal` :
@@ -38,8 +41,7 @@ open_locale big_operators
3841
open finset
3942

4043
variables (𝕜 : Type*) {E : Type*} [is_R_or_C 𝕜] [inner_product_space 𝕜 E]
41-
variables {ι : Type*} [linear_order ι] [order_bot ι]
42-
variables [locally_finite_order ι] [is_well_order ι (<)]
44+
variables {ι : Type*} [linear_order ι] [locally_finite_order_bot ι] [is_well_order ι (<)]
4345

4446
local attribute [instance] is_well_order.to_has_well_founded
4547

@@ -49,7 +51,7 @@ local notation `⟪`x`, `y`⟫` := @inner 𝕜 _ _ x y
4951
and outputs a set of orthogonal vectors which have the same span. -/
5052
noncomputable def gram_schmidt (f : ι → E) : ι → E
5153
| n := f n - ∑ i : Iio n, orthogonal_projection (𝕜 ∙ gram_schmidt i) (f n)
52-
using_well_founded { dec_tac := `[exact (mem_Ico.1 i.2).2] }
54+
using_well_founded { dec_tac := `[exact mem_Iio.1 i.2] }
5355

5456
/-- This lemma uses `∑ i in` instead of `∑ i :`.-/
5557
lemma gram_schmidt_def (f : ι → E) (n : ι):
@@ -62,7 +64,8 @@ lemma gram_schmidt_def' (f : ι → E) (n : ι):
6264
orthogonal_projection (𝕜 ∙ gram_schmidt 𝕜 f i) (f n) :=
6365
by rw [gram_schmidt_def, sub_add_cancel]
6466

65-
@[simp] lemma gram_schmidt_zero (f : ι → E) : gram_schmidt 𝕜 f ⊥ = f ⊥ :=
67+
@[simp] lemma gram_schmidt_zero {ι : Type*} [linear_order ι] [locally_finite_order ι]
68+
[order_bot ι] [is_well_order ι (<)] (f : ι → E) : gram_schmidt 𝕜 f ⊥ = f ⊥ :=
6669
by rw [gram_schmidt_def, Iio_eq_Ico, finset.Ico_self, finset.sum_empty, sub_zero]
6770

6871
/-- **Gram-Schmidt Orthogonalisation**:
@@ -93,7 +96,7 @@ begin
9396
cases hia.lt_or_lt with hia₁ hia₂,
9497
{ rw inner_eq_zero_sym,
9598
exact ih a h₀ i hia₁ },
96-
{ exact ih i (mem_Ico.1 hi).2 a hia₂ }
99+
{ exact ih i (mem_Iio.1 hi) a hia₂ }
97100
end
98101

99102
/-- This is another version of `gram_schmidt_orthogonal` using `pairwise` instead. -/
@@ -103,46 +106,56 @@ theorem gram_schmidt_pairwise_orthogonal (f : ι → E) :
103106

104107
open submodule set order
105108

106-
/-- `gram_schmidt` preserves span of vectors. -/
107-
lemma span_gram_schmidt [succ_order ι] [is_succ_archimedean ι] (f : ι → E) (c : ι) :
108-
span 𝕜 (gram_schmidt 𝕜 f '' Iio c) = span 𝕜 (f '' Iio c) :=
109+
lemma mem_span_gram_schmidt (f : ι → E) {i j : ι} (hij : i ≤ j) :
110+
f i ∈ span 𝕜 (gram_schmidt 𝕜 f '' Iic j) :=
111+
begin
112+
rw [gram_schmidt_def' 𝕜 f i],
113+
simp_rw orthogonal_projection_singleton,
114+
exact submodule.add_mem _ (subset_span $ mem_image_of_mem _ hij)
115+
(submodule.sum_mem _ $ λ k hk, smul_mem (span 𝕜 (gram_schmidt 𝕜 f '' Iic j)) _ $
116+
subset_span $ mem_image_of_mem (gram_schmidt 𝕜 f) $ (finset.mem_Iio.1 hk).le.trans hij),
117+
end
118+
119+
lemma gram_schmidt_mem_span (f : ι → E) :
120+
∀ {j i}, i ≤ j → gram_schmidt 𝕜 f i ∈ span 𝕜 (f '' Iic j)
121+
| j := λ i hij,
109122
begin
110-
apply @succ.rec ι _ _ _ (λ c, span 𝕜 (gram_schmidt 𝕜 f '' Iio c) = span 𝕜 (f '' Iio c)) ⊥
111-
_ _ _ bot_le,
112-
{ simp only [set.Iio_bot, set.image_empty] },
113-
intros c _ hc,
114-
by_cases h : succ c = c,
115-
{ rwa h },
116-
have h₀ : ∀ b, b ∈ finset.Iio c → gram_schmidt 𝕜 f b ∈ span 𝕜 (f '' Iio c),
117-
{ simp_intros b hb only [finset.mem_range, nat.succ_eq_add_one],
118-
rw ← hc,
119-
refine subset_span _,
120-
simp only [set.mem_image, set.mem_Iio],
121-
refine ⟨b, (finset.mem_Ico.1 hb).2, by refl⟩ },
122-
rw not_iff_not.2 order.succ_eq_iff_is_max at h,
123-
rw [order.Iio_succ_eq_insert_of_not_is_max h],
124-
simp only [span_insert, image_insert_eq, hc],
125-
apply le_antisymm,
126-
{ simp only [nat.succ_eq_succ,gram_schmidt_def 𝕜 f c, orthogonal_projection_singleton,
127-
sup_le_iff, span_singleton_le_iff_mem, le_sup_right, and_true],
128-
apply submodule.sub_mem _ _ _,
129-
{ exact mem_sup_left (mem_span_singleton_self (f c)) },
130-
{ exact submodule.sum_mem _ (λ b hb, mem_sup_right (smul_mem _ _ (h₀ b hb))) } },
131-
{ rw [gram_schmidt_def' 𝕜 f c],
132-
simp only [orthogonal_projection_singleton,
133-
sup_le_iff, span_singleton_le_iff_mem, le_sup_right, and_true],
134-
apply submodule.add_mem _ _ _,
135-
{ exact mem_sup_left (mem_span_singleton_self (gram_schmidt 𝕜 f c)), },
136-
{ exact submodule.sum_mem _ (λ b hb, mem_sup_right (smul_mem _ _ (h₀ b hb))) } }
123+
rw [gram_schmidt_def 𝕜 f i],
124+
simp_rw orthogonal_projection_singleton,
125+
refine submodule.sub_mem _ (subset_span (mem_image_of_mem _ hij))
126+
(submodule.sum_mem _ $ λ k hk, _),
127+
let hkj : k < j := (finset.mem_Iio.1 hk).trans_le hij,
128+
exact smul_mem _ _ (span_mono (image_subset f $ Iic_subset_Iic.2 hkj.le) $
129+
gram_schmidt_mem_span le_rfl),
137130
end
131+
using_well_founded { dec_tac := `[assumption] }
132+
133+
lemma span_gram_schmidt_Iic (f : ι → E) (c : ι) :
134+
span 𝕜 (gram_schmidt 𝕜 f '' Iic c) = span 𝕜 (f '' Iic c) :=
135+
span_eq_span (set.image_subset_iff.2 $ λ i, gram_schmidt_mem_span _ _) $
136+
set.image_subset_iff.2 $ λ i, mem_span_gram_schmidt _ _
138137

139-
lemma gram_schmidt_ne_zero_coe [succ_order ι] [is_succ_archimedean ι]
138+
lemma span_gram_schmidt_Iio (f : ι → E) (c : ι) :
139+
span 𝕜 (gram_schmidt 𝕜 f '' Iio c) = span 𝕜 (f '' Iio c) :=
140+
span_eq_span
141+
(set.image_subset_iff.2 $ λ i hi, span_mono (image_subset _ $ Iic_subset_Iio.2 hi) $
142+
gram_schmidt_mem_span _ _ le_rfl) $
143+
set.image_subset_iff.2 $ λ i hi, span_mono (image_subset _ $ Iic_subset_Iio.2 hi) $
144+
mem_span_gram_schmidt _ _ le_rfl
145+
146+
/-- `gram_schmidt` preserves span of vectors. -/
147+
lemma span_gram_schmidt (f : ι → E) : span 𝕜 (range (gram_schmidt 𝕜 f)) = span 𝕜 (range f) :=
148+
span_eq_span (range_subset_iff.2 $ λ i, span_mono (image_subset_range _ _) $
149+
gram_schmidt_mem_span _ _ le_rfl) $
150+
range_subset_iff.2 $ λ i, span_mono (image_subset_range _ _) $ mem_span_gram_schmidt _ _ le_rfl
151+
152+
lemma gram_schmidt_ne_zero_coe
140153
(f : ι → E) (n : ι) (h₀ : linear_independent 𝕜 (f ∘ (coe : set.Iic n → ι))) :
141154
gram_schmidt 𝕜 f n ≠ 0 :=
142155
begin
143156
by_contra h,
144157
have h₁ : f n ∈ span 𝕜 (f '' Iio n),
145-
{ rw [← span_gram_schmidt 𝕜 f n, gram_schmidt_def' _ f, h, zero_add],
158+
{ rw [← span_gram_schmidt_Iio 𝕜 f n, gram_schmidt_def' _ f, h, zero_add],
146159
apply submodule.sum_mem _ _,
147160
simp_intros a ha only [finset.mem_Ico],
148161
simp only [set.mem_image, set.mem_Iio, orthogonal_projection_singleton],
@@ -161,39 +174,89 @@ end
161174

162175
/-- If the input vectors of `gram_schmidt` are linearly independent,
163176
then the output vectors are non-zero. -/
164-
lemma gram_schmidt_ne_zero [succ_order ι] [is_succ_archimedean ι]
165-
(f : ι → E) (n : ι) (h₀ : linear_independent 𝕜 f) :
177+
lemma gram_schmidt_ne_zero (f : ι → E) (n : ι) (h₀ : linear_independent 𝕜 f) :
166178
gram_schmidt 𝕜 f n ≠ 0 :=
167179
gram_schmidt_ne_zero_coe _ _ _ (linear_independent.comp h₀ _ subtype.coe_injective)
168180

181+
/-- `gram_schmidt` produces a triangular matrix of vectors when given a basis. -/
182+
lemma gram_schmidt_triangular {i j : ι} (hij : i < j) (b : basis ι 𝕜 E) :
183+
b.repr (gram_schmidt 𝕜 b i) j = 0 :=
184+
begin
185+
have : gram_schmidt 𝕜 b i ∈ span 𝕜 (gram_schmidt 𝕜 b '' set.Iio j),
186+
from subset_span ((set.mem_image _ _ _).2 ⟨i, hij, rfl⟩),
187+
have : gram_schmidt 𝕜 b i ∈ span 𝕜 (b '' set.Iio j),
188+
by rwa [← span_gram_schmidt_Iio 𝕜 b j],
189+
have : ↑(((b.repr) (gram_schmidt 𝕜 b i)).support) ⊆ set.Iio j,
190+
from basis.repr_support_subset_of_mem_span b (set.Iio j) this,
191+
exact (finsupp.mem_supported' _ _).1
192+
((finsupp.mem_supported 𝕜 _).2 this) j (not_mem_Iio.2 (le_refl j)),
193+
end
194+
195+
/-- `gram_schmidt` produces linearly independent vectors when given linearly independent vectors. -/
196+
lemma gram_schmidt_linear_independent (f : ι → E) (h₀ : linear_independent 𝕜 f) :
197+
linear_independent 𝕜 (gram_schmidt 𝕜 f) :=
198+
linear_independent_of_ne_zero_of_inner_eq_zero
199+
(λ i, gram_schmidt_ne_zero _ _ _ h₀) (λ i j, gram_schmidt_orthogonal 𝕜 f)
200+
201+
/-- When given a basis, `gram_schmidt` produces a basis. -/
202+
noncomputable def gram_schmidt_basis (b : basis ι 𝕜 E) : basis ι 𝕜 E :=
203+
basis.mk
204+
(gram_schmidt_linear_independent 𝕜 b b.linear_independent)
205+
((span_gram_schmidt 𝕜 b).trans b.span_eq)
206+
207+
lemma coe_gram_schmidt_basis (b : basis ι 𝕜 E) :
208+
(gram_schmidt_basis 𝕜 b : ι → E) = gram_schmidt 𝕜 b := basis.coe_mk _ _
209+
169210
/-- the normalized `gram_schmidt`
170211
(i.e each vector in `gram_schmidt_normed` has unit length.) -/
171212
noncomputable def gram_schmidt_normed (f : ι → E) (n : ι) : E :=
172213
(∥gram_schmidt 𝕜 f n∥ : 𝕜)⁻¹ • (gram_schmidt 𝕜 f n)
173214

174-
lemma gram_schmidt_normed_unit_length_coe [succ_order ι] [is_succ_archimedean ι]
215+
lemma gram_schmidt_normed_unit_length_coe
175216
(f : ι → E) (n : ι) (h₀ : linear_independent 𝕜 (f ∘ (coe : set.Iic n → ι))) :
176217
∥gram_schmidt_normed 𝕜 f n∥ = 1 :=
177218
by simp only [gram_schmidt_ne_zero_coe 𝕜 f n h₀,
178219
gram_schmidt_normed, norm_smul_inv_norm, ne.def, not_false_iff]
179220

180-
lemma gram_schmidt_normed_unit_length [succ_order ι] [is_succ_archimedean ι]
181-
(f : ι → E) (n : ι) (h₀ : linear_independent 𝕜 f) :
221+
lemma gram_schmidt_normed_unit_length (f : ι → E) (n : ι) (h₀ : linear_independent 𝕜 f) :
182222
∥gram_schmidt_normed 𝕜 f n∥ = 1 :=
183223
gram_schmidt_normed_unit_length_coe _ _ _ (linear_independent.comp h₀ _ subtype.coe_injective)
184224

185225
/-- **Gram-Schmidt Orthonormalization**:
186226
`gram_schmidt_normed` produces an orthornormal system of vectors. -/
187-
theorem gram_schmidt_orthonormal [succ_order ι] [is_succ_archimedean ι]
188-
(f : ι → E) (h₀ : linear_independent 𝕜 f) :
227+
theorem gram_schmidt_orthonormal (f : ι → E) (h₀ : linear_independent 𝕜 f) :
189228
orthonormal 𝕜 (gram_schmidt_normed 𝕜 f) :=
190229
begin
191230
unfold orthonormal,
192231
split,
193-
{ simp only [gram_schmidt_normed_unit_length, h₀, forall_const] },
232+
{ simp only [gram_schmidt_normed_unit_length, h₀, eq_self_iff_true, implies_true_iff], },
194233
{ intros i j hij,
195234
simp only [gram_schmidt_normed, inner_smul_left, inner_smul_right, is_R_or_C.conj_inv,
196235
is_R_or_C.conj_of_real, mul_eq_zero, inv_eq_zero, is_R_or_C.of_real_eq_zero, norm_eq_zero],
197236
repeat { right },
198237
exact gram_schmidt_orthogonal 𝕜 f hij }
199238
end
239+
240+
lemma span_gram_schmidt_normed (f : ι → E) (s : set ι) :
241+
span 𝕜 (gram_schmidt_normed 𝕜 f '' s) = span 𝕜 (gram_schmidt 𝕜 f '' s) :=
242+
begin
243+
refine span_eq_span (set.image_subset_iff.2 $ λ i hi, smul_mem _ _ $ subset_span $
244+
mem_image_of_mem _ hi)
245+
(set.image_subset_iff.2 $ λ i hi, span_mono (image_subset _ $ singleton_subset_set_iff.2 hi) _),
246+
simp only [coe_singleton, set.image_singleton],
247+
by_cases h : gram_schmidt 𝕜 f i = 0,
248+
{ simp [h] },
249+
{ refine mem_span_singleton.2 ⟨∥gram_schmidt 𝕜 f i∥, smul_inv_smul₀ _ _⟩,
250+
exact_mod_cast (norm_ne_zero_iff.2 h) }
251+
end
252+
253+
lemma span_gram_schmidt_normed_range (f : ι → E) :
254+
span 𝕜 (range (gram_schmidt_normed 𝕜 f)) = span 𝕜 (range (gram_schmidt 𝕜 f)) :=
255+
by simpa only [image_univ.symm] using span_gram_schmidt_normed 𝕜 f univ
256+
257+
/-- When given a basis, `gram_schmidt_normed` produces an orthonormal basis. -/
258+
noncomputable def gram_schmidt_orthonormal_basis [fintype ι] (b : basis ι 𝕜 E) :
259+
orthonormal_basis ι 𝕜 E :=
260+
orthonormal_basis.mk
261+
(gram_schmidt_orthonormal 𝕜 b b.linear_independent)
262+
(((span_gram_schmidt_normed_range 𝕜 b).trans (span_gram_schmidt 𝕜 b)).trans b.span_eq)

src/linear_algebra/span.lean

Lines changed: 3 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -57,6 +57,9 @@ le_antisymm (span_le.2 h₁) h₂
5757
lemma span_eq : span R (p : set M) = p :=
5858
span_eq_of_le _ (subset.refl _) subset_span
5959

60+
lemma span_eq_span (hs : s ⊆ span R t) (ht : t ⊆ span R s) : span R s = span R t :=
61+
le_antisymm (span_le.2 hs) (span_le.2 ht)
62+
6063
/-- A version of `submodule.span_eq` for when the span is by a smaller ring. -/
6164
@[simp] lemma span_coe_eq_restrict_scalars
6265
[semiring S] [has_scalar S R] [module S M] [is_scalar_tower S R M] :

0 commit comments

Comments
 (0)