Skip to content

Commit

Permalink
feat: totally unimodular matrices
Browse files Browse the repository at this point in the history
  • Loading branch information
madvorak committed Oct 31, 2024
1 parent ecb2a9e commit 1619e6a
Show file tree
Hide file tree
Showing 2 changed files with 97 additions and 0 deletions.
1 change: 1 addition & 0 deletions Mathlib.lean
Original file line number Diff line number Diff line change
Expand Up @@ -3175,6 +3175,7 @@ import Mathlib.LinearAlgebra.Matrix.Charpoly.Univ
import Mathlib.LinearAlgebra.Matrix.Circulant
import Mathlib.LinearAlgebra.Matrix.Determinant.Basic
import Mathlib.LinearAlgebra.Matrix.Determinant.Misc
import Mathlib.LinearAlgebra.Matrix.Determinant.TotallyUnimodular
import Mathlib.LinearAlgebra.Matrix.Diagonal
import Mathlib.LinearAlgebra.Matrix.DotProduct
import Mathlib.LinearAlgebra.Matrix.Dual
Expand Down
96 changes: 96 additions & 0 deletions Mathlib/LinearAlgebra/Matrix/Determinant/TotallyUnimodular.lean
Original file line number Diff line number Diff line change
@@ -0,0 +1,96 @@
/-
Copyright (c) 2024 Martin Dvorak. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Martin Dvorak, Vladimir Kolmogorov, Ivan Sergeev
-/
import Mathlib.LinearAlgebra.Matrix.Determinant.Basic

/-!
# Totally unimodular matrices
This file defines totally unimodular matrices and provides basic API for them.
## Main definitions
- `Matrix.IsTotallyUnimodular`: a matrix is totally unimodular iff every square submatrix
(not necessarily contiguous) has determinant `0` or `1` or `-1`.
## Main results
- `Matrix.IsTotallyUnimodular_iff`: a matrix is totally unimodular iff every square submatrix
(possibly with repeated rows and/or repeated columns) has determinant `0` or `1` or `-1`.
- `Matrix.IsTotallyUnimodular.apply`: entry in a totally unimodular matrix is `0` or `1` or `-1`.
-/

open scoped Matrix

variable {m n R : Type*} [CommRing R]

/-- Is the matrix `A` totally unimodular? -/
def Matrix.IsTotallyUnimodular (A : Matrix m n R) : Prop :=
∀ k : ℕ, ∀ f : Fin k → m, ∀ g : Fin k → n, f.Injective → g.Injective →
(A.submatrix f g).det = 0
(A.submatrix f g).det = 1
(A.submatrix f g).det = -1

lemma Matrix.IsTotallyUnimodular_iff (A : Matrix m n R) : A.IsTotallyUnimodular ↔
∀ k : ℕ, ∀ f : Fin k → m, ∀ g : Fin k → n,
(A.submatrix f g).det = 0
(A.submatrix f g).det = 1
(A.submatrix f g).det = -1
:= by

Check failure on line 42 in Mathlib/LinearAlgebra/Matrix/Determinant/TotallyUnimodular.lean

View workflow job for this annotation

GitHub Actions / Lint style

Mathlib/LinearAlgebra/Matrix/Determinant/TotallyUnimodular.lean:42 ERR_CLN: Put : and := before line breaks, not after

Check failure on line 42 in Mathlib/LinearAlgebra/Matrix/Determinant/TotallyUnimodular.lean

View workflow job for this annotation

GitHub Actions / Lint style

Mathlib/LinearAlgebra/Matrix/Determinant/TotallyUnimodular.lean:42 ERR_CLN: Put : and := before line breaks, not after
constructor <;> intro hA
· intro k f g
if hf : f.Injective then
if hg : g.Injective then
exact hA k f g hf hg
else
left
unfold Function.Injective at hg
push_neg at hg
obtain ⟨i, j, hqij, hij⟩ := hg
apply Matrix.det_zero_of_column_eq hij
intro
simp [hqij]
else
left
unfold Function.Injective at hf
push_neg at hf
obtain ⟨i, j, hpij, hij⟩ := hf
apply Matrix.det_zero_of_row_eq
· exact hij
show (A (f i)) ∘ (g ·) = (A (f j)) ∘ (g ·)
rw [hpij]
· intro _ _ _ _ _
apply hA

lemma Matrix.IsTotallyUnimodular.apply {A : Matrix m n R} (hA : A.IsTotallyUnimodular)
(i : m) (j : n) :
A i j = 0 ∨ A i j = 1 ∨ A i j = -1 := by
let f : Fin 1 → m := (fun _ => i)
let g : Fin 1 → n := (fun _ => j)
convert hA 1 f g (Function.injective_of_subsingleton f) (Function.injective_of_subsingleton g) <;>
exact (det_fin_one (A.submatrix f g)).symm

lemma Matrix.IsTotallyUnimodular.submatrix {A : Matrix m n R} (hA : A.IsTotallyUnimodular) {k : ℕ}
(f : Fin k → m) (g : Fin k → n) :
(A.submatrix f g).IsTotallyUnimodular := by
intro _ _ _ _ _
rw [Matrix.submatrix_submatrix]
rw [Matrix.IsTotallyUnimodular_iff] at hA
apply hA

lemma Matrix.IsTotallyUnimodular.transpose {A : Matrix m n R} (hA : A.IsTotallyUnimodular) :
Aᵀ.IsTotallyUnimodular := by
intro _ _ _ _ _
simp only [←Matrix.transpose_submatrix, Matrix.det_transpose]

Check failure on line 87 in Mathlib/LinearAlgebra/Matrix/Determinant/TotallyUnimodular.lean

View workflow job for this annotation

GitHub Actions / Lint style

Mathlib/LinearAlgebra/Matrix/Determinant/TotallyUnimodular.lean:87 ERR_ARR: Missing space after '←'.

Check failure on line 87 in Mathlib/LinearAlgebra/Matrix/Determinant/TotallyUnimodular.lean

View workflow job for this annotation

GitHub Actions / Lint style

Mathlib/LinearAlgebra/Matrix/Determinant/TotallyUnimodular.lean:87 ERR_ARR: Missing space after '←'.
apply hA <;> assumption

lemma Matrix.mapEquiv_IsTotallyUnimodular {X' Y' : Type*}
(A : Matrix m n R) (eX : X' ≃ m) (eY : Y' ≃ n) :
Matrix.IsTotallyUnimodular ((A · ∘ eY) ∘ eX) ↔ A.IsTotallyUnimodular := by
rw [Matrix.IsTotallyUnimodular_iff, Matrix.IsTotallyUnimodular_iff]
constructor <;> intro hA k f g
· simpa [Matrix.submatrix] using hA k (eX.symm ∘ f) (eY.symm ∘ g)
· simpa [Matrix.submatrix] using hA k (eX ∘ f) (eY ∘ g)

0 comments on commit 1619e6a

Please sign in to comment.