Skip to content

Commit facf1f9

Browse files
4hma4deric-wieser
andcommitted
feat(LinearAlgebra/Matrix): Woodbury Identity (#16325)
This adds the [Woodbury Identity](https://en.wikipedia.org/wiki/Woodbury_matrix_identity). [Zulip discussion](https://leanprover.zulipchat.com/#narrow/stream/287929-mathlib4/topic/Woodbury.20identity/near/462245284) Also corrects some bad deprecations introduced in #16590, which affected development of this PR. Co-authored-by: Mohanad Ahmad Co-authored-by: Eric Wieser <wieser.eric@gmail.com>
1 parent b272ab8 commit facf1f9

File tree

2 files changed

+90
-5
lines changed

2 files changed

+90
-5
lines changed

Mathlib/Data/Matrix/Invertible.lean

Lines changed: 72 additions & 5 deletions
Original file line numberDiff line numberDiff line change
@@ -1,9 +1,10 @@
11
/-
22
Copyright (c) 2023 Eric Wieser. All rights reserved.
33
Released under Apache 2.0 license as described in the file LICENSE.
4-
Authors: Eric Wieser
4+
Authors: Eric Wieser, Ahmad Alkhalawi
55
-/
66
import Mathlib.Data.Matrix.Basic
7+
import Mathlib.Tactic.Abel
78

89
/-! # Extra lemmas about invertible matrices
910
@@ -47,10 +48,12 @@ protected theorem invOf_mul_cancel_right (A : Matrix m n α) (B : Matrix n n α)
4748
protected theorem mul_invOf_cancel_right (A : Matrix m n α) (B : Matrix n n α) [Invertible B] :
4849
A * B * ⅟ B = A := by rw [Matrix.mul_assoc, mul_invOf_self, Matrix.mul_one]
4950

50-
@[deprecated (since := "2024-09-07")] alias invOf_mul_self_assoc := invOf_mul_cancel_left
51-
@[deprecated (since := "2024-09-07")] alias mul_invOf_self_assoc := mul_invOf_cancel_left
52-
@[deprecated (since := "2024-09-07")] alias mul_invOf_mul_self_cancel := invOf_mul_cancel_right
53-
@[deprecated (since := "2024-09-07")] alias mul_mul_invOf_self_cancel := mul_invOf_cancel_right
51+
@[deprecated (since := "2024-09-07")] alias invOf_mul_self_assoc := Matrix.invOf_mul_cancel_left
52+
@[deprecated (since := "2024-09-07")] alias mul_invOf_self_assoc := Matrix.mul_invOf_cancel_left
53+
@[deprecated (since := "2024-09-07")]
54+
alias mul_invOf_mul_self_cancel := Matrix.invOf_mul_cancel_right
55+
@[deprecated (since := "2024-09-07")]
56+
alias mul_mul_invOf_self_cancel := Matrix.mul_invOf_cancel_right
5457

5558
section ConjTranspose
5659
variable [StarRing α] (A : Matrix n n α)
@@ -106,4 +109,68 @@ def transposeInvertibleEquivInvertible : Invertible Aᵀ ≃ Invertible A where
106109

107110
end CommSemiring
108111

112+
section Ring
113+
114+
section Woodbury
115+
116+
variable [Fintype m] [DecidableEq m] [Ring α]
117+
(A : Matrix n n α) (U : Matrix n m α) (C : Matrix m m α) (V : Matrix m n α)
118+
[Invertible A] [Invertible C] [Invertible (⅟C + V * ⅟A * U)]
119+
120+
-- No spaces around multiplication signs for better clarity
121+
lemma add_mul_mul_invOf_mul_eq_one :
122+
(A + U*C*V)*(⅟A - ⅟A*U*⅟(⅟C + V*⅟A*U)*V*⅟A) = 1 := by
123+
calc
124+
(A + U*C*V)*(⅟A - ⅟A*U*⅟(⅟C + V*⅟A*U)*V*⅟A)
125+
_ = A*⅟A - A*⅟A*U*⅟(⅟C + V*⅟A*U)*V*⅟A + U*C*V*⅟A - U*C*V*⅟A*U*⅟(⅟C + V*⅟A*U)*V*⅟A := by
126+
simp_rw [add_sub_assoc, add_mul, mul_sub, Matrix.mul_assoc]
127+
_ = (1 + U*C*V*⅟A) - (U*⅟(⅟C + V*⅟A*U)*V*⅟A + U*C*V*⅟A*U*⅟(⅟C + V*⅟A*U)*V*⅟A) := by
128+
rw [mul_invOf_self, Matrix.one_mul]
129+
abel
130+
_ = 1 + U*C*V*⅟A - (U + U*C*V*⅟A*U)*⅟(⅟C + V*⅟A*U)*V*⅟A := by
131+
rw [sub_right_inj, Matrix.add_mul, Matrix.add_mul, Matrix.add_mul]
132+
_ = 1 + U*C*V*⅟A - U*C*(⅟C + V*⅟A*U)*⅟(⅟C + V*⅟A*U)*V*⅟A := by
133+
congr
134+
simp only [Matrix.mul_add, Matrix.mul_invOf_cancel_right, ← Matrix.mul_assoc]
135+
_ = 1 := by
136+
rw [Matrix.mul_invOf_cancel_right]
137+
abel
138+
139+
/-- Like `add_mul_mul_invOf_mul_eq_one`, but with multiplication reversed. -/
140+
lemma add_mul_mul_invOf_mul_eq_one' :
141+
(⅟A - ⅟A*U*⅟(⅟C + V*⅟A*U)*V*⅟A)*(A + U*C*V) = 1 := by
142+
calc
143+
(⅟A - ⅟A*U*⅟(⅟C + V*⅟A*U)*V*⅟A)*(A + U*C*V)
144+
_ = ⅟A*A - ⅟A*U*⅟(⅟C + V*⅟A*U)*V*⅟A*A + ⅟A*U*C*V - ⅟A*U*⅟(⅟C + V*⅟A*U)*V*⅟A*U*C*V := by
145+
simp_rw [add_sub_assoc, _root_.mul_add, _root_.sub_mul, Matrix.mul_assoc]
146+
_ = (1 + ⅟A*U*C*V) - (⅟A*U*⅟(⅟C + V*⅟A*U)*V + ⅟A*U*⅟(⅟C + V*⅟A*U)*V*⅟A*U*C*V) := by
147+
rw [invOf_mul_self, Matrix.invOf_mul_cancel_right]
148+
abel
149+
_ = 1 + ⅟A*U*C*V - ⅟A*U*⅟(⅟C + V*⅟A*U)*(V + V*⅟A*U*C*V) := by
150+
rw [sub_right_inj, Matrix.mul_add]
151+
simp_rw [Matrix.mul_assoc]
152+
_ = 1 + ⅟A*U*C*V - ⅟A*U*⅟(⅟C + V*⅟A*U)*(⅟C + V*⅟A*U)*C*V := by
153+
congr 1
154+
simp only [Matrix.mul_add, Matrix.add_mul, ← Matrix.mul_assoc,
155+
Matrix.invOf_mul_cancel_right]
156+
_ = 1 := by
157+
rw [Matrix.invOf_mul_cancel_right]
158+
abel
159+
160+
/-- If matrices `A`, `C`, and `C⁻¹ + V * A⁻¹ * U` are invertible, then so is `A + U * C * V`-/
161+
def invertibleAddMulMul : Invertible (A + U*C*V) where
162+
invOf := ⅟A - ⅟A*U*⅟(⅟C + V*⅟A*U)*V*⅟A
163+
invOf_mul_self := add_mul_mul_invOf_mul_eq_one' _ _ _ _
164+
mul_invOf_self := add_mul_mul_invOf_mul_eq_one _ _ _ _
165+
166+
/-- The **Woodbury Identity** (`⅟` version). -/
167+
theorem invOf_add_mul_mul [Invertible (A + U*C*V)] :
168+
⅟(A + U*C*V) = ⅟A - ⅟A*U*⅟(⅟C + V*⅟A*U)*V*⅟A := by
169+
letI := invertibleAddMulMul A U C V
170+
convert (rfl : ⅟(A + U*C*V) = _)
171+
172+
end Woodbury
173+
174+
end Ring
175+
109176
end Matrix

Mathlib/LinearAlgebra/Matrix/NonsingularInverse.lean

Lines changed: 18 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -607,6 +607,24 @@ theorem inv_diagonal (v : n → α) : (diagonal v)⁻¹ = diagonal (Ring.inverse
607607

608608
end Diagonal
609609

610+
section Woodbury
611+
612+
variable [Fintype m] [DecidableEq m]
613+
variable (A : Matrix n n α) (U : Matrix n m α) (C : Matrix m m α) (V : Matrix m n α)
614+
615+
/-- The **Woodbury Identity** (`⁻¹` version). -/
616+
theorem add_mul_mul_inv_eq_sub (hA : IsUnit A) (hC : IsUnit C) (hAC : IsUnit (C⁻¹ + V * A⁻¹ * U)) :
617+
(A + U * C * V)⁻¹ = A⁻¹ - A⁻¹ * U * (C⁻¹ + V * A⁻¹ * U)⁻¹ * V * A⁻¹ := by
618+
obtain ⟨_⟩ := hA.nonempty_invertible
619+
obtain ⟨_⟩ := hC.nonempty_invertible
620+
obtain ⟨iAC⟩ := hAC.nonempty_invertible
621+
simp only [← invOf_eq_nonsing_inv] at iAC
622+
letI := invertibleAddMulMul A U C V
623+
simp only [← invOf_eq_nonsing_inv]
624+
apply invOf_add_mul_mul
625+
626+
end Woodbury
627+
610628
@[simp]
611629
theorem inv_inv_inv (A : Matrix n n α) : A⁻¹⁻¹⁻¹ = A⁻¹ := by
612630
by_cases h : IsUnit A.det

0 commit comments

Comments
 (0)