diff --git a/PolyhedralCombinatorics/LinearSystem/Basic.lean b/PolyhedralCombinatorics/LinearSystem/Basic.lean index 9d0668d..b9d5820 100644 --- a/PolyhedralCombinatorics/LinearSystem/Basic.lean +++ b/PolyhedralCombinatorics/LinearSystem/Basic.lean @@ -25,7 +25,7 @@ with `LinearSystem 𝔽 n`. -/ variable - (𝔽 : Type*) [LinearOrderedField 𝔽] -- Field type + (𝔽 : Type*) -- Field type (n : β„•) -- Dimension of the space /-- `LinearSystem 𝔽 n` is the type of linear systems of inequalities in `𝔽^n`, where `𝔽` is a @@ -41,14 +41,51 @@ structure LinearSystem where namespace LinearSystem open Matrix -variable {𝔽 m n} (p q : LinearSystem 𝔽 n) +variable {𝔽 m n} /-- Constructs a linear system defined by a `m Γ— n` coefficient matrix `A` and an `m Γ— 1` column vector `b`. -/ -abbrev of (A : Matrix (Fin m) (Fin n) 𝔽) (b : Fin m β†’ 𝔽) : LinearSystem 𝔽 n := ⟨m, A, b⟩ +@[match_pattern] abbrev of (A : Matrix (Fin m) (Fin n) 𝔽) (b : Fin m β†’ 𝔽) : LinearSystem 𝔽 n := + ⟨m, A, b⟩ + +def empty : LinearSystem 𝔽 n := of vecEmpty vecEmpty + +/-- Prepends a linear inequality `yα΅€ x ≀ b` to a linear system `p`. -/ +def cons (y : Fin n β†’ 𝔽) (b : 𝔽) (p : LinearSystem 𝔽 n) : LinearSystem 𝔽 n := + of (vecCons y p.mat) (vecCons b p.vec) + +@[simp] lemma cons_m : (cons y b p : LinearSystem 𝔽 n).m = p.m + 1 := rfl + +@[simp] lemma cons_mat : (cons y b p : LinearSystem 𝔽 n).mat = vecCons y p.mat := rfl + +@[simp] lemma cons_vec : (cons y b p : LinearSystem 𝔽 n).vec = vecCons b p.vec := rfl + +@[elab_as_elim] def induction {motive : LinearSystem 𝔽 n β†’ Sort _} + (empty : motive empty) (cons : βˆ€ y b p, motive p β†’ motive (cons y b p)) : βˆ€ p, motive p + | ⟨m, A, b⟩ => by + induction m + case zero => + have : A = vecEmpty := by funext x; exact Fin.elim0 x + have : b = vecEmpty := by funext x; exact Fin.elim0 x + subst A b + exact empty + case succ m ih => + have : ⟨_, A, b⟩ = LinearSystem.cons (vecHead A) (vecHead b) (of (vecTail A) (vecTail b)) := by + simp only [LinearSystem.cons, Matrix.cons_head_tail] + simp_rw [this] + apply cons + apply ih + +@[elab_as_elim] def inductionOn {motive : LinearSystem 𝔽 n β†’ Sort _} + (p : LinearSystem 𝔽 n) (empty : motive empty) (cons : βˆ€ y b p, motive p β†’ motive (cons y b p)) + : motive p := induction empty cons p + +@[elab_as_elim] def cases {motive : LinearSystem 𝔽 n β†’ Sort _} + (empty : motive empty) (cons : βˆ€ y b p, motive (cons y b p)) + : βˆ€ p, motive p := induction empty fun y b p _ => cons y b p /-- The concatenation of two linear systems. -/ -def concat : LinearSystem 𝔽 n := +def concat (p q : LinearSystem 𝔽 n) : LinearSystem 𝔽 n := { m := p.m + q.m, mat := Matrix.of fun i j ↦ @@ -63,6 +100,8 @@ def concat : LinearSystem 𝔽 n := q.vec ⟨i - p.m, Nat.sub_lt_left_of_lt_add (not_lt.mp h) i.prop⟩ } +variable [LinearOrderedField 𝔽] (p : LinearSystem 𝔽 n) + /-- The set of solutions of the linear system. -/ @[simp] def solutions : Set (Fin n β†’ 𝔽) := { x | p.mat *α΅₯ x ≀ p.vec } diff --git a/PolyhedralCombinatorics/LinearSystem/LinearConstraints.lean b/PolyhedralCombinatorics/LinearSystem/LinearConstraints.lean index d1bf221..c02be8f 100644 --- a/PolyhedralCombinatorics/LinearSystem/LinearConstraints.lean +++ b/PolyhedralCombinatorics/LinearSystem/LinearConstraints.lean @@ -28,44 +28,22 @@ variable (𝔽 : Type*) [LinearOrderedField 𝔽] -- Field type (n : β„•) -- Dimension of the space -inductive LinearConstraint.Comparator | le | eq | ge - /-- `LinearConstraint n` is the type of linear constraints on `n` variables. -/ -structure LinearConstraint where - coefficients : Fin n β†’ 𝔽 - comparator : LinearConstraint.Comparator - rhs : 𝔽 +inductive LinearConstraint (𝔽 : Type*) (n : β„•) where +| le : (Fin n β†’ 𝔽) β†’ 𝔽 β†’ LinearConstraint 𝔽 n +| eq : (Fin n β†’ 𝔽) β†’ 𝔽 β†’ LinearConstraint 𝔽 n +| ge : (Fin n β†’ 𝔽) β†’ 𝔽 β†’ LinearConstraint 𝔽 n variable {𝔽 n} {y : Fin n β†’ 𝔽} {b : 𝔽} namespace LinearConstraint open Matrix -@[match_pattern] abbrev le (y : Fin n β†’ 𝔽) (b : 𝔽) : LinearConstraint 𝔽 n := ⟨y, .le, b⟩ -@[match_pattern] abbrev eq (y : Fin n β†’ 𝔽) (b : 𝔽) : LinearConstraint 𝔽 n := ⟨y, .eq, b⟩ -@[match_pattern] abbrev ge (y : Fin n β†’ 𝔽) (b : 𝔽) : LinearConstraint 𝔽 n := ⟨y, .ge, b⟩ - -/-- Converts a linear constraint into a list o extended matrix rows for a - linear system. -/ -def toExtendedRows : LinearConstraint 𝔽 n β†’ List ((Fin n β†’ 𝔽) Γ— 𝔽) -| le y b => [(y, b)] -| eq y b => [(y, b), (-y, -b)] -| ge y b => [(-y, -b)] - -@[simp] lemma le_toExtendedRows : toExtendedRows (le y b) = [(y, b)] := rfl -@[simp] lemma eq_toExtendedRows : toExtendedRows (eq y b) = [(y, b), (-y, -b)] := rfl -@[simp] lemma ge_toExtendedRows : toExtendedRows (ge y b) = [(-y, -b)] := rfl - -def Comparator.valid : Comparator β†’ 𝔽 β†’ 𝔽 β†’ Prop -| le => (Β· ≀ Β·) | eq => (Β· = Β·) | ge => (Β· β‰₯ Β·) - /-- Whether a linear constraint is valid for a given vector. -/ -def valid : LinearConstraint 𝔽 n β†’ (Fin n β†’ 𝔽) β†’ Prop -| ⟨y, cmp, b⟩, x => cmp.valid (y ⬝α΅₯ x) b - -@[simp] lemma le_valid : (le y b).valid x = (y ⬝α΅₯ x ≀ b) := rfl -@[simp] lemma eq_valid : (eq y b).valid x = (y ⬝α΅₯ x = b) := rfl -@[simp] lemma ge_valid : (ge y b).valid x = (y ⬝α΅₯ x β‰₯ b) := rfl +@[simp] def valid : LinearConstraint 𝔽 n β†’ (Fin n β†’ 𝔽) β†’ Prop +| le y b, x => y ⬝α΅₯ x ≀ b +| eq y b, x => y ⬝α΅₯ x = b +| ge y b, x => y ⬝α΅₯ x β‰₯ b lemma eq_valid_iff : (eq y b).valid x ↔ (le y b).valid x ∧ (ge y b).valid x := by simp [le_antisymm_iff] @@ -74,86 +52,72 @@ end LinearConstraint namespace LinearSystem -open LinearConstraint +open LinearConstraint Matrix /-- Constructs a linear system from a list of linear constraints. -/ -def ofConstraints (cs : List (LinearConstraint 𝔽 n)) : LinearSystem 𝔽 n := - let rows := cs.bind toExtendedRows - ⟨rows.length, Matrix.of (Prod.fst ∘ rows.get), Prod.snd ∘ rows.get⟩ +def ofConstraints : List (LinearConstraint 𝔽 n) β†’ LinearSystem 𝔽 n + | [] => of vecEmpty vecEmpty + | c :: cs => + match c with + | le y b => cons y b $ ofConstraints cs + | ge y b => cons (-y) (-b) $ ofConstraints cs + | eq y b => cons y b $ cons (-y) (-b) $ ofConstraints cs section toSet_ofConstraints open Matrix LinearConstraint variable {cs : List (LinearConstraint 𝔽 n)} {y : Fin n β†’ 𝔽} {b : 𝔽} -@[simp] -lemma ofConstraints_mat_apply : (ofConstraints cs).mat i = (cs.bind toExtendedRows)[i].1 := rfl - -@[simp] -lemma ofConstraints_vec_apply : (ofConstraints cs).vec i = (cs.bind toExtendedRows)[i].2 := rfl - -@[simp] lemma ofConstraints_apply_eq_lemma - : let s := ofConstraints cs - s.mat i = y ∧ s.vec i = b ↔ (cs.bind toExtendedRows)[i] = (y, b) := by - simp only [ofConstraints_mat_apply, ofConstraints_vec_apply] - apply Prod.mk.inj_iff.symm - -macro "linear_constraint_lemma_tac" h:ident : tactic => - `(tactic| - simp_rw [ofConstraints_apply_eq_lemma] - <;> exact List.mem_iff_get.mp $ List.mem_bind.mpr ⟨_, $h, by simp [toExtendedRows]⟩) - -private lemma le_lemma (h : ⟨y, Comparator.le, b⟩ ∈ cs) - : let s := ofConstraints cs; βˆƒ i, s.mat i = y ∧ s.vec i = b := by - linear_constraint_lemma_tac h - -private lemma ge_lemma (h : ⟨y, Comparator.ge, b⟩ ∈ cs) - : let s := ofConstraints cs; βˆƒ i, -(s.mat i) = y ∧ -(s.vec i) = b := by - simp_rw [neg_eq_iff_eq_neg] - linear_constraint_lemma_tac h - -private lemma eq_lemma (h : ⟨y, Comparator.eq, b⟩ ∈ cs) - : let s := ofConstraints cs; - (βˆƒ i, s.mat i = y ∧ s.vec i = b) ∧ (βˆƒ i, -(s.mat i) = y ∧ -(s.vec i) = b) := by - simp_rw [neg_eq_iff_eq_neg] - constructor <;> linear_constraint_lemma_tac h +@[simp] lemma ofConstraints_nil + : (@ofConstraints 𝔽 _ n []) = of vecEmpty vecEmpty := rfl + +theorem mem_ofConstraints_nil_solutions : x ∈ (@ofConstraints 𝔽 _ n []).solutions := by simp + +@[simp] private lemma vecCons_mulVec + {m n : β„•} (y : Fin n β†’ 𝔽) (A : Matrix (Fin m) (Fin n) 𝔽) (x : Fin n β†’ 𝔽) + : vecCons y A *α΅₯ x = vecCons (y ⬝α΅₯ x) (A *α΅₯ x) := by + funext x + cases x using Fin.cases <;> rfl + +@[simp] private lemma vecCons_le_vecCons {n : β„•} (a b : 𝔽) (x y : Fin n β†’ 𝔽) + : vecCons a x ≀ vecCons b y ↔ a ≀ b ∧ x ≀ y := by + simp_rw [Pi.le_def] + constructor <;> intro h + . constructor + . exact h 0 + . intro i + exact h i.succ + . intro i + cases i using Fin.cases + . simp only [cons_val_zero, h.1] + . simp only [cons_val_succ, h.2] + +@[simp] lemma mem_ofConstraints_cons_solutions + : x ∈ (@ofConstraints 𝔽 _ n $ c :: cs).solutions ↔ c.valid x ∧ x ∈ (ofConstraints cs).solutions := by + simp only [solutions, Set.mem_setOf_eq, valid, ge_iff_le] + rw [ofConstraints] + split <;> simp [le_antisymm_iff, and_assoc] + +@[simp] lemma ofConstraints_le_cons + : (ofConstraints $ le y b :: cs) = cons y b (ofConstraints cs) := by + rfl + +@[simp] lemma ofConstraints_ge_cons + : (ofConstraints $ ge y b :: cs) = cons (-y) (-b) (ofConstraints cs) := by + rfl + +@[simp] lemma ofConstraints_eq_cons + : (ofConstraints $ eq y b :: cs) = + cons y b (cons (-y) (-b) $ ofConstraints cs) := by + rfl @[simp] theorem mem_solutions_ofConstraints (x : Fin n β†’ 𝔽) : x ∈ (ofConstraints cs).solutions ↔ βˆ€ c ∈ cs, c.valid x := by - let rows := cs.bind toExtendedRows - constructor - . show x ∈ (ofConstraints cs).solutions β†’ βˆ€ c ∈ cs, c.valid _ - intro h ⟨y, cmp, b⟩ mem_cs - cases cmp - case' le => have ⟨_, hy, hb⟩ := le_lemma mem_cs - case' ge => have ⟨_, hy, hb⟩ := ge_lemma mem_cs - case' eq => - rw [eq_valid_iff] - constructor - case' left => have ⟨_, hy, hb⟩ := (eq_lemma mem_cs).1 - case' right => have ⟨_, hy, hb⟩ := (eq_lemma mem_cs).2 - all_goals subst hy hb - case le | eq.left => - apply h - case ge | eq.right => - simp_rw [ge_valid, neg_dotProduct, neg_le_neg_iff] - apply h - . show (βˆ€ c ∈ cs, c.valid x) β†’ βˆ€ (i : Fin rows.length), rows[i].1 ⬝α΅₯ _ ≀ rows[i].2 - intro h i - have : rows[i] ∈ rows := List.get_mem rows i i.prop - have ⟨⟨y, cmp, b⟩, mem_cs, h'⟩ := List.mem_bind.mp this - have c_valid := h _ mem_cs - cases cmp - case' eq => - rw [eq_valid_iff] at c_valid - rw [toExtendedRows, List.mem_pair] at h' - cases h' - case le | eq.inl => - rw [le_valid] at c_valid - simp_all only [toExtendedRows, List.mem_singleton] - case ge | eq.inr => - rw [ge_valid] at c_valid - simp_all only [toExtendedRows, List.mem_singleton, neg_dotProduct, neg_le_neg_iff] + induction cs + case nil => + simp + case cons c cs ih => + simp only [mem_ofConstraints_cons_solutions, List.forall_mem_cons, ih] /-- The set of solutions of a linear system of constraints is the set of all points that satisfy all constraints. -/ diff --git a/PolyhedralCombinatorics/Polyhedron/Notation.lean b/PolyhedralCombinatorics/Polyhedron/Notation.lean index b28627a..0dbd0eb 100644 --- a/PolyhedralCombinatorics/Polyhedron/Notation.lean +++ b/PolyhedralCombinatorics/Polyhedron/Notation.lean @@ -5,7 +5,7 @@ open Matrix LinearSystem variable {𝔽} [LinearOrderedField 𝔽] {n : β„•} -open Lean.Parser Lean.Elab.Term LinearConstraint.Comparator +open Lean.Parser Lean.Elab.Term @[inherit_doc ofLinearSystem] scoped notation:max (name := matVecPolyhedron) "P(" A " , " b ")" => ofLinearSystem $ of A b @@ -31,11 +31,11 @@ macro_rules | `(!P[$t^$n]{$[$constraints],*}) => `((!P{$constraints,*} : Polyhedron $t $n)) | `(!P{$[$constraints],*}) => do let constraints ← constraints.mapM (fun - | `(linConstraint| $x:term ≀ $y:term) => `(⟨$x, le, $y⟩) - | `(linConstraint| $x:term <= $y:term) => `(⟨$x, le, $y⟩) - | `(linConstraint| $x:term = $y:term) => `(⟨$x, eq, $y⟩) - | `(linConstraint| $x:term β‰₯ $y:term) => `(⟨$x, ge, $y⟩) - | `(linConstraint| $x:term >= $y:term) => `(⟨$x, ge, $y⟩) + | `(linConstraint| $x:term ≀ $y:term) => `(LinearConstraint.le $x $y) + | `(linConstraint| $x:term <= $y:term) => `(LinearConstraint.le $x $y) + | `(linConstraint| $x:term = $y:term) => `(LinearConstraint.eq $x $y) + | `(linConstraint| $x:term β‰₯ $y:term) => `(LinearConstraint.ge $x $y) + | `(linConstraint| $x:term >= $y:term) => `(LinearConstraint.ge $x $y) | _ => Lean.Macro.throwUnsupported) `(ofLinearSystem $ ofConstraints [$constraints,*]) diff --git a/PolyhedralCombinatorics/Polyhedron/Projection.lean b/PolyhedralCombinatorics/Polyhedron/Projection.lean index 0ff0ba6..cba837c 100644 --- a/PolyhedralCombinatorics/Polyhedron/Projection.lean +++ b/PolyhedralCombinatorics/Polyhedron/Projection.lean @@ -198,7 +198,7 @@ theorem mem_fourierMotzkin (p : Polyhedron 𝔽 n) (j : Fin n) : x ∈ p.fourierMotzkin j ↔ x j = 0 ∧ βˆƒ (Ξ± : 𝔽), x + Pi.single j Ξ± ∈ p := by simp_rw [ fourierMotzkin, mem_projection, mem_ofConstraints, - List.mem_singleton, forall_eq, LinearConstraint.eq_valid, + List.mem_singleton, forall_eq, LinearConstraint.valid, single_dotProduct, one_mul, and_congr_right_iff, ←Pi.single_smul, smul_eq_mul, mul_one, implies_true ]